src/HOL/Analysis/Convex_Euclidean_Space.thy
author paulson <lp15@cam.ac.uk>
Mon, 04 Nov 2019 17:06:18 +0000
changeset 71029 934e0044e94b
parent 71008 e892f7a1fd83
child 71030 b6e69c71a9f6
permissions -rw-r--r--
Moved or deleted some out of place material, also eliminating obsolete naming conventions
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
63969
f4b4fba60b1d HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents: 63967
diff changeset
     1
(* Title:      HOL/Analysis/Convex_Euclidean_Space.thy
f4b4fba60b1d HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents: 63967
diff changeset
     2
   Author:     L C Paulson, University of Cambridge
f4b4fba60b1d HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents: 63967
diff changeset
     3
   Author:     Robert Himmelmann, TU Muenchen
f4b4fba60b1d HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents: 63967
diff changeset
     4
   Author:     Bogdan Grechuk, University of Edinburgh
f4b4fba60b1d HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents: 63967
diff changeset
     5
   Author:     Armin Heller, TU Muenchen
f4b4fba60b1d HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents: 63967
diff changeset
     6
   Author:     Johannes Hoelzl, TU Muenchen
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
     7
*)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
     8
69619
3f7d8e05e0f2 split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents: 69618
diff changeset
     9
section \<open>Convex Sets and Functions on (Normed) Euclidean Spaces\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
    10
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
    11
theory Convex_Euclidean_Space
44132
0f35a870ecf1 full import paths
huffman
parents: 44125
diff changeset
    12
imports
69619
3f7d8e05e0f2 split off Convex.thy: material that does not require Topology_Euclidean_Space
immler
parents: 69618
diff changeset
    13
  Convex
69617
63ee37c519a3 reduced dependencies of Connected.thy
immler
parents: 69529
diff changeset
    14
  Topology_Euclidean_Space
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
    15
begin
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
    16
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70097
diff changeset
    17
subsection\<^marker>\<open>tag unimportant\<close> \<open>Topological Properties of Convex Sets and Functions\<close>
63969
f4b4fba60b1d HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents: 63967
diff changeset
    18
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
    19
lemma convex_supp_sum:
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
    20
  assumes "convex S" and 1: "supp_sum u I = 1"
63969
f4b4fba60b1d HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents: 63967
diff changeset
    21
      and "\<And>i. i \<in> I \<Longrightarrow> 0 \<le> u i \<and> (u i = 0 \<or> f i \<in> S)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
    22
    shows "supp_sum (\<lambda>i. u i *\<^sub>R f i) I \<in> S"
63969
f4b4fba60b1d HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents: 63967
diff changeset
    23
proof -
f4b4fba60b1d HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents: 63967
diff changeset
    24
  have fin: "finite {i \<in> I. u i \<noteq> 0}"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
    25
    using 1 sum.infinite by (force simp: supp_sum_def support_on_def)
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
    26
  then have eq: "supp_sum (\<lambda>i. u i *\<^sub>R f i) I = sum (\<lambda>i. u i *\<^sub>R f i) {i \<in> I. u i \<noteq> 0}"
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
    27
    by (force intro: sum.mono_neutral_left simp: supp_sum_def support_on_def)
63969
f4b4fba60b1d HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents: 63967
diff changeset
    28
  show ?thesis
f4b4fba60b1d HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents: 63967
diff changeset
    29
    apply (simp add: eq)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
    30
    apply (rule convex_sum [OF fin \<open>convex S\<close>])
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
    31
    using 1 assms apply (auto simp: supp_sum_def support_on_def)
63969
f4b4fba60b1d HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents: 63967
diff changeset
    32
    done
f4b4fba60b1d HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents: 63967
diff changeset
    33
qed
f4b4fba60b1d HOL-Analysis: move Library/Convex to Convex_Euclidean_Space
hoelzl
parents: 63967
diff changeset
    34
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    35
lemma closure_bounded_linear_image_subset:
44524
04ad69081646 generalize some lemmas
huffman
parents: 44523
diff changeset
    36
  assumes f: "bounded_linear f"
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
    37
  shows "f ` closure S \<subseteq> closure (f ` S)"
44524
04ad69081646 generalize some lemmas
huffman
parents: 44523
diff changeset
    38
  using linear_continuous_on [OF f] closed_closure closure_subset
04ad69081646 generalize some lemmas
huffman
parents: 44523
diff changeset
    39
  by (rule image_closure_subset)
04ad69081646 generalize some lemmas
huffman
parents: 44523
diff changeset
    40
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    41
lemma closure_linear_image_subset:
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
    42
  fixes f :: "'m::euclidean_space \<Rightarrow> 'n::real_normed_vector"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
    43
  assumes "linear f"
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    44
  shows "f ` (closure S) \<subseteq> closure (f ` S)"
44524
04ad69081646 generalize some lemmas
huffman
parents: 44523
diff changeset
    45
  using assms unfolding linear_conv_bounded_linear
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    46
  by (rule closure_bounded_linear_image_subset)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    47
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    48
lemma closed_injective_linear_image:
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    49
    fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    50
    assumes S: "closed S" and f: "linear f" "inj f"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    51
    shows "closed (f ` S)"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    52
proof -
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    53
  obtain g where g: "linear g" "g \<circ> f = id"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    54
    using linear_injective_left_inverse [OF f] by blast
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    55
  then have confg: "continuous_on (range f) g"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    56
    using linear_continuous_on linear_conv_bounded_linear by blast
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    57
  have [simp]: "g ` f ` S = S"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    58
    using g by (simp add: image_comp)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    59
  have cgf: "closed (g ` f ` S)"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61762
diff changeset
    60
    by (simp add: \<open>g \<circ> f = id\<close> S image_comp)
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66827
diff changeset
    61
  have [simp]: "(range f \<inter> g -` S) = f ` S"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66827
diff changeset
    62
    using g unfolding o_def id_def image_def by auto metis+
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    63
  show ?thesis
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66827
diff changeset
    64
  proof (rule closedin_closed_trans [of "range f"])
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69710
diff changeset
    65
    show "closedin (top_of_set (range f)) (f ` S)"
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66827
diff changeset
    66
      using continuous_closedin_preimage [OF confg cgf] by simp
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66827
diff changeset
    67
    show "closed (range f)"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66827
diff changeset
    68
      apply (rule closed_injective_image_subspace)
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66827
diff changeset
    69
      using f apply (auto simp: linear_linear linear_injective_0)
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66827
diff changeset
    70
      done
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66827
diff changeset
    71
  qed
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    72
qed
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    73
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    74
lemma closed_injective_linear_image_eq:
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    75
    fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    76
    assumes f: "linear f" "inj f"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    77
      shows "(closed(image f s) \<longleftrightarrow> closed s)"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    78
  by (metis closed_injective_linear_image closure_eq closure_linear_image_subset closure_subset_eq f(1) f(2) inj_image_subset_iff)
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
    79
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
    80
lemma closure_injective_linear_image:
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    81
    fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    82
    shows "\<lbrakk>linear f; inj f\<rbrakk> \<Longrightarrow> f ` (closure S) = closure (f ` S)"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    83
  apply (rule subset_antisym)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    84
  apply (simp add: closure_linear_image_subset)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    85
  by (simp add: closure_minimal closed_injective_linear_image closure_subset image_mono)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    86
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    87
lemma closure_bounded_linear_image:
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    88
    fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    89
    shows "\<lbrakk>linear f; bounded S\<rbrakk> \<Longrightarrow> f ` (closure S) = closure (f ` S)"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    90
  apply (rule subset_antisym, simp add: closure_linear_image_subset)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    91
  apply (rule closure_minimal, simp add: closure_subset image_mono)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
    92
  by (meson bounded_closure closed_closure compact_continuous_image compact_eq_bounded_closed linear_continuous_on linear_conv_bounded_linear)
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
    93
44524
04ad69081646 generalize some lemmas
huffman
parents: 44523
diff changeset
    94
lemma closure_scaleR:
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
    95
  fixes S :: "'a::real_normed_vector set"
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68607
diff changeset
    96
  shows "((*\<^sub>R) c) ` (closure S) = closure (((*\<^sub>R) c) ` S)"
44524
04ad69081646 generalize some lemmas
huffman
parents: 44523
diff changeset
    97
proof
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68607
diff changeset
    98
  show "((*\<^sub>R) c) ` (closure S) \<subseteq> closure (((*\<^sub>R) c) ` S)"
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
    99
    using bounded_linear_scaleR_right
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   100
    by (rule closure_bounded_linear_image_subset)
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68607
diff changeset
   101
  show "closure (((*\<^sub>R) c) ` S) \<subseteq> ((*\<^sub>R) c) ` (closure S)"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   102
    by (intro closure_minimal image_mono closure_subset closed_scaling closed_closure)
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   103
qed
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   104
62381
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
   105
lemma sphere_eq_empty [simp]:
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
   106
  fixes a :: "'a::{real_normed_vector, perfect_space}"
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
   107
  shows "sphere a r = {} \<longleftrightarrow> r < 0"
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
   108
by (auto simp: sphere_def dist_norm) (metis dist_norm le_less_linear vector_choose_dist)
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   109
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   110
lemma cone_closure:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   111
  fixes S :: "'a::real_normed_vector set"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   112
  assumes "cone S"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   113
  shows "cone (closure S)"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   114
proof (cases "S = {}")
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   115
  case True
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   116
  then show ?thesis by auto
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   117
next
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   118
  case False
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68607
diff changeset
   119
  then have "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> (*\<^sub>R) c ` S = S)"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   120
    using cone_iff[of S] assms by auto
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68607
diff changeset
   121
  then have "0 \<in> closure S \<and> (\<forall>c. c > 0 \<longrightarrow> (*\<^sub>R) c ` closure S = closure S)"
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
   122
    using closure_subset by (auto simp: closure_scaleR)
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   123
  then show ?thesis
60974
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
   124
    using False cone_iff[of "closure S"] by auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   125
qed
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   126
66939
04678058308f New results in topology, mostly from HOL Light's moretop.ml
paulson <lp15@cam.ac.uk>
parents: 66884
diff changeset
   127
corollary component_complement_connected:
04678058308f New results in topology, mostly from HOL Light's moretop.ml
paulson <lp15@cam.ac.uk>
parents: 66884
diff changeset
   128
  fixes S :: "'a::real_normed_vector set"
04678058308f New results in topology, mostly from HOL Light's moretop.ml
paulson <lp15@cam.ac.uk>
parents: 66884
diff changeset
   129
  assumes "connected S" "C \<in> components (-S)"
04678058308f New results in topology, mostly from HOL Light's moretop.ml
paulson <lp15@cam.ac.uk>
parents: 66884
diff changeset
   130
  shows "connected(-C)"
04678058308f New results in topology, mostly from HOL Light's moretop.ml
paulson <lp15@cam.ac.uk>
parents: 66884
diff changeset
   131
  using component_diff_connected [of S UNIV] assms
04678058308f New results in topology, mostly from HOL Light's moretop.ml
paulson <lp15@cam.ac.uk>
parents: 66884
diff changeset
   132
  by (auto simp: Compl_eq_Diff_UNIV)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   133
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62097
diff changeset
   134
proposition clopen:
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66827
diff changeset
   135
  fixes S :: "'a :: real_normed_vector set"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66827
diff changeset
   136
  shows "closed S \<and> open S \<longleftrightarrow> S = {} \<or> S = UNIV"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66827
diff changeset
   137
    by (force intro!: connected_UNIV [unfolded connected_clopen, rule_format])
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62097
diff changeset
   138
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62097
diff changeset
   139
corollary compact_open:
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66827
diff changeset
   140
  fixes S :: "'a :: euclidean_space set"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66827
diff changeset
   141
  shows "compact S \<and> open S \<longleftrightarrow> S = {}"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62097
diff changeset
   142
  by (auto simp: compact_eq_bounded_closed clopen)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62097
diff changeset
   143
62948
7700f467892b lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 62843
diff changeset
   144
corollary finite_imp_not_open:
7700f467892b lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 62843
diff changeset
   145
    fixes S :: "'a::{real_normed_vector, perfect_space} set"
7700f467892b lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 62843
diff changeset
   146
    shows "\<lbrakk>finite S; open S\<rbrakk> \<Longrightarrow> S={}"
7700f467892b lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 62843
diff changeset
   147
  using clopen [of S] finite_imp_closed not_bounded_UNIV by blast
7700f467892b lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 62843
diff changeset
   148
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
   149
corollary empty_interior_finite:
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
   150
    fixes S :: "'a::{real_normed_vector, perfect_space} set"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
   151
    shows "finite S \<Longrightarrow> interior S = {}"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
   152
  by (metis interior_subset finite_subset open_interior [of S] finite_imp_not_open)
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
   153
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
   154
text \<open>Balls, being convex, are connected.\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   155
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   156
lemma convex_local_global_minimum:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   157
  fixes s :: "'a::real_normed_vector set"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   158
  assumes "e > 0"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   159
    and "convex_on s f"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   160
    and "ball x e \<subseteq> s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   161
    and "\<forall>y\<in>ball x e. f x \<le> f y"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   162
  shows "\<forall>y\<in>s. f x \<le> f y"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   163
proof (rule ccontr)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   164
  have "x \<in> s" using assms(1,3) by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   165
  assume "\<not> ?thesis"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   166
  then obtain y where "y\<in>s" and y: "f x > f y" by auto
62087
44841d07ef1d revisions to limits and derivatives, plus new lemmas
paulson
parents: 61952
diff changeset
   167
  then have xy: "0 < dist x y"  by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   168
  then obtain u where "0 < u" "u \<le> 1" and u: "u < e / dist x y"
68527
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68074
diff changeset
   169
    using field_lbound_gt_zero[of 1 "e / dist x y"] xy \<open>e>0\<close> by auto
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   170
  then have "f ((1-u) *\<^sub>R x + u *\<^sub>R y) \<le> (1-u) * f x + u * f y"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
   171
    using \<open>x\<in>s\<close> \<open>y\<in>s\<close>
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   172
    using assms(2)[unfolded convex_on_def,
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   173
      THEN bspec[where x=x], THEN bspec[where x=y], THEN spec[where x="1-u"]]
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
   174
    by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   175
  moreover
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
   176
  have *: "x - ((1 - u) *\<^sub>R x + u *\<^sub>R y) = u *\<^sub>R (x - y)"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
   177
    by (simp add: algebra_simps)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
   178
  have "(1 - u) *\<^sub>R x + u *\<^sub>R y \<in> ball x e"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   179
    unfolding mem_ball dist_norm
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
   180
    unfolding * and norm_scaleR and abs_of_pos[OF \<open>0<u\<close>]
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
   181
    unfolding dist_norm[symmetric]
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   182
    using u
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   183
    unfolding pos_less_divide_eq[OF xy]
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   184
    by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   185
  then have "f x \<le> f ((1 - u) *\<^sub>R x + u *\<^sub>R y)"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   186
    using assms(4) by auto
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
   187
  ultimately show False
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
   188
    using mult_strict_left_mono[OF y \<open>u>0\<close>]
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   189
    unfolding left_diff_distrib
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   190
    by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   191
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   192
60800
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
   193
lemma convex_ball [iff]:
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   194
  fixes x :: "'a::real_normed_vector"
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
   195
  shows "convex (ball x e)"
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
   196
proof (auto simp: convex_def)
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
   197
  fix y z
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
   198
  assume yz: "dist x y < e" "dist x z < e"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
   199
  fix u v :: real
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
   200
  assume uv: "0 \<le> u" "0 \<le> v" "u + v = 1"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
   201
  have "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> u * dist x y + v * dist x z"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
   202
    using uv yz
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
   203
    using convex_on_dist [of "ball x e" x, unfolded convex_on_def,
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   204
      THEN bspec[where x=y], THEN bspec[where x=z]]
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
   205
    by auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
   206
  then show "dist x (u *\<^sub>R y + v *\<^sub>R z) < e"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
   207
    using convex_bound_lt[OF yz uv] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   208
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   209
60800
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
   210
lemma convex_cball [iff]:
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   211
  fixes x :: "'a::real_normed_vector"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   212
  shows "convex (cball x e)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   213
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   214
  {
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   215
    fix y z
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   216
    assume yz: "dist x y \<le> e" "dist x z \<le> e"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   217
    fix u v :: real
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   218
    assume uv: "0 \<le> u" "0 \<le> v" "u + v = 1"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   219
    have "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> u * dist x y + v * dist x z"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   220
      using uv yz
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
   221
      using convex_on_dist [of "cball x e" x, unfolded convex_on_def,
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   222
        THEN bspec[where x=y], THEN bspec[where x=z]]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   223
      by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   224
    then have "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> e"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   225
      using convex_bound_le[OF yz uv] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   226
  }
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
   227
  then show ?thesis by (auto simp: convex_def Ball_def)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   228
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   229
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   230
lemma connected_ball [iff]:
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   231
  fixes x :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   232
  shows "connected (ball x e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   233
  using convex_connected convex_ball by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   234
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   235
lemma connected_cball [iff]:
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   236
  fixes x :: "'a::real_normed_vector"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   237
  shows "connected (cball x e)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   238
  using convex_connected convex_cball by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   239
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
   240
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   241
lemma bounded_convex_hull:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   242
  fixes s :: "'a::real_normed_vector set"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   243
  assumes "bounded s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   244
  shows "bounded (convex hull s)"
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
   245
proof -
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
   246
  from assms obtain B where B: "\<forall>x\<in>s. norm x \<le> B"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
   247
    unfolding bounded_iff by auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
   248
  show ?thesis
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
   249
    apply (rule bounded_subset[OF bounded_cball, of _ 0 B])
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44142
diff changeset
   250
    unfolding subset_hull[of convex, OF convex_cball]
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   251
    unfolding subset_eq mem_cball dist_norm using B
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   252
    apply auto
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
   253
    done
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
   254
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   255
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   256
lemma finite_imp_bounded_convex_hull:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   257
  fixes s :: "'a::real_normed_vector set"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   258
  shows "finite s \<Longrightarrow> bounded (convex hull s)"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   259
  using bounded_convex_hull finite_imp_bounded
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   260
  by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   261
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   262
lemma aff_dim_cball:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   263
  fixes a :: "'n::euclidean_space"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   264
  assumes "e > 0"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   265
  shows "aff_dim (cball a e) = int (DIM('n))"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   266
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   267
  have "(\<lambda>x. a + x) ` (cball 0 e) \<subseteq> cball a e"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   268
    unfolding cball_def dist_norm by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   269
  then have "aff_dim (cball (0 :: 'n::euclidean_space) e) \<le> aff_dim (cball a e)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   270
    using aff_dim_translation_eq[of a "cball 0 e"]
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67155
diff changeset
   271
          aff_dim_subset[of "(+) a ` cball 0 e" "cball a e"]
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   272
    by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   273
  moreover have "aff_dim (cball (0 :: 'n::euclidean_space) e) = int (DIM('n))"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   274
    using hull_inc[of "(0 :: 'n::euclidean_space)" "cball 0 e"]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   275
      centre_in_cball[of "(0 :: 'n::euclidean_space)"] assms
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   276
    by (simp add: dim_cball[of e] aff_dim_zero[of "cball 0 e"])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   277
  ultimately show ?thesis
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
   278
    using aff_dim_le_DIM[of "cball a e"] by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   279
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   280
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   281
lemma aff_dim_open:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   282
  fixes S :: "'n::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   283
  assumes "open S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   284
    and "S \<noteq> {}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   285
  shows "aff_dim S = int (DIM('n))"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   286
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   287
  obtain x where "x \<in> S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   288
    using assms by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   289
  then obtain e where e: "e > 0" "cball x e \<subseteq> S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   290
    using open_contains_cball[of S] assms by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   291
  then have "aff_dim (cball x e) \<le> aff_dim S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   292
    using aff_dim_subset by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   293
  with e show ?thesis
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
   294
    using aff_dim_cball[of e x] aff_dim_le_DIM[of S] by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   295
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   296
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   297
lemma low_dim_interior:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   298
  fixes S :: "'n::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   299
  assumes "\<not> aff_dim S = int (DIM('n))"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   300
  shows "interior S = {}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   301
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   302
  have "aff_dim(interior S) \<le> aff_dim S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   303
    using interior_subset aff_dim_subset[of "interior S" S] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   304
  then show ?thesis
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
   305
    using aff_dim_open[of "interior S"] aff_dim_le_DIM[of S] assms by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   306
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   307
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
   308
corollary empty_interior_lowdim:
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
   309
  fixes S :: "'n::euclidean_space set"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
   310
  shows "dim S < DIM ('n) \<Longrightarrow> interior S = {}"
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
   311
by (metis low_dim_interior affine_hull_UNIV dim_affine_hull less_not_refl dim_UNIV)
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
   312
63016
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
   313
corollary aff_dim_nonempty_interior:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
   314
  fixes S :: "'a::euclidean_space set"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
   315
  shows "interior S \<noteq> {} \<Longrightarrow> aff_dim S = DIM('a)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
   316
by (metis low_dim_interior)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
   317
63881
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
   318
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
   319
subsection \<open>Relative interior of a set\<close>
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   320
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70097
diff changeset
   321
definition\<^marker>\<open>tag important\<close> "rel_interior S =
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69710
diff changeset
   322
  {x. \<exists>T. openin (top_of_set (affine hull S)) T \<and> x \<in> T \<and> T \<subseteq> S}"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   323
64287
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   324
lemma rel_interior_mono:
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   325
   "\<lbrakk>S \<subseteq> T; affine hull S = affine hull T\<rbrakk>
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   326
   \<Longrightarrow> (rel_interior S) \<subseteq> (rel_interior T)"
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   327
  by (auto simp: rel_interior_def)
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   328
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   329
lemma rel_interior_maximal:
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69710
diff changeset
   330
   "\<lbrakk>T \<subseteq> S; openin(top_of_set (affine hull S)) T\<rbrakk> \<Longrightarrow> T \<subseteq> (rel_interior S)"
64287
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   331
  by (auto simp: rel_interior_def)
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   332
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   333
lemma rel_interior:
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   334
  "rel_interior S = {x \<in> S. \<exists>T. open T \<and> x \<in> T \<and> T \<inter> affine hull S \<subseteq> S}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   335
  unfolding rel_interior_def[of S] openin_open[of "affine hull S"]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   336
  apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   337
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   338
  fix x T
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   339
  assume *: "x \<in> S" "open T" "x \<in> T" "T \<inter> affine hull S \<subseteq> S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   340
  then have **: "x \<in> T \<inter> affine hull S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   341
    using hull_inc by auto
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
   342
  show "\<exists>Tb. (\<exists>Ta. open Ta \<and> Tb = affine hull S \<inter> Ta) \<and> x \<in> Tb \<and> Tb \<subseteq> S"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
   343
    apply (rule_tac x = "T \<inter> (affine hull S)" in exI)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   344
    using * **
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   345
    apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   346
    done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   347
qed
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   348
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   349
lemma mem_rel_interior: "x \<in> rel_interior S \<longleftrightarrow> (\<exists>T. open T \<and> x \<in> T \<inter> S \<and> T \<inter> affine hull S \<subseteq> S)"
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
   350
  by (auto simp: rel_interior)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   351
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   352
lemma mem_rel_interior_ball:
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   353
  "x \<in> rel_interior S \<longleftrightarrow> x \<in> S \<and> (\<exists>e. e > 0 \<and> ball x e \<inter> affine hull S \<subseteq> S)"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   354
  apply (simp add: rel_interior, safe)
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
   355
  apply (force simp: open_contains_ball)
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
   356
  apply (rule_tac x = "ball x e" in exI, simp)
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   357
  done
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   358
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
   359
lemma rel_interior_ball:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   360
  "rel_interior S = {x \<in> S. \<exists>e. e > 0 \<and> ball x e \<inter> affine hull S \<subseteq> S}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   361
  using mem_rel_interior_ball [of _ S] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   362
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   363
lemma mem_rel_interior_cball:
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   364
  "x \<in> rel_interior S \<longleftrightarrow> x \<in> S \<and> (\<exists>e. e > 0 \<and> cball x e \<inter> affine hull S \<subseteq> S)"
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
   365
  apply (simp add: rel_interior, safe)
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
   366
  apply (force simp: open_contains_cball)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   367
  apply (rule_tac x = "ball x e" in exI)
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
   368
  apply (simp add: subset_trans [OF ball_subset_cball], auto)
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   369
  done
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   370
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   371
lemma rel_interior_cball:
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   372
  "rel_interior S = {x \<in> S. \<exists>e. e > 0 \<and> cball x e \<inter> affine hull S \<subseteq> S}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   373
  using mem_rel_interior_cball [of _ S] by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   374
60303
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60176
diff changeset
   375
lemma rel_interior_empty [simp]: "rel_interior {} = {}"
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
   376
   by (auto simp: rel_interior_def)
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   377
60303
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60176
diff changeset
   378
lemma affine_hull_sing [simp]: "affine hull {a :: 'n::euclidean_space} = {a}"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   379
  by (metis affine_hull_eq affine_sing)
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   380
63114
27afe7af7379 Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 63092
diff changeset
   381
lemma rel_interior_sing [simp]:
27afe7af7379 Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 63092
diff changeset
   382
    fixes a :: "'n::euclidean_space"  shows "rel_interior {a} = {a}"
27afe7af7379 Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 63092
diff changeset
   383
  apply (auto simp: rel_interior_ball)
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
   384
  apply (rule_tac x=1 in exI, force)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   385
  done
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   386
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   387
lemma subset_rel_interior:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   388
  fixes S T :: "'n::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   389
  assumes "S \<subseteq> T"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   390
    and "affine hull S = affine hull T"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   391
  shows "rel_interior S \<subseteq> rel_interior T"
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
   392
  using assms by (auto simp: rel_interior_def)
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
   393
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   394
lemma rel_interior_subset: "rel_interior S \<subseteq> S"
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
   395
  by (auto simp: rel_interior_def)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   396
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   397
lemma rel_interior_subset_closure: "rel_interior S \<subseteq> closure S"
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
   398
  using rel_interior_subset by (auto simp: closure_def)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   399
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   400
lemma interior_subset_rel_interior: "interior S \<subseteq> rel_interior S"
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
   401
  by (auto simp: rel_interior interior_def)
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   402
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   403
lemma interior_rel_interior:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   404
  fixes S :: "'n::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   405
  assumes "aff_dim S = int(DIM('n))"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   406
  shows "rel_interior S = interior S"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   407
proof -
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   408
  have "affine hull S = UNIV"
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
   409
    using assms affine_hull_UNIV[of S] by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   410
  then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   411
    unfolding rel_interior interior_def by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   412
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   413
60303
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60176
diff changeset
   414
lemma rel_interior_interior:
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60176
diff changeset
   415
  fixes S :: "'n::euclidean_space set"
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60176
diff changeset
   416
  assumes "affine hull S = UNIV"
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60176
diff changeset
   417
  shows "rel_interior S = interior S"
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60176
diff changeset
   418
  using assms unfolding rel_interior interior_def by auto
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60176
diff changeset
   419
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   420
lemma rel_interior_open:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   421
  fixes S :: "'n::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   422
  assumes "open S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   423
  shows "rel_interior S = S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   424
  by (metis assms interior_eq interior_subset_rel_interior rel_interior_subset set_eq_subset)
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   425
60800
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
   426
lemma interior_ball [simp]: "interior (ball x e) = ball x e"
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
   427
  by (simp add: interior_open)
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
   428
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   429
lemma interior_rel_interior_gen:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   430
  fixes S :: "'n::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   431
  shows "interior S = (if aff_dim S = int(DIM('n)) then rel_interior S else {})"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   432
  by (metis interior_rel_interior low_dim_interior)
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   433
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
   434
lemma rel_interior_nonempty_interior:
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
   435
  fixes S :: "'n::euclidean_space set"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
   436
  shows "interior S \<noteq> {} \<Longrightarrow> rel_interior S = interior S"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
   437
by (metis interior_rel_interior_gen)
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
   438
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
   439
lemma affine_hull_nonempty_interior:
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
   440
  fixes S :: "'n::euclidean_space set"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
   441
  shows "interior S \<noteq> {} \<Longrightarrow> affine hull S = UNIV"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
   442
by (metis affine_hull_UNIV interior_rel_interior_gen)
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
   443
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
   444
lemma rel_interior_affine_hull [simp]:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   445
  fixes S :: "'n::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   446
  shows "rel_interior (affine hull S) = affine hull S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   447
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   448
  have *: "rel_interior (affine hull S) \<subseteq> affine hull S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   449
    using rel_interior_subset by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   450
  {
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   451
    fix x
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   452
    assume x: "x \<in> affine hull S"
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
   453
    define e :: real where "e = 1"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   454
    then have "e > 0" "ball x e \<inter> affine hull (affine hull S) \<subseteq> affine hull S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   455
      using hull_hull[of _ S] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   456
    then have "x \<in> rel_interior (affine hull S)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   457
      using x rel_interior_ball[of "affine hull S"] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   458
  }
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   459
  then show ?thesis using * by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   460
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   461
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
   462
lemma rel_interior_UNIV [simp]: "rel_interior (UNIV :: ('n::euclidean_space) set) = UNIV"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   463
  by (metis open_UNIV rel_interior_open)
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   464
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   465
lemma rel_interior_convex_shrink:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   466
  fixes S :: "'a::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   467
  assumes "convex S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   468
    and "c \<in> rel_interior S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   469
    and "x \<in> S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   470
    and "0 < e"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   471
    and "e \<le> 1"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   472
  shows "x - e *\<^sub>R (x - c) \<in> rel_interior S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   473
proof -
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
   474
  obtain d where "d > 0" and d: "ball c d \<inter> affine hull S \<subseteq> S"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   475
    using assms(2) unfolding  mem_rel_interior_ball by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   476
  {
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   477
    fix y
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   478
    assume as: "dist (x - e *\<^sub>R (x - c)) y < e * d" "y \<in> affine hull S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   479
    have *: "y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x"
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
   480
      using \<open>e > 0\<close> by (auto simp: scaleR_left_diff_distrib scaleR_right_diff_distrib)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   481
    have "x \<in> affine hull S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   482
      using assms hull_subset[of S] by auto
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
   483
    moreover have "1 / e + - ((1 - e) / e) = 1"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
   484
      using \<open>e > 0\<close> left_diff_distrib[of "1" "(1-e)" "1/e"] by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   485
    ultimately have **: "(1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x \<in> affine hull S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   486
      using as affine_affine_hull[of S] mem_affine[of "affine hull S" y x "(1 / e)" "-((1 - e) / e)"]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   487
      by (simp add: algebra_simps)
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61880
diff changeset
   488
    have "dist c ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = \<bar>1/e\<bar> * norm (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   489
      unfolding dist_norm norm_scaleR[symmetric]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   490
      apply (rule arg_cong[where f=norm])
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
   491
      using \<open>e > 0\<close>
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
   492
      apply (auto simp: euclidean_eq_iff[where 'a='a] field_simps inner_simps)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   493
      done
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61880
diff changeset
   494
    also have "\<dots> = \<bar>1/e\<bar> * norm (x - e *\<^sub>R (x - c) - y)"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   495
      by (auto intro!:arg_cong[where f=norm] simp add: algebra_simps)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   496
    also have "\<dots> < d"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
   497
      using as[unfolded dist_norm] and \<open>e > 0\<close>
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
   498
      by (auto simp:pos_divide_less_eq[OF \<open>e > 0\<close>] mult.commute)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   499
    finally have "y \<in> S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   500
      apply (subst *)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   501
      apply (rule assms(1)[unfolded convex_alt,rule_format])
68058
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
   502
      apply (rule d[THEN subsetD])
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   503
      unfolding mem_ball
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   504
      using assms(3-5) **
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   505
      apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   506
      done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   507
  }
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   508
  then have "ball (x - e *\<^sub>R (x - c)) (e*d) \<inter> affine hull S \<subseteq> S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   509
    by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   510
  moreover have "e * d > 0"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
   511
    using \<open>e > 0\<close> \<open>d > 0\<close> by simp
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   512
  moreover have c: "c \<in> S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   513
    using assms rel_interior_subset by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   514
  moreover from c have "x - e *\<^sub>R (x - c) \<in> S"
61426
d53db136e8fd new material on path_component_sets, inside, outside, etc. And more default simprules
paulson <lp15@cam.ac.uk>
parents: 61222
diff changeset
   515
    using convexD_alt[of S x c e]
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   516
    apply (simp add: algebra_simps)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   517
    using assms
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   518
    apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   519
    done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   520
  ultimately show ?thesis
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
   521
    using mem_rel_interior_ball[of "x - e *\<^sub>R (x - c)" S] \<open>e > 0\<close> by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   522
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   523
69710
61372780515b some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents: 69619
diff changeset
   524
lemma interior_real_atLeast [simp]:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   525
  fixes a :: real
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   526
  shows "interior {a..} = {a<..}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   527
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   528
  {
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   529
    fix y
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   530
    assume "a < y"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   531
    then have "y \<in> interior {a..}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   532
      apply (simp add: mem_interior)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   533
      apply (rule_tac x="(y-a)" in exI)
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
   534
      apply (auto simp: dist_norm)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   535
      done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   536
  }
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   537
  moreover
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   538
  {
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   539
    fix y
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   540
    assume "y \<in> interior {a..}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   541
    then obtain e where e: "e > 0" "cball y e \<subseteq> {a..}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   542
      using mem_interior_cball[of y "{a..}"] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   543
    moreover from e have "y - e \<in> cball y e"
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
   544
      by (auto simp: cball_def dist_norm)
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
   545
    ultimately have "a \<le> y - e" by blast
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   546
    then have "a < y" using e by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   547
  }
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   548
  ultimately show ?thesis by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   549
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   550
61880
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   551
lemma continuous_ge_on_Ioo:
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   552
  assumes "continuous_on {c..d} g" "\<And>x. x \<in> {c<..<d} \<Longrightarrow> g x \<ge> a" "c < d" "x \<in> {c..d}"
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   553
  shows "g (x::real) \<ge> (a::real)"
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   554
proof-
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   555
  from assms(3) have "{c..d} = closure {c<..<d}" by (rule closure_greaterThanLessThan[symmetric])
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   556
  also from assms(2) have "{c<..<d} \<subseteq> (g -` {a..} \<inter> {c..d})" by auto
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   557
  hence "closure {c<..<d} \<subseteq> closure (g -` {a..} \<inter> {c..d})" by (rule closure_mono)
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   558
  also from assms(1) have "closed (g -` {a..} \<inter> {c..d})"
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   559
    by (auto simp: continuous_on_closed_vimage)
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   560
  hence "closure (g -` {a..} \<inter> {c..d}) = g -` {a..} \<inter> {c..d}" by simp
62087
44841d07ef1d revisions to limits and derivatives, plus new lemmas
paulson
parents: 61952
diff changeset
   561
  finally show ?thesis using \<open>x \<in> {c..d}\<close> by auto
44841d07ef1d revisions to limits and derivatives, plus new lemmas
paulson
parents: 61952
diff changeset
   562
qed
61880
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   563
69710
61372780515b some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents: 69619
diff changeset
   564
lemma interior_real_atMost [simp]:
61880
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   565
  fixes a :: real
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   566
  shows "interior {..a} = {..<a}"
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   567
proof -
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   568
  {
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   569
    fix y
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   570
    assume "a > y"
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   571
    then have "y \<in> interior {..a}"
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   572
      apply (simp add: mem_interior)
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   573
      apply (rule_tac x="(a-y)" in exI)
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
   574
      apply (auto simp: dist_norm)
61880
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   575
      done
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   576
  }
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   577
  moreover
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   578
  {
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   579
    fix y
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   580
    assume "y \<in> interior {..a}"
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   581
    then obtain e where e: "e > 0" "cball y e \<subseteq> {..a}"
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   582
      using mem_interior_cball[of y "{..a}"] by auto
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   583
    moreover from e have "y + e \<in> cball y e"
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
   584
      by (auto simp: cball_def dist_norm)
61880
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   585
    ultimately have "a \<ge> y + e" by auto
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   586
    then have "a > y" using e by auto
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   587
  }
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   588
  ultimately show ?thesis by auto
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   589
qed
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   590
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64394
diff changeset
   591
lemma interior_atLeastAtMost_real [simp]: "interior {a..b} = {a<..<b :: real}"
61880
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   592
proof-
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   593
  have "{a..b} = {a..} \<inter> {..b}" by auto
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
   594
  also have "interior \<dots> = {a<..} \<inter> {..<b}"
69710
61372780515b some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents: 69619
diff changeset
   595
    by (simp add: interior_real_atLeast interior_real_atMost)
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
   596
  also have "\<dots> = {a<..<b}" by auto
61880
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   597
  finally show ?thesis .
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   598
qed
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   599
66793
deabce3ccf1f new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents: 66641
diff changeset
   600
lemma interior_atLeastLessThan [simp]:
deabce3ccf1f new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents: 66641
diff changeset
   601
  fixes a::real shows "interior {a..<b} = {a<..<b}"
69710
61372780515b some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents: 69619
diff changeset
   602
  by (metis atLeastLessThan_def greaterThanLessThan_def interior_atLeastAtMost_real interior_Int interior_interior interior_real_atLeast)
66793
deabce3ccf1f new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents: 66641
diff changeset
   603
deabce3ccf1f new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents: 66641
diff changeset
   604
lemma interior_lessThanAtMost [simp]:
deabce3ccf1f new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents: 66641
diff changeset
   605
  fixes a::real shows "interior {a<..b} = {a<..<b}"
deabce3ccf1f new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents: 66641
diff changeset
   606
  by (metis atLeastAtMost_def greaterThanAtMost_def interior_atLeastAtMost_real interior_Int
69710
61372780515b some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents: 69619
diff changeset
   607
            interior_interior interior_real_atLeast)
66793
deabce3ccf1f new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents: 66641
diff changeset
   608
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64394
diff changeset
   609
lemma interior_greaterThanLessThan_real [simp]: "interior {a<..<b} = {a<..<b :: real}"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64394
diff changeset
   610
  by (metis interior_atLeastAtMost_real interior_interior)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64394
diff changeset
   611
69710
61372780515b some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents: 69619
diff changeset
   612
lemma frontier_real_atMost [simp]:
61880
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   613
  fixes a :: real
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   614
  shows "frontier {..a} = {a}"
69710
61372780515b some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents: 69619
diff changeset
   615
  unfolding frontier_def by (auto simp: interior_real_atMost)
61372780515b some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents: 69619
diff changeset
   616
61372780515b some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents: 69619
diff changeset
   617
lemma frontier_real_atLeast [simp]: "frontier {a..} = {a::real}"
61372780515b some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents: 69619
diff changeset
   618
  by (auto simp: frontier_def)
61372780515b some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents: 69619
diff changeset
   619
61372780515b some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents: 69619
diff changeset
   620
lemma frontier_real_greaterThan [simp]: "frontier {a<..} = {a::real}"
61372780515b some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents: 69619
diff changeset
   621
  by (auto simp: interior_open frontier_def)
61372780515b some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents: 69619
diff changeset
   622
61372780515b some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents: 69619
diff changeset
   623
lemma frontier_real_lessThan [simp]: "frontier {..<a} = {a::real}"
61372780515b some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents: 69619
diff changeset
   624
  by (auto simp: interior_open frontier_def)
61880
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
   625
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64394
diff changeset
   626
lemma rel_interior_real_box [simp]:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   627
  fixes a b :: real
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   628
  assumes "a < b"
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
   629
  shows "rel_interior {a .. b} = {a <..< b}"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   630
proof -
54775
2d3df8633dad prefer box over greaterThanLessThan on euclidean_space
immler
parents: 54465
diff changeset
   631
  have "box a b \<noteq> {}"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   632
    using assms
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   633
    unfolding set_eq_iff
56189
c4daa97ac57a removed dependencies on theory Ordered_Euclidean_Space
immler
parents: 56188
diff changeset
   634
    by (auto intro!: exI[of _ "(a + b) / 2"] simp: box_def)
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   635
  then show ?thesis
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
   636
    using interior_rel_interior_gen[of "cbox a b", symmetric]
62390
842917225d56 more canonical names
nipkow
parents: 62131
diff changeset
   637
    by (simp split: if_split_asm del: box_real add: box_real[symmetric] interior_cbox)
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   638
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   639
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64394
diff changeset
   640
lemma rel_interior_real_semiline [simp]:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   641
  fixes a :: real
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   642
  shows "rel_interior {a..} = {a<..}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   643
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   644
  have *: "{a<..} \<noteq> {}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   645
    unfolding set_eq_iff by (auto intro!: exI[of _ "a + 1"])
69710
61372780515b some renamings and a bit of new material
paulson <lp15@cam.ac.uk>
parents: 69619
diff changeset
   646
  then show ?thesis using interior_real_atLeast interior_rel_interior_gen[of "{a..}"]
62390
842917225d56 more canonical names
nipkow
parents: 62131
diff changeset
   647
    by (auto split: if_split_asm)
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   648
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   649
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
   650
subsubsection \<open>Relative open sets\<close>
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   651
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70097
diff changeset
   652
definition\<^marker>\<open>tag important\<close> "rel_open S \<longleftrightarrow> rel_interior S = S"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   653
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69710
diff changeset
   654
lemma rel_open: "rel_open S \<longleftrightarrow> openin (top_of_set (affine hull S)) S"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   655
  unfolding rel_open_def rel_interior_def
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   656
  apply auto
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69710
diff changeset
   657
  using openin_subopen[of "top_of_set (affine hull S)" S]
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   658
  apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   659
  done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   660
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69710
diff changeset
   661
lemma openin_rel_interior: "openin (top_of_set (affine hull S)) (rel_interior S)"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   662
  apply (simp add: rel_interior_def)
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
   663
  apply (subst openin_subopen, blast)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   664
  done
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   665
63469
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   666
lemma openin_set_rel_interior:
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69710
diff changeset
   667
   "openin (top_of_set S) (rel_interior S)"
63469
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   668
by (rule openin_subset_trans [OF openin_rel_interior rel_interior_subset hull_subset])
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   669
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
   670
lemma affine_rel_open:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   671
  fixes S :: "'n::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   672
  assumes "affine S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   673
  shows "rel_open S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   674
  unfolding rel_open_def
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
   675
  using assms rel_interior_affine_hull[of S] affine_hull_eq[of S]
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   676
  by metis
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   677
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
   678
lemma affine_closed:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   679
  fixes S :: "'n::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   680
  assumes "affine S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   681
  shows "closed S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   682
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   683
  {
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   684
    assume "S \<noteq> {}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   685
    then obtain L where L: "subspace L" "affine_parallel S L"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   686
      using assms affine_parallel_subspace[of S] by auto
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67155
diff changeset
   687
    then obtain a where a: "S = ((+) a ` L)"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   688
      using affine_parallel_def[of L S] affine_parallel_commut by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   689
    from L have "closed L" using closed_subspace by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   690
    then have "closed S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   691
      using closed_translation a by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   692
  }
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   693
  then show ?thesis by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   694
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   695
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   696
lemma closure_affine_hull:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   697
  fixes S :: "'n::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   698
  shows "closure S \<subseteq> affine hull S"
44524
04ad69081646 generalize some lemmas
huffman
parents: 44523
diff changeset
   699
  by (intro closure_minimal hull_subset affine_closed affine_affine_hull)
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   700
62948
7700f467892b lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 62843
diff changeset
   701
lemma closure_same_affine_hull [simp]:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   702
  fixes S :: "'n::euclidean_space set"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   703
  shows "affine hull (closure S) = affine hull S"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   704
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   705
  have "affine hull (closure S) \<subseteq> affine hull S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   706
    using hull_mono[of "closure S" "affine hull S" "affine"]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   707
      closure_affine_hull[of S] hull_hull[of "affine" S]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   708
    by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   709
  moreover have "affine hull (closure S) \<supseteq> affine hull S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   710
    using hull_mono[of "S" "closure S" "affine"] closure_subset by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   711
  ultimately show ?thesis by auto
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
   712
qed
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
   713
63114
27afe7af7379 Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 63092
diff changeset
   714
lemma closure_aff_dim [simp]:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   715
  fixes S :: "'n::euclidean_space set"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   716
  shows "aff_dim (closure S) = aff_dim S"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   717
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   718
  have "aff_dim S \<le> aff_dim (closure S)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   719
    using aff_dim_subset closure_subset by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   720
  moreover have "aff_dim (closure S) \<le> aff_dim (affine hull S)"
63075
60a42a4166af lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents: 63072
diff changeset
   721
    using aff_dim_subset closure_affine_hull by blast
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   722
  moreover have "aff_dim (affine hull S) = aff_dim S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   723
    using aff_dim_affine_hull by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   724
  ultimately show ?thesis by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   725
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   726
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   727
lemma rel_interior_closure_convex_shrink:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   728
  fixes S :: "_::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   729
  assumes "convex S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   730
    and "c \<in> rel_interior S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   731
    and "x \<in> closure S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   732
    and "e > 0"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   733
    and "e \<le> 1"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   734
  shows "x - e *\<^sub>R (x - c) \<in> rel_interior S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   735
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   736
  obtain d where "d > 0" and d: "ball c d \<inter> affine hull S \<subseteq> S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   737
    using assms(2) unfolding mem_rel_interior_ball by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   738
  have "\<exists>y \<in> S. norm (y - x) * (1 - e) < e * d"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   739
  proof (cases "x \<in> S")
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   740
    case True
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
   741
    then show ?thesis using \<open>e > 0\<close> \<open>d > 0\<close>
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
   742
      apply (rule_tac bexI[where x=x], auto)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   743
      done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   744
  next
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   745
    case False
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   746
    then have x: "x islimpt S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   747
      using assms(3)[unfolded closure_def] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   748
    show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   749
    proof (cases "e = 1")
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   750
      case True
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   751
      obtain y where "y \<in> S" "y \<noteq> x" "dist y x < 1"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   752
        using x[unfolded islimpt_approachable,THEN spec[where x=1]] by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   753
      then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   754
        apply (rule_tac x=y in bexI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   755
        unfolding True
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
   756
        using \<open>d > 0\<close>
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   757
        apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   758
        done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   759
    next
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   760
      case False
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   761
      then have "0 < e * d / (1 - e)" and *: "1 - e > 0"
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
   762
        using \<open>e \<le> 1\<close> \<open>e > 0\<close> \<open>d > 0\<close> by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   763
      then obtain y where "y \<in> S" "y \<noteq> x" "dist y x < e * d / (1 - e)"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   764
        using x[unfolded islimpt_approachable,THEN spec[where x="e*d / (1 - e)"]] by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   765
      then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   766
        apply (rule_tac x=y in bexI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   767
        unfolding dist_norm
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   768
        using pos_less_divide_eq[OF *]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   769
        apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   770
        done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   771
    qed
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   772
  qed
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   773
  then obtain y where "y \<in> S" and y: "norm (y - x) * (1 - e) < e * d"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   774
    by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
   775
  define z where "z = c + ((1 - e) / e) *\<^sub>R (x - y)"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   776
  have *: "x - e *\<^sub>R (x - c) = y - e *\<^sub>R (y - z)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
   777
    unfolding z_def using \<open>e > 0\<close>
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
   778
    by (auto simp: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   779
  have zball: "z \<in> ball c d"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   780
    using mem_ball z_def dist_norm[of c]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   781
    using y and assms(4,5)
70802
160eaf566bcb formally augmented corresponding rules for field_simps
haftmann
parents: 70136
diff changeset
   782
    by (simp add: norm_minus_commute) (simp add: field_simps)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   783
  have "x \<in> affine hull S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   784
    using closure_affine_hull assms by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   785
  moreover have "y \<in> affine hull S"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
   786
    using \<open>y \<in> S\<close> hull_subset[of S] by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   787
  moreover have "c \<in> affine hull S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   788
    using assms rel_interior_subset hull_subset[of S] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   789
  ultimately have "z \<in> affine hull S"
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
   790
    using z_def affine_affine_hull[of S]
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   791
      mem_affine_3_minus [of "affine hull S" c x y "(1 - e) / e"]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   792
      assms
70802
160eaf566bcb formally augmented corresponding rules for field_simps
haftmann
parents: 70136
diff changeset
   793
    by simp
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   794
  then have "z \<in> S" using d zball by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   795
  obtain d1 where "d1 > 0" and d1: "ball z d1 \<le> ball c d"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   796
    using zball open_ball[of c d] openE[of "ball c d" z] by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   797
  then have "ball z d1 \<inter> affine hull S \<subseteq> ball c d \<inter> affine hull S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   798
    by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   799
  then have "ball z d1 \<inter> affine hull S \<subseteq> S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   800
    using d by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   801
  then have "z \<in> rel_interior S"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
   802
    using mem_rel_interior_ball using \<open>d1 > 0\<close> \<open>z \<in> S\<close> by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   803
  then have "y - e *\<^sub>R (y - z) \<in> rel_interior S"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
   804
    using rel_interior_convex_shrink[of S z y e] assms \<open>y \<in> S\<close> by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   805
  then show ?thesis using * by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   806
qed
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   807
62620
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
   808
lemma rel_interior_eq:
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69710
diff changeset
   809
   "rel_interior s = s \<longleftrightarrow> openin(top_of_set (affine hull s)) s"
62620
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
   810
using rel_open rel_open_def by blast
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
   811
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
   812
lemma rel_interior_openin:
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69710
diff changeset
   813
   "openin(top_of_set (affine hull s)) s \<Longrightarrow> rel_interior s = s"
62620
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
   814
by (simp add: rel_interior_eq)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
   815
63469
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   816
lemma rel_interior_affine:
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   817
  fixes S :: "'n::euclidean_space set"
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   818
  shows  "affine S \<Longrightarrow> rel_interior S = S"
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   819
using affine_rel_open rel_open_def by auto
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   820
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   821
lemma rel_interior_eq_closure:
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   822
  fixes S :: "'n::euclidean_space set"
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   823
  shows "rel_interior S = closure S \<longleftrightarrow> affine S"
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   824
proof (cases "S = {}")
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   825
  case True
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   826
 then show ?thesis
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   827
    by auto
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   828
next
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   829
  case False show ?thesis
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   830
  proof
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   831
    assume eq: "rel_interior S = closure S"
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   832
    have "S = {} \<or> S = affine hull S"
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   833
      apply (rule connected_clopen [THEN iffD1, rule_format])
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   834
       apply (simp add: affine_imp_convex convex_connected)
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   835
      apply (rule conjI)
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   836
       apply (metis eq closure_subset openin_rel_interior rel_interior_subset subset_antisym)
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   837
      apply (metis closed_subset closure_subset_eq eq hull_subset rel_interior_subset)
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   838
      done
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   839
    with False have "affine hull S = S"
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   840
      by auto
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   841
    then show "affine S"
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   842
      by (metis affine_hull_eq)
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   843
  next
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   844
    assume "affine S"
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   845
    then show "rel_interior S = closure S"
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   846
      by (simp add: rel_interior_affine affine_closed)
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   847
  qed
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   848
qed
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   849
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   850
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70097
diff changeset
   851
subsubsection\<^marker>\<open>tag unimportant\<close>\<open>Relative interior preserves under linear transformations\<close>
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   852
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   853
lemma rel_interior_translation_aux:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   854
  fixes a :: "'n::euclidean_space"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   855
  shows "((\<lambda>x. a + x) ` rel_interior S) \<subseteq> rel_interior ((\<lambda>x. a + x) ` S)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   856
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   857
  {
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   858
    fix x
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   859
    assume x: "x \<in> rel_interior S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   860
    then obtain T where "open T" "x \<in> T \<inter> S" "T \<inter> affine hull S \<subseteq> S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   861
      using mem_rel_interior[of x S] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   862
    then have "open ((\<lambda>x. a + x) ` T)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   863
      and "a + x \<in> ((\<lambda>x. a + x) ` T) \<inter> ((\<lambda>x. a + x) ` S)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   864
      and "((\<lambda>x. a + x) ` T) \<inter> affine hull ((\<lambda>x. a + x) ` S) \<subseteq> (\<lambda>x. a + x) ` S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   865
      using affine_hull_translation[of a S] open_translation[of T a] x by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   866
    then have "a + x \<in> rel_interior ((\<lambda>x. a + x) ` S)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   867
      using mem_rel_interior[of "a+x" "((\<lambda>x. a + x) ` S)"] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   868
  }
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   869
  then show ?thesis by auto
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
   870
qed
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   871
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   872
lemma rel_interior_translation:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   873
  fixes a :: "'n::euclidean_space"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   874
  shows "rel_interior ((\<lambda>x. a + x) ` S) = (\<lambda>x. a + x) ` rel_interior S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   875
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   876
  have "(\<lambda>x. (-a) + x) ` rel_interior ((\<lambda>x. a + x) ` S) \<subseteq> rel_interior S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   877
    using rel_interior_translation_aux[of "-a" "(\<lambda>x. a + x) ` S"]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   878
      translation_assoc[of "-a" "a"]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   879
    by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   880
  then have "((\<lambda>x. a + x) ` rel_interior S) \<supseteq> rel_interior ((\<lambda>x. a + x) ` S)"
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67155
diff changeset
   881
    using translation_inverse_subset[of a "rel_interior ((+) a ` S)" "rel_interior S"]
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   882
    by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   883
  then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   884
    using rel_interior_translation_aux[of a S] by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   885
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   886
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   887
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   888
lemma affine_hull_linear_image:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   889
  assumes "bounded_linear f"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   890
  shows "f ` (affine hull s) = affine hull f ` s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   891
proof -
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   892
  interpret f: bounded_linear f by fact
68058
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
   893
  have "affine {x. f x \<in> affine hull f ` s}"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   894
    unfolding affine_def
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
   895
    by (auto simp: f.scaleR f.add affine_affine_hull[unfolded affine_def, rule_format])
68058
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
   896
  moreover have "affine {x. x \<in> f ` (affine hull s)}"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   897
    using affine_affine_hull[unfolded affine_def, of s]
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
   898
    unfolding affine_def by (auto simp: f.scaleR [symmetric] f.add [symmetric])
68058
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
   899
  ultimately show ?thesis
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
   900
    by (auto simp: hull_inc elim!: hull_induct)
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
   901
qed 
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   902
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   903
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   904
lemma rel_interior_injective_on_span_linear_image:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   905
  fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   906
    and S :: "'m::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   907
  assumes "bounded_linear f"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   908
    and "inj_on f (span S)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   909
  shows "rel_interior (f ` S) = f ` (rel_interior S)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   910
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   911
  {
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   912
    fix z
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   913
    assume z: "z \<in> rel_interior (f ` S)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   914
    then have "z \<in> f ` S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   915
      using rel_interior_subset[of "f ` S"] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   916
    then obtain x where x: "x \<in> S" "f x = z" by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   917
    obtain e2 where e2: "e2 > 0" "cball z e2 \<inter> affine hull (f ` S) \<subseteq> (f ` S)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   918
      using z rel_interior_cball[of "f ` S"] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   919
    obtain K where K: "K > 0" "\<And>x. norm (f x) \<le> norm x * K"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   920
     using assms Real_Vector_Spaces.bounded_linear.pos_bounded[of f] by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
   921
    define e1 where "e1 = 1 / K"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   922
    then have e1: "e1 > 0" "\<And>x. e1 * norm (f x) \<le> norm x"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   923
      using K pos_le_divide_eq[of e1] by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
   924
    define e where "e = e1 * e2"
56544
b60d5d119489 made mult_pos_pos a simp rule
nipkow
parents: 56541
diff changeset
   925
    then have "e > 0" using e1 e2 by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   926
    {
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   927
      fix y
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   928
      assume y: "y \<in> cball x e \<inter> affine hull S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   929
      then have h1: "f y \<in> affine hull (f ` S)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   930
        using affine_hull_linear_image[of f S] assms by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   931
      from y have "norm (x-y) \<le> e1 * e2"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   932
        using cball_def[of x e] dist_norm[of x y] e_def by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   933
      moreover have "f x - f y = f (x - y)"
63469
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   934
        using assms linear_diff[of f x y] linear_conv_bounded_linear[of f] by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   935
      moreover have "e1 * norm (f (x-y)) \<le> norm (x - y)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   936
        using e1 by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   937
      ultimately have "e1 * norm ((f x)-(f y)) \<le> e1 * e2"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   938
        by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   939
      then have "f y \<in> cball z e2"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   940
        using cball_def[of "f x" e2] dist_norm[of "f x" "f y"] e1 x by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   941
      then have "f y \<in> f ` S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   942
        using y e2 h1 by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   943
      then have "y \<in> S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   944
        using assms y hull_subset[of S] affine_hull_subset_span
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
   945
          inj_on_image_mem_iff [OF \<open>inj_on f (span S)\<close>]
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   946
        by (metis Int_iff span_superset subsetCE)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   947
    }
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   948
    then have "z \<in> f ` (rel_interior S)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
   949
      using mem_rel_interior_cball[of x S] \<open>e > 0\<close> x by auto
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
   950
  }
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   951
  moreover
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   952
  {
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   953
    fix x
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   954
    assume x: "x \<in> rel_interior S"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
   955
    then obtain e2 where e2: "e2 > 0" "cball x e2 \<inter> affine hull S \<subseteq> S"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   956
      using rel_interior_cball[of S] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   957
    have "x \<in> S" using x rel_interior_subset by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   958
    then have *: "f x \<in> f ` S" by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   959
    have "\<forall>x\<in>span S. f x = 0 \<longrightarrow> x = 0"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   960
      using assms subspace_span linear_conv_bounded_linear[of f]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   961
        linear_injective_on_subspace_0[of f "span S"]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   962
      by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   963
    then obtain e1 where e1: "e1 > 0" "\<forall>x \<in> span S. e1 * norm x \<le> norm (f x)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   964
      using assms injective_imp_isometric[of "span S" f]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   965
        subspace_span[of S] closed_subspace[of "span S"]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   966
      by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
   967
    define e where "e = e1 * e2"
56544
b60d5d119489 made mult_pos_pos a simp rule
nipkow
parents: 56541
diff changeset
   968
    hence "e > 0" using e1 e2 by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   969
    {
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   970
      fix y
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   971
      assume y: "y \<in> cball (f x) e \<inter> affine hull (f ` S)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   972
      then have "y \<in> f ` (affine hull S)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   973
        using affine_hull_linear_image[of f S] assms by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   974
      then obtain xy where xy: "xy \<in> affine hull S" "f xy = y" by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   975
      with y have "norm (f x - f xy) \<le> e1 * e2"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   976
        using cball_def[of "f x" e] dist_norm[of "f x" y] e_def by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   977
      moreover have "f x - f xy = f (x - xy)"
63469
b6900858dcb9 lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents: 63332
diff changeset
   978
        using assms linear_diff[of f x xy] linear_conv_bounded_linear[of f] by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   979
      moreover have *: "x - xy \<in> span S"
63114
27afe7af7379 Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 63092
diff changeset
   980
        using subspace_diff[of "span S" x xy] subspace_span \<open>x \<in> S\<close> xy
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
   981
          affine_hull_subset_span[of S] span_superset
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   982
        by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   983
      moreover from * have "e1 * norm (x - xy) \<le> norm (f (x - xy))"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   984
        using e1 by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   985
      ultimately have "e1 * norm (x - xy) \<le> e1 * e2"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   986
        by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   987
      then have "xy \<in> cball x e2"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   988
        using cball_def[of x e2] dist_norm[of x xy] e1 by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   989
      then have "y \<in> f ` S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   990
        using xy e2 by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   991
    }
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   992
    then have "f x \<in> rel_interior (f ` S)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
   993
      using mem_rel_interior_cball[of "(f x)" "(f ` S)"] * \<open>e > 0\<close> by auto
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
   994
  }
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   995
  ultimately show ?thesis by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   996
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   997
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   998
lemma rel_interior_injective_linear_image:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   999
  fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1000
  assumes "bounded_linear f"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1001
    and "inj f"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1002
  shows "rel_interior (f ` S) = f ` (rel_interior S)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1003
  using assms rel_interior_injective_on_span_linear_image[of f S]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1004
    subset_inj_on[of f "UNIV" "span S"]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1005
  by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1006
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1007
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70097
diff changeset
  1008
subsection\<^marker>\<open>tag unimportant\<close> \<open>Openness and compactness are preserved by convex hull operation\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1009
34964
4e8be3c04d37 Replaced vec1 and dest_vec1 by abbreviation.
hoelzl
parents: 34915
diff changeset
  1010
lemma open_convex_hull[intro]:
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1011
  fixes S :: "'a::real_normed_vector set"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1012
  assumes "open S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1013
  shows "open (convex hull S)"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1014
proof (clarsimp simp: open_contains_cball convex_hull_explicit)
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1015
  fix T and u :: "'a\<Rightarrow>real"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1016
  assume obt: "finite T" "T\<subseteq>S" "\<forall>x\<in>T. 0 \<le> u x" "sum u T = 1" 
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1017
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1018
  from assms[unfolded open_contains_cball] obtain b
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1019
    where b: "\<And>x. x\<in>S \<Longrightarrow> 0 < b x \<and> cball x (b x) \<subseteq> S" by metis
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1020
  have "b ` T \<noteq> {}"
56889
48a745e1bde7 avoid the Complex constructor, use the more natural Re/Im view; moved csqrt to Complex.
hoelzl
parents: 56571
diff changeset
  1021
    using obt by auto
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1022
  define i where "i = b ` T"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1023
  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)"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1024
  let ?a = "\<Sum>v\<in>T. u v *\<^sub>R v"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1025
  show "\<exists>e > 0. cball ?a e \<subseteq> {y. ?\<Phi> y}"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1026
  proof (intro exI subsetI conjI)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1027
    show "0 < Min i"
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1028
      unfolding i_def and Min_gr_iff[OF finite_imageI[OF obt(1)] \<open>b ` T\<noteq>{}\<close>]
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1029
      using b \<open>T\<subseteq>S\<close> by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1030
  next
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1031
    fix y
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1032
    assume "y \<in> cball ?a (Min i)"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1033
    then have y: "norm (?a - y) \<le> Min i"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1034
      unfolding dist_norm[symmetric] by auto
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1035
    { fix x
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1036
      assume "x \<in> T"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1037
      then have "Min i \<le> b x"
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1038
        by (simp add: i_def obt(1))
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1039
      then have "x + (y - ?a) \<in> cball x (b x)"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1040
        using y unfolding mem_cball dist_norm by auto
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1041
      moreover have "x \<in> S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1042
        using \<open>x\<in>T\<close> \<open>T\<subseteq>S\<close> by auto
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1043
      ultimately have "x + (y - ?a) \<in> S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1044
        using y b by blast
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1045
    }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1046
    moreover
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1047
    have *: "inj_on (\<lambda>v. v + (y - ?a)) T"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1048
      unfolding inj_on_def by auto
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1049
    have "(\<Sum>v\<in>(\<lambda>v. v + (y - ?a)) ` T. u (v - (y - ?a)) *\<^sub>R v) = y"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1050
      unfolding sum.reindex[OF *] o_def using obt(4)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
  1051
      by (simp add: sum.distrib sum_subtractf scaleR_left.sum[symmetric] scaleR_right_distrib)
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1052
    ultimately show "y \<in> {y. ?\<Phi> y}"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1053
    proof (intro CollectI exI conjI)
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1054
      show "finite ((\<lambda>v. v + (y - ?a)) ` T)"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1055
        by (simp add: obt(1))
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1056
      show "sum (\<lambda>v. u (v - (y - ?a))) ((\<lambda>v. v + (y - ?a)) ` T) = 1"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1057
        unfolding sum.reindex[OF *] o_def using obt(4) by auto
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1058
    qed (use obt(1, 3) in auto)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1059
  qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1060
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1061
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1062
lemma compact_convex_combinations:
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1063
  fixes S T :: "'a::real_normed_vector set"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1064
  assumes "compact S" "compact T"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1065
  shows "compact { (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> T}"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1066
proof -
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1067
  let ?X = "{0..1} \<times> S \<times> T"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1068
  let ?h = "(\<lambda>z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))"
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1069
  have *: "{ (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> T} = ?h ` ?X"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1070
    by force
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  1071
  have "continuous_on ?X (\<lambda>z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1072
    unfolding continuous_on by (rule ballI) (intro tendsto_intros)
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1073
  with assms show ?thesis
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1074
    by (simp add: * compact_Times compact_continuous_image)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1075
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1076
44525
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  1077
lemma finite_imp_compact_convex_hull:
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1078
  fixes S :: "'a::real_normed_vector set"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1079
  assumes "finite S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1080
  shows "compact (convex hull S)"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1081
proof (cases "S = {}")
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1082
  case True
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1083
  then show ?thesis by simp
44525
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  1084
next
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1085
  case False
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1086
  with assms show ?thesis
44525
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  1087
  proof (induct rule: finite_ne_induct)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1088
    case (singleton x)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1089
    show ?case by simp
44525
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  1090
  next
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  1091
    case (insert x A)
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  1092
    let ?f = "\<lambda>(u, y::'a). u *\<^sub>R x + (1 - u) *\<^sub>R y"
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  1093
    let ?T = "{0..1::real} \<times> (convex hull A)"
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  1094
    have "continuous_on ?T ?f"
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  1095
      unfolding split_def continuous_on by (intro ballI tendsto_intros)
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  1096
    moreover have "compact ?T"
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  1097
      by (intro compact_Times compact_Icc insert)
44525
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  1098
    ultimately have "compact (?f ` ?T)"
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  1099
      by (rule compact_continuous_image)
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  1100
    also have "?f ` ?T = convex hull (insert x A)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  1101
      unfolding convex_hull_insert [OF \<open>A \<noteq> {}\<close>]
44525
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  1102
      apply safe
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  1103
      apply (rule_tac x=a in exI, simp)
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1104
      apply (rule_tac x="1 - a" in exI, simp, fast)
44525
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  1105
      apply (rule_tac x="(u, b)" in image_eqI, simp_all)
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  1106
      done
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  1107
    finally show "compact (convex hull (insert x A))" .
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  1108
  qed
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  1109
qed
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  1110
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1111
lemma compact_convex_hull:
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1112
  fixes S :: "'a::euclidean_space set"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1113
  assumes "compact S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1114
  shows "compact (convex hull S)"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1115
proof (cases "S = {}")
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1116
  case True
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1117
  then show ?thesis using compact_empty by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1118
next
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1119
  case False
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1120
  then obtain w where "w \<in> S" by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1121
  show ?thesis
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1122
    unfolding caratheodory[of S]
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1123
  proof (induct ("DIM('a) + 1"))
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1124
    case 0
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1125
    have *: "{x.\<exists>sa. finite sa \<and> sa \<subseteq> S \<and> card sa \<le> 0 \<and> x \<in> convex hull sa} = {}"
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36341
diff changeset
  1126
      using compact_empty by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1127
    from 0 show ?case unfolding * by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1128
  next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1129
    case (Suc n)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1130
    show ?case
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1131
    proof (cases "n = 0")
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1132
      case True
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1133
      have "{x. \<exists>T. finite T \<and> T \<subseteq> S \<and> card T \<le> Suc n \<and> x \<in> convex hull T} = S"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1134
        unfolding set_eq_iff and mem_Collect_eq
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1135
      proof (rule, rule)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1136
        fix x
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1137
        assume "\<exists>T. finite T \<and> T \<subseteq> S \<and> card T \<le> Suc n \<and> x \<in> convex hull T"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1138
        then obtain T where T: "finite T" "T \<subseteq> S" "card T \<le> Suc n" "x \<in> convex hull T"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1139
          by auto
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1140
        show "x \<in> S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1141
        proof (cases "card T = 0")
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1142
          case True
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1143
          then show ?thesis
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1144
            using T(4) unfolding card_0_eq[OF T(1)] by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1145
        next
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1146
          case False
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1147
          then have "card T = Suc 0" using T(3) \<open>n=0\<close> by auto
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1148
          then obtain a where "T = {a}" unfolding card_Suc_eq by auto
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1149
          then show ?thesis using T(2,4) by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1150
        qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1151
      next
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1152
        fix x assume "x\<in>S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1153
        then show "\<exists>T. finite T \<and> T \<subseteq> S \<and> card T \<le> Suc n \<and> x \<in> convex hull T"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1154
          apply (rule_tac x="{x}" in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1155
          unfolding convex_hull_singleton
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1156
          apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1157
          done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1158
      qed
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1159
      then show ?thesis using assms by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1160
    next
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1161
      case False
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1162
      have "{x. \<exists>T. finite T \<and> T \<subseteq> S \<and> card T \<le> Suc n \<and> x \<in> convex hull T} =
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1163
        {(1 - u) *\<^sub>R x + u *\<^sub>R y | x y u.
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1164
          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}}"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1165
        unfolding set_eq_iff and mem_Collect_eq
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1166
      proof (rule, rule)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1167
        fix x
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1168
        assume "\<exists>u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \<and>
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1169
          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)"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1170
        then obtain u v c T where obt: "x = (1 - c) *\<^sub>R u + c *\<^sub>R v"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1171
          "0 \<le> c \<and> c \<le> 1" "u \<in> S" "finite T" "T \<subseteq> S" "card T \<le> n"  "v \<in> convex hull T"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1172
          by auto
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1173
        moreover have "(1 - c) *\<^sub>R u + c *\<^sub>R v \<in> convex hull insert u T"
61426
d53db136e8fd new material on path_component_sets, inside, outside, etc. And more default simprules
paulson <lp15@cam.ac.uk>
parents: 61222
diff changeset
  1174
          apply (rule convexD_alt)
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1175
          using obt(2) and convex_convex_hull and hull_subset[of "insert u T" convex]
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1176
          using obt(7) and hull_mono[of T "insert u T"]
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1177
          apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1178
          done
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1179
        ultimately show "\<exists>T. finite T \<and> T \<subseteq> S \<and> card T \<le> Suc n \<and> x \<in> convex hull T"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1180
          apply (rule_tac x="insert u T" in exI)
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1181
          apply (auto simp: card_insert_if)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1182
          done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1183
      next
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1184
        fix x
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1185
        assume "\<exists>T. finite T \<and> T \<subseteq> S \<and> card T \<le> Suc n \<and> x \<in> convex hull T"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1186
        then obtain T where T: "finite T" "T \<subseteq> S" "card T \<le> Suc n" "x \<in> convex hull T"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1187
          by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1188
        show "\<exists>u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \<and>
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1189
          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)"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1190
        proof (cases "card T = Suc n")
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1191
          case False
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1192
          then have "card T \<le> n" using T(3) by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1193
          then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1194
            apply (rule_tac x=w in exI, rule_tac x=x in exI, rule_tac x=1 in exI)
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1195
            using \<open>w\<in>S\<close> and T
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1196
            apply (auto intro!: exI[where x=T])
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1197
            done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1198
        next
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1199
          case True
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1200
          then obtain a u where au: "T = insert a u" "a\<notin>u"
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1201
            apply (drule_tac card_eq_SucD, auto)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1202
            done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1203
          show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1204
          proof (cases "u = {}")
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1205
            case True
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1206
            then have "x = a" using T(4)[unfolded au] by auto
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  1207
            show ?thesis unfolding \<open>x = a\<close>
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1208
              apply (rule_tac x=a in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1209
              apply (rule_tac x=a in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1210
              apply (rule_tac x=1 in exI)
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1211
              using T and \<open>n \<noteq> 0\<close>
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1212
              unfolding au
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1213
              apply (auto intro!: exI[where x="{a}"])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1214
              done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1215
          next
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1216
            case False
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1217
            obtain ux vx b where obt: "ux\<ge>0" "vx\<ge>0" "ux + vx = 1"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1218
              "b \<in> convex hull u" "x = ux *\<^sub>R a + vx *\<^sub>R b"
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1219
              using T(4)[unfolded au convex_hull_insert[OF False]]
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1220
              by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1221
            have *: "1 - vx = ux" using obt(3) by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1222
            show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1223
              apply (rule_tac x=a in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1224
              apply (rule_tac x=b in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1225
              apply (rule_tac x=vx in exI)
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1226
              using obt and T(1-3)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1227
              unfolding au and * using card_insert_disjoint[OF _ au(2)]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1228
              apply (auto intro!: exI[where x=u])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1229
              done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1230
          qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1231
        qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1232
      qed
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1233
      then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1234
        using compact_convex_combinations[OF assms Suc] by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1235
    qed
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36341
diff changeset
  1236
  qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1237
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1238
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1239
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70097
diff changeset
  1240
subsection\<^marker>\<open>tag unimportant\<close> \<open>Extremal points of a simplex are some vertices\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1241
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1242
lemma dist_increases_online:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1243
  fixes a b d :: "'a::real_inner"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1244
  assumes "d \<noteq> 0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1245
  shows "dist a (b + d) > dist a b \<or> dist a (b - d) > dist a b"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1246
proof (cases "inner a d - inner b d > 0")
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1247
  case True
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1248
  then have "0 < inner d d + (inner a d * 2 - inner b d * 2)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1249
    apply (rule_tac add_pos_pos)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1250
    using assms
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1251
    apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1252
    done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1253
  then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1254
    apply (rule_tac disjI2)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1255
    unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1256
    apply  (simp add: algebra_simps inner_commute)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1257
    done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1258
next
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1259
  case False
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1260
  then have "0 < inner d d + (inner b d * 2 - inner a d * 2)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1261
    apply (rule_tac add_pos_nonneg)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1262
    using assms
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1263
    apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1264
    done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1265
  then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1266
    apply (rule_tac disjI1)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1267
    unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1268
    apply (simp add: algebra_simps inner_commute)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1269
    done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1270
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1271
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1272
lemma norm_increases_online:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1273
  fixes d :: "'a::real_inner"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1274
  shows "d \<noteq> 0 \<Longrightarrow> norm (a + d) > norm a \<or> norm(a - d) > norm a"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1275
  using dist_increases_online[of d a 0] unfolding dist_norm by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1276
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1277
lemma simplex_furthest_lt:
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1278
  fixes S :: "'a::real_inner set"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1279
  assumes "finite S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1280
  shows "\<forall>x \<in> convex hull S.  x \<notin> S \<longrightarrow> (\<exists>y \<in> convex hull S. norm (x - a) < norm(y - a))"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1281
  using assms
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1282
proof induct
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1283
  fix x S
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1284
  assume as: "finite S" "x\<notin>S" "\<forall>x\<in>convex hull S. x \<notin> S \<longrightarrow> (\<exists>y\<in>convex hull S. norm (x - a) < norm (y - a))"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1285
  show "\<forall>xa\<in>convex hull insert x S. xa \<notin> insert x S \<longrightarrow>
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1286
    (\<exists>y\<in>convex hull insert x S. norm (xa - a) < norm (y - a))"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1287
  proof (intro impI ballI, cases "S = {}")
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1288
    case False
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1289
    fix y
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1290
    assume y: "y \<in> convex hull insert x S" "y \<notin> insert x S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1291
    obtain u v b where obt: "u\<ge>0" "v\<ge>0" "u + v = 1" "b \<in> convex hull S" "y = u *\<^sub>R x + v *\<^sub>R b"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1292
      using y(1)[unfolded convex_hull_insert[OF False]] by auto
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1293
    show "\<exists>z\<in>convex hull insert x S. norm (y - a) < norm (z - a)"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1294
    proof (cases "y \<in> convex hull S")
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1295
      case True
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1296
      then obtain z where "z \<in> convex hull S" "norm (y - a) < norm (z - a)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1297
        using as(3)[THEN bspec[where x=y]] and y(2) by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1298
      then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1299
        apply (rule_tac x=z in bexI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1300
        unfolding convex_hull_insert[OF False]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1301
        apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1302
        done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1303
    next
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1304
      case False
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1305
      show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1306
        using obt(3)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1307
      proof (cases "u = 0", case_tac[!] "v = 0")
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1308
        assume "u = 0" "v \<noteq> 0"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1309
        then have "y = b" using obt by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1310
        then show ?thesis using False and obt(4) by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1311
      next
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1312
        assume "u \<noteq> 0" "v = 0"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1313
        then have "y = x" using obt by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1314
        then show ?thesis using y(2) by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1315
      next
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1316
        assume "u \<noteq> 0" "v \<noteq> 0"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1317
        then obtain w where w: "w>0" "w<u" "w<v"
68527
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68074
diff changeset
  1318
          using field_lbound_gt_zero[of u v] and obt(1,2) by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1319
        have "x \<noteq> b"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1320
        proof
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1321
          assume "x = b"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1322
          then have "y = b" unfolding obt(5)
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1323
            using obt(3) by (auto simp: scaleR_left_distrib[symmetric])
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1324
          then show False using obt(4) and False by simp
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1325
        qed
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1326
        then have *: "w *\<^sub>R (x - b) \<noteq> 0" using w(1) by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1327
        show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1328
          using dist_increases_online[OF *, of a y]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1329
        proof (elim disjE)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1330
          assume "dist a y < dist a (y + w *\<^sub>R (x - b))"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1331
          then have "norm (y - a) < norm ((u + w) *\<^sub>R x + (v - w) *\<^sub>R b - a)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1332
            unfolding dist_commute[of a]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1333
            unfolding dist_norm obt(5)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1334
            by (simp add: algebra_simps)
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1335
          moreover have "(u + w) *\<^sub>R x + (v - w) *\<^sub>R b \<in> convex hull insert x S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1336
            unfolding convex_hull_insert[OF \<open>S\<noteq>{}\<close>]
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1337
          proof (intro CollectI conjI exI)
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1338
            show "u + w \<ge> 0" "v - w \<ge> 0"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1339
              using obt(1) w by auto
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1340
          qed (use obt in auto)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1341
          ultimately show ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1342
        next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1343
          assume "dist a y < dist a (y - w *\<^sub>R (x - b))"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1344
          then have "norm (y - a) < norm ((u - w) *\<^sub>R x + (v + w) *\<^sub>R b - a)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1345
            unfolding dist_commute[of a]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1346
            unfolding dist_norm obt(5)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1347
            by (simp add: algebra_simps)
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1348
          moreover have "(u - w) *\<^sub>R x + (v + w) *\<^sub>R b \<in> convex hull insert x S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1349
            unfolding convex_hull_insert[OF \<open>S\<noteq>{}\<close>]
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1350
          proof (intro CollectI conjI exI)
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1351
            show "u - w \<ge> 0" "v + w \<ge> 0"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1352
              using obt(1) w by auto
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1353
          qed (use obt in auto)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1354
          ultimately show ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1355
        qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1356
      qed auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1357
    qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1358
  qed auto
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1359
qed (auto simp: assms)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1360
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1361
lemma simplex_furthest_le:
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1362
  fixes S :: "'a::real_inner set"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1363
  assumes "finite S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1364
    and "S \<noteq> {}"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1365
  shows "\<exists>y\<in>S. \<forall>x\<in> convex hull S. norm (x - a) \<le> norm (y - a)"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1366
proof -
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1367
  have "convex hull S \<noteq> {}"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1368
    using hull_subset[of S convex] and assms(2) by auto
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1369
  then obtain x where x: "x \<in> convex hull S" "\<forall>y\<in>convex hull S. norm (y - a) \<le> norm (x - a)"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1370
    using distance_attains_sup[OF finite_imp_compact_convex_hull[OF \<open>finite S\<close>], of a]
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1371
    unfolding dist_commute[of a]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1372
    unfolding dist_norm
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1373
    by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1374
  show ?thesis
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1375
  proof (cases "x \<in> S")
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1376
    case False
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1377
    then obtain y where "y \<in> convex hull S" "norm (x - a) < norm (y - a)"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1378
      using simplex_furthest_lt[OF assms(1), THEN bspec[where x=x]] and x(1)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1379
      by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1380
    then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1381
      using x(2)[THEN bspec[where x=y]] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1382
  next
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1383
    case True
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1384
    with x show ?thesis by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1385
  qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1386
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1387
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1388
lemma simplex_furthest_le_exists:
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1389
  fixes S :: "('a::real_inner) set"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1390
  shows "finite S \<Longrightarrow> \<forall>x\<in>(convex hull S). \<exists>y\<in>S. norm (x - a) \<le> norm (y - a)"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1391
  using simplex_furthest_le[of S] by (cases "S = {}") auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1392
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1393
lemma simplex_extremal_le:
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1394
  fixes S :: "'a::real_inner set"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1395
  assumes "finite S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1396
    and "S \<noteq> {}"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1397
  shows "\<exists>u\<in>S. \<exists>v\<in>S. \<forall>x\<in>convex hull S. \<forall>y \<in> convex hull S. norm (x - y) \<le> norm (u - v)"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1398
proof -
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1399
  have "convex hull S \<noteq> {}"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1400
    using hull_subset[of S convex] and assms(2) by auto
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1401
  then obtain u v where obt: "u \<in> convex hull S" "v \<in> convex hull S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1402
    "\<forall>x\<in>convex hull S. \<forall>y\<in>convex hull S. norm (x - y) \<le> norm (u - v)"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1403
    using compact_sup_maxdistance[OF finite_imp_compact_convex_hull[OF assms(1)]]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1404
    by (auto simp: dist_norm)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1405
  then show ?thesis
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1406
  proof (cases "u\<notin>S \<or> v\<notin>S", elim disjE)
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1407
    assume "u \<notin> S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1408
    then obtain y where "y \<in> convex hull S" "norm (u - v) < norm (y - v)"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1409
      using simplex_furthest_lt[OF assms(1), THEN bspec[where x=u]] and obt(1)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1410
      by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1411
    then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1412
      using obt(3)[THEN bspec[where x=y], THEN bspec[where x=v]] and obt(2)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1413
      by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1414
  next
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1415
    assume "v \<notin> S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1416
    then obtain y where "y \<in> convex hull S" "norm (v - u) < norm (y - u)"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1417
      using simplex_furthest_lt[OF assms(1), THEN bspec[where x=v]] and obt(2)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1418
      by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1419
    then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1420
      using obt(3)[THEN bspec[where x=u], THEN bspec[where x=y]] and obt(1)
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1421
      by (auto simp: norm_minus_commute)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1422
  qed auto
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  1423
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1424
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1425
lemma simplex_extremal_le_exists:
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1426
  fixes S :: "'a::real_inner set"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1427
  shows "finite S \<Longrightarrow> x \<in> convex hull S \<Longrightarrow> y \<in> convex hull S \<Longrightarrow>
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1428
    \<exists>u\<in>S. \<exists>v\<in>S. norm (x - y) \<le> norm (u - v)"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1429
  using convex_hull_empty simplex_extremal_le[of S]
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1430
  by(cases "S = {}") auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1431
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1432
67968
a5ad4c015d1c removed dots at the end of (sub)titles
nipkow
parents: 67962
diff changeset
  1433
subsection \<open>Closest point of a convex set is unique, with a continuous projection\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1434
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70097
diff changeset
  1435
definition\<^marker>\<open>tag important\<close> closest_point :: "'a::{real_inner,heine_borel} set \<Rightarrow> 'a \<Rightarrow> 'a"
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1436
  where "closest_point S a = (SOME x. x \<in> S \<and> (\<forall>y\<in>S. dist a x \<le> dist a y))"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1437
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1438
lemma closest_point_exists:
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1439
  assumes "closed S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1440
    and "S \<noteq> {}"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1441
  shows "closest_point S a \<in> S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1442
    and "\<forall>y\<in>S. dist a (closest_point S a) \<le> dist a y"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1443
  unfolding closest_point_def
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1444
  apply(rule_tac[!] someI2_ex)
62381
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
  1445
  apply (auto intro: distance_attains_inf[OF assms(1,2), of a])
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1446
  done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1447
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1448
lemma closest_point_in_set: "closed S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> closest_point S a \<in> S"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1449
  by (meson closest_point_exists)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1450
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1451
lemma closest_point_le: "closed S \<Longrightarrow> x \<in> S \<Longrightarrow> dist a (closest_point S a) \<le> dist a x"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1452
  using closest_point_exists[of S] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1453
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1454
lemma closest_point_self:
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1455
  assumes "x \<in> S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1456
  shows "closest_point S x = x"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1457
  unfolding closest_point_def
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1458
  apply (rule some1_equality, rule ex1I[of _ x])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1459
  using assms
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1460
  apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1461
  done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1462
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1463
lemma closest_point_refl: "closed S \<Longrightarrow> S \<noteq> {} \<Longrightarrow> closest_point S x = x \<longleftrightarrow> x \<in> S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1464
  using closest_point_in_set[of S x] closest_point_self[of x S]
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1465
  by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1466
36337
87b6c83d7ed7 generalize constant closest_point
huffman
parents: 36071
diff changeset
  1467
lemma closer_points_lemma:
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1468
  assumes "inner y z > 0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1469
  shows "\<exists>u>0. \<forall>v>0. v \<le> u \<longrightarrow> norm(v *\<^sub>R z - y) < norm y"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1470
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1471
  have z: "inner z z > 0"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1472
    unfolding inner_gt_zero_iff using assms by auto
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1473
  have "norm (v *\<^sub>R z - y) < norm y"
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1474
    if "0 < v" and "v \<le> inner y z / inner z z" for v
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1475
    unfolding norm_lt using z assms that
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1476
    by (simp add: field_simps inner_diff inner_commute mult_strict_left_mono[OF _ \<open>0<v\<close>])
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1477
  then show ?thesis
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1478
    using assms z
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1479
    by (rule_tac x = "inner y z / inner z z" in exI) auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1480
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1481
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1482
lemma closer_point_lemma:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1483
  assumes "inner (y - x) (z - x) > 0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1484
  shows "\<exists>u>0. u \<le> 1 \<and> dist (x + u *\<^sub>R (z - x)) y < dist x y"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1485
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1486
  obtain u where "u > 0"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1487
    and u: "\<forall>v>0. v \<le> u \<longrightarrow> norm (v *\<^sub>R (z - x) - (y - x)) < norm (y - x)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1488
    using closer_points_lemma[OF assms] by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1489
  show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1490
    apply (rule_tac x="min u 1" in exI)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  1491
    using u[THEN spec[where x="min u 1"]] and \<open>u > 0\<close>
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1492
    unfolding dist_norm by (auto simp: norm_minus_commute field_simps)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1493
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1494
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1495
lemma any_closest_point_dot:
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1496
  assumes "convex S" "closed S" "x \<in> S" "y \<in> S" "\<forall>z\<in>S. dist a x \<le> dist a z"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1497
  shows "inner (a - x) (y - x) \<le> 0"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1498
proof (rule ccontr)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1499
  assume "\<not> ?thesis"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1500
  then obtain u where u: "u>0" "u\<le>1" "dist (x + u *\<^sub>R (y - x)) a < dist x a"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1501
    using closer_point_lemma[of a x y] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1502
  let ?z = "(1 - u) *\<^sub>R x + u *\<^sub>R y"
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1503
  have "?z \<in> S"
61426
d53db136e8fd new material on path_component_sets, inside, outside, etc. And more default simprules
paulson <lp15@cam.ac.uk>
parents: 61222
diff changeset
  1504
    using convexD_alt[OF assms(1,3,4), of u] using u by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1505
  then show False
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1506
    using assms(5)[THEN bspec[where x="?z"]] and u(3)
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1507
    by (auto simp: dist_commute algebra_simps)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1508
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1509
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1510
lemma any_closest_point_unique:
36337
87b6c83d7ed7 generalize constant closest_point
huffman
parents: 36071
diff changeset
  1511
  fixes x :: "'a::real_inner"
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1512
  assumes "convex S" "closed S" "x \<in> S" "y \<in> S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1513
    "\<forall>z\<in>S. dist a x \<le> dist a z" "\<forall>z\<in>S. dist a y \<le> dist a z"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1514
  shows "x = y"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1515
  using any_closest_point_dot[OF assms(1-4,5)] and any_closest_point_dot[OF assms(1-2,4,3,6)]
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1516
  unfolding norm_pths(1) and norm_le_square
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1517
  by (auto simp: algebra_simps)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1518
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1519
lemma closest_point_unique:
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1520
  assumes "convex S" "closed S" "x \<in> S" "\<forall>z\<in>S. dist a x \<le> dist a z"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1521
  shows "x = closest_point S a"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1522
  using any_closest_point_unique[OF assms(1-3) _ assms(4), of "closest_point S a"]
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1523
  using closest_point_exists[OF assms(2)] and assms(3) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1524
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1525
lemma closest_point_dot:
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1526
  assumes "convex S" "closed S" "x \<in> S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1527
  shows "inner (a - closest_point S a) (x - closest_point S a) \<le> 0"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1528
  apply (rule any_closest_point_dot[OF assms(1,2) _ assms(3)])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1529
  using closest_point_exists[OF assms(2)] and assms(3)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1530
  apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1531
  done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1532
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1533
lemma closest_point_lt:
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1534
  assumes "convex S" "closed S" "x \<in> S" "x \<noteq> closest_point S a"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1535
  shows "dist a (closest_point S a) < dist a x"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1536
  apply (rule ccontr)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1537
  apply (rule_tac notE[OF assms(4)])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1538
  apply (rule closest_point_unique[OF assms(1-3), of a])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1539
  using closest_point_le[OF assms(2), of _ a]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1540
  apply fastforce
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1541
  done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1542
69618
2be1baf40351 moved setdist to more appropriate places
immler
parents: 69617
diff changeset
  1543
lemma setdist_closest_point:
2be1baf40351 moved setdist to more appropriate places
immler
parents: 69617
diff changeset
  1544
    "\<lbrakk>closed S; S \<noteq> {}\<rbrakk> \<Longrightarrow> setdist {a} S = dist a (closest_point S a)"
2be1baf40351 moved setdist to more appropriate places
immler
parents: 69617
diff changeset
  1545
  apply (rule setdist_unique)
2be1baf40351 moved setdist to more appropriate places
immler
parents: 69617
diff changeset
  1546
  using closest_point_le
2be1baf40351 moved setdist to more appropriate places
immler
parents: 69617
diff changeset
  1547
  apply (auto simp: closest_point_in_set)
2be1baf40351 moved setdist to more appropriate places
immler
parents: 69617
diff changeset
  1548
  done
2be1baf40351 moved setdist to more appropriate places
immler
parents: 69617
diff changeset
  1549
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1550
lemma closest_point_lipschitz:
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1551
  assumes "convex S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1552
    and "closed S" "S \<noteq> {}"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1553
  shows "dist (closest_point S x) (closest_point S y) \<le> dist x y"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1554
proof -
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1555
  have "inner (x - closest_point S x) (closest_point S y - closest_point S x) \<le> 0"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1556
    and "inner (y - closest_point S y) (closest_point S x - closest_point S y) \<le> 0"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1557
    apply (rule_tac[!] any_closest_point_dot[OF assms(1-2)])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1558
    using closest_point_exists[OF assms(2-3)]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1559
    apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1560
    done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1561
  then show ?thesis unfolding dist_norm and norm_le
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1562
    using inner_ge_zero[of "(x - closest_point S x) - (y - closest_point S y)"]
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1563
    by (simp add: inner_add inner_diff inner_commute)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1564
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1565
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1566
lemma continuous_at_closest_point:
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1567
  assumes "convex S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1568
    and "closed S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1569
    and "S \<noteq> {}"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1570
  shows "continuous (at x) (closest_point S)"
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  1571
  unfolding continuous_at_eps_delta
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1572
  using le_less_trans[OF closest_point_lipschitz[OF assms]] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1573
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1574
lemma continuous_on_closest_point:
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1575
  assumes "convex S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1576
    and "closed S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1577
    and "S \<noteq> {}"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1578
  shows "continuous_on t (closest_point S)"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1579
  by (metis continuous_at_imp_continuous_on continuous_at_closest_point[OF assms])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1580
63881
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1581
proposition closest_point_in_rel_interior:
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1582
  assumes "closed S" "S \<noteq> {}" and x: "x \<in> affine hull S"
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1583
    shows "closest_point S x \<in> rel_interior S \<longleftrightarrow> x \<in> rel_interior S"
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1584
proof (cases "x \<in> S")
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1585
  case True
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1586
  then show ?thesis
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1587
    by (simp add: closest_point_self)
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1588
next
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1589
  case False
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1590
  then have "False" if asm: "closest_point S x \<in> rel_interior S"
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1591
  proof -
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1592
    obtain e where "e > 0" and clox: "closest_point S x \<in> S"
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1593
               and e: "cball (closest_point S x) e \<inter> affine hull S \<subseteq> S"
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1594
      using asm mem_rel_interior_cball by blast
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1595
    then have clo_notx: "closest_point S x \<noteq> x"
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1596
      using \<open>x \<notin> S\<close> by auto
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1597
    define y where "y \<equiv> closest_point S x -
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1598
                        (min 1 (e / norm(closest_point S x - x))) *\<^sub>R (closest_point S x - x)"
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1599
    have "x - y = (1 - min 1 (e / norm (closest_point S x - x))) *\<^sub>R (x - closest_point S x)"
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1600
      by (simp add: y_def algebra_simps)
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1601
    then have "norm (x - y) = abs ((1 - min 1 (e / norm (closest_point S x - x)))) * norm(x - closest_point S x)"
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1602
      by simp
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1603
    also have "\<dots> < norm(x - closest_point S x)"
63881
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1604
      using clo_notx \<open>e > 0\<close>
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70802
diff changeset
  1605
      by (auto simp: mult_less_cancel_right2 field_split_simps)
63881
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1606
    finally have no_less: "norm (x - y) < norm (x - closest_point S x)" .
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1607
    have "y \<in> affine hull S"
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1608
      unfolding y_def
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1609
      by (meson affine_affine_hull clox hull_subset mem_affine_3_minus2 subsetD x)
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1610
    moreover have "dist (closest_point S x) y \<le> e"
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1611
      using \<open>e > 0\<close> by (auto simp: y_def min_mult_distrib_right)
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1612
    ultimately have "y \<in> S"
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1613
      using subsetD [OF e] by simp
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1614
    then have "dist x (closest_point S x) \<le> dist x y"
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1615
      by (simp add: closest_point_le \<open>closed S\<close>)
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1616
    with no_less show False
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1617
      by (simp add: dist_norm)
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1618
  qed
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1619
  moreover have "x \<notin> rel_interior S"
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1620
    using rel_interior_subset False by blast
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1621
  ultimately show ?thesis by blast
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1622
qed
b746b19197bd lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
  1623
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1624
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70097
diff changeset
  1625
subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Various point-to-set separating/supporting hyperplane theorems\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1626
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1627
lemma supporting_hyperplane_closed_point:
36337
87b6c83d7ed7 generalize constant closest_point
huffman
parents: 36071
diff changeset
  1628
  fixes z :: "'a::{real_inner,heine_borel}"
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1629
  assumes "convex S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1630
    and "closed S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1631
    and "S \<noteq> {}"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1632
    and "z \<notin> S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1633
  shows "\<exists>a b. \<exists>y\<in>S. inner a z < b \<and> inner a y = b \<and> (\<forall>x\<in>S. inner a x \<ge> b)"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1634
proof -
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1635
  obtain y where "y \<in> S" and y: "\<forall>x\<in>S. dist z y \<le> dist z x"
63075
60a42a4166af lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents: 63072
diff changeset
  1636
    by (metis distance_attains_inf[OF assms(2-3)])
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1637
  show ?thesis
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1638
  proof (intro exI bexI conjI ballI)
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1639
    show "(y - z) \<bullet> z < (y - z) \<bullet> y"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1640
      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)
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1641
    show "(y - z) \<bullet> y \<le> (y - z) \<bullet> x" if "x \<in> S" for x
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1642
    proof (rule ccontr)
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1643
      have *: "\<And>u. 0 \<le> u \<and> u \<le> 1 \<longrightarrow> dist z y \<le> dist z ((1 - u) *\<^sub>R y + u *\<^sub>R x)"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1644
        using assms(1)[unfolded convex_alt] and y and \<open>x\<in>S\<close> and \<open>y\<in>S\<close> by auto
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1645
      assume "\<not> (y - z) \<bullet> y \<le> (y - z) \<bullet> x"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1646
      then obtain v where "v > 0" "v \<le> 1" "dist (y + v *\<^sub>R (x - y)) z < dist y z"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1647
        using closer_point_lemma[of z y x] by (auto simp: inner_diff)
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1648
      then show False
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1649
        using *[of v] by (auto simp: dist_commute algebra_simps)
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1650
    qed
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1651
  qed (use \<open>y \<in> S\<close> in auto)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1652
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1653
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1654
lemma separating_hyperplane_closed_point:
36337
87b6c83d7ed7 generalize constant closest_point
huffman
parents: 36071
diff changeset
  1655
  fixes z :: "'a::{real_inner,heine_borel}"
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1656
  assumes "convex S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1657
    and "closed S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1658
    and "z \<notin> S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1659
  shows "\<exists>a b. inner a z < b \<and> (\<forall>x\<in>S. inner a x > b)"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1660
proof (cases "S = {}")
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1661
  case True
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1662
  then show ?thesis
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1663
    by (simp add: gt_ex)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1664
next
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1665
  case False
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1666
  obtain y where "y \<in> S" and y: "\<And>x. x \<in> S \<Longrightarrow> dist z y \<le> dist z x"
62381
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
  1667
    by (metis distance_attains_inf[OF assms(2) False])
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1668
  show ?thesis
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1669
  proof (intro exI conjI ballI)
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1670
    show "(y - z) \<bullet> z < inner (y - z) z + (norm (y - z))\<^sup>2 / 2"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1671
      using \<open>y\<in>S\<close> \<open>z\<notin>S\<close> by auto
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1672
  next
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1673
    fix x
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1674
    assume "x \<in> S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1675
    have "False" if *: "0 < inner (z - y) (x - y)"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1676
    proof -
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1677
      obtain u where "u > 0" "u \<le> 1" "dist (y + u *\<^sub>R (x - y)) z < dist y z"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1678
        using * closer_point_lemma by blast
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1679
      then show False using y[of "y + u *\<^sub>R (x - y)"] convexD_alt [OF \<open>convex S\<close>]
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1680
        using \<open>x\<in>S\<close> \<open>y\<in>S\<close> by (auto simp: dist_commute algebra_simps)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1681
    qed
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1682
    moreover have "0 < (norm (y - z))\<^sup>2"
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1683
      using \<open>y\<in>S\<close> \<open>z\<notin>S\<close> by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1684
    then have "0 < inner (y - z) (y - z)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1685
      unfolding power2_norm_eq_inner by simp
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1686
    ultimately show "(y - z) \<bullet> z + (norm (y - z))\<^sup>2 / 2 < (y - z) \<bullet> x"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1687
      by (force simp: field_simps power2_norm_eq_inner inner_commute inner_diff)
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1688
  qed 
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1689
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1690
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1691
lemma separating_hyperplane_closed_0:
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1692
  assumes "convex (S::('a::euclidean_space) set)"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1693
    and "closed S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1694
    and "0 \<notin> S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1695
  shows "\<exists>a b. a \<noteq> 0 \<and> 0 < b \<and> (\<forall>x\<in>S. inner a x > b)"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1696
proof (cases "S = {}")
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50104
diff changeset
  1697
  case True
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1698
  have "(SOME i. i\<in>Basis) \<noteq> (0::'a)"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1699
    by (metis Basis_zero SOME_Basis)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1700
  then show ?thesis
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1701
    using True zero_less_one by blast
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1702
next
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1703
  case False
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1704
  then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1705
    using False using separating_hyperplane_closed_point[OF assms]
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1706
    by (metis all_not_in_conv inner_zero_left inner_zero_right less_eq_real_def not_le)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1707
qed
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1708
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1709
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70097
diff changeset
  1710
subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Now set-to-set for closed/compact sets\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1711
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1712
lemma separating_hyperplane_closed_compact:
65038
9391ea7daa17 new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  1713
  fixes S :: "'a::euclidean_space set"
9391ea7daa17 new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  1714
  assumes "convex S"
9391ea7daa17 new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  1715
    and "closed S"
9391ea7daa17 new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  1716
    and "convex T"
9391ea7daa17 new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  1717
    and "compact T"
9391ea7daa17 new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  1718
    and "T \<noteq> {}"
9391ea7daa17 new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  1719
    and "S \<inter> T = {}"
9391ea7daa17 new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  1720
  shows "\<exists>a b. (\<forall>x\<in>S. inner a x < b) \<and> (\<forall>x\<in>T. inner a x > b)"
9391ea7daa17 new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  1721
proof (cases "S = {}")
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1722
  case True
65038
9391ea7daa17 new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  1723
  obtain b where b: "b > 0" "\<forall>x\<in>T. norm x \<le> b"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1724
    using compact_imp_bounded[OF assms(4)] unfolding bounded_pos by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1725
  obtain z :: 'a where z: "norm z = b + 1"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1726
    using vector_choose_size[of "b + 1"] and b(1) by auto
65038
9391ea7daa17 new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  1727
  then have "z \<notin> T" using b(2)[THEN bspec[where x=z]] by auto
9391ea7daa17 new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  1728
  then obtain a b where ab: "inner a z < b" "\<forall>x\<in>T. b < inner a x"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1729
    using separating_hyperplane_closed_point[OF assms(3) compact_imp_closed[OF assms(4)], of z]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1730
    by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1731
  then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1732
    using True by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1733
next
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1734
  case False
65038
9391ea7daa17 new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  1735
  then obtain y where "y \<in> S" by auto
9391ea7daa17 new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  1736
  obtain a b where "0 < b" "\<forall>x \<in> (\<Union>x\<in> S. \<Union>y \<in> T. {x - y}). b < inner a x"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1737
    using separating_hyperplane_closed_point[OF convex_differences[OF assms(1,3)], of 0]
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1738
    using closed_compact_differences[OF assms(2,4)]
65038
9391ea7daa17 new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  1739
    using assms(6) by auto 
9391ea7daa17 new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  1740
  then have ab: "\<forall>x\<in>S. \<forall>y\<in>T. b + inner a y < inner a x"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1741
    apply -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1742
    apply rule
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1743
    apply rule
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1744
    apply (erule_tac x="x - y" in ballE)
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1745
    apply (auto simp: inner_diff)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1746
    done
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69064
diff changeset
  1747
  define k where "k = (SUP x\<in>T. a \<bullet> x)"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1748
  show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1749
    apply (rule_tac x="-a" in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1750
    apply (rule_tac x="-(k + b / 2)" in exI)
54263
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54258
diff changeset
  1751
    apply (intro conjI ballI)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1752
    unfolding inner_minus_left and neg_less_iff_less
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1753
  proof -
65038
9391ea7daa17 new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  1754
    fix x assume "x \<in> T"
54263
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54258
diff changeset
  1755
    then have "inner a x - b / 2 < k"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1756
      unfolding k_def
54263
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54258
diff changeset
  1757
    proof (subst less_cSUP_iff)
65038
9391ea7daa17 new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  1758
      show "T \<noteq> {}" by fact
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67155
diff changeset
  1759
      show "bdd_above ((\<bullet>) a ` T)"
65038
9391ea7daa17 new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  1760
        using ab[rule_format, of y] \<open>y \<in> S\<close>
54263
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54258
diff changeset
  1761
        by (intro bdd_aboveI2[where M="inner a y - b"]) (auto simp: field_simps intro: less_imp_le)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  1762
    qed (auto intro!: bexI[of _ x] \<open>0<b\<close>)
54263
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54258
diff changeset
  1763
    then show "inner a x < k + b / 2"
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54258
diff changeset
  1764
      by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1765
  next
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1766
    fix x
65038
9391ea7daa17 new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  1767
    assume "x \<in> S"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1768
    then have "k \<le> inner a x - b"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1769
      unfolding k_def
54263
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54258
diff changeset
  1770
      apply (rule_tac cSUP_least)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1771
      using assms(5)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1772
      using ab[THEN bspec[where x=x]]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1773
      apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1774
      done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1775
    then show "k + b / 2 < inner a x"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  1776
      using \<open>0 < b\<close> by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1777
  qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1778
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1779
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1780
lemma separating_hyperplane_compact_closed:
65038
9391ea7daa17 new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  1781
  fixes S :: "'a::euclidean_space set"
9391ea7daa17 new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  1782
  assumes "convex S"
9391ea7daa17 new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  1783
    and "compact S"
9391ea7daa17 new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  1784
    and "S \<noteq> {}"
9391ea7daa17 new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  1785
    and "convex T"
9391ea7daa17 new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  1786
    and "closed T"
9391ea7daa17 new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  1787
    and "S \<inter> T = {}"
9391ea7daa17 new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  1788
  shows "\<exists>a b. (\<forall>x\<in>S. inner a x < b) \<and> (\<forall>x\<in>T. inner a x > b)"
9391ea7daa17 new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  1789
proof -
9391ea7daa17 new lemmas about segments, etc. Also recast some theorems to use Union rather than general set comprehensions
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  1790
  obtain a b where "(\<forall>x\<in>T. inner a x < b) \<and> (\<forall>x\<in>S. b < inner a x)"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1791
    using separating_hyperplane_closed_compact[OF assms(4-5,1-2,3)] and assms(6)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1792
    by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1793
  then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1794
    apply (rule_tac x="-a" in exI)
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1795
    apply (rule_tac x="-b" in exI, auto)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1796
    done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1797
qed
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1798
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1799
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70097
diff changeset
  1800
subsubsection\<^marker>\<open>tag unimportant\<close> \<open>General case without assuming closure and getting non-strict separation\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1801
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1802
lemma separating_hyperplane_set_0:
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1803
  assumes "convex S" "(0::'a::euclidean_space) \<notin> S"
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1804
  shows "\<exists>a. a \<noteq> 0 \<and> (\<forall>x\<in>S. 0 \<le> inner a x)"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1805
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1806
  let ?k = "\<lambda>c. {x::'a. 0 \<le> inner c x}"
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1807
  have *: "frontier (cball 0 1) \<inter> \<Inter>f \<noteq> {}" if as: "f \<subseteq> ?k ` S" "finite f" for f
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1808
  proof -
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1809
    obtain c where c: "f = ?k ` c" "c \<subseteq> S" "finite c"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1810
      using finite_subset_image[OF as(2,1)] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1811
    then obtain a b where ab: "a \<noteq> 0" "0 < b" "\<forall>x\<in>convex hull c. b < inner a x"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1812
      using separating_hyperplane_closed_0[OF convex_convex_hull, of c]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1813
      using finite_imp_compact_convex_hull[OF c(3), THEN compact_imp_closed] and assms(2)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1814
      using subset_hull[of convex, OF assms(1), symmetric, of c]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61531
diff changeset
  1815
      by force
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1816
    then have "\<exists>x. norm x = 1 \<and> (\<forall>y\<in>c. 0 \<le> inner y x)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1817
      apply (rule_tac x = "inverse(norm a) *\<^sub>R a" in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1818
      using hull_subset[of c convex]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1819
      unfolding subset_eq and inner_scaleR
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1820
      by (auto simp: inner_commute del: ballE elim!: ballE)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1821
    then show "frontier (cball 0 1) \<inter> \<Inter>f \<noteq> {}"
62381
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
  1822
      unfolding c(1) frontier_cball sphere_def dist_norm by auto
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
  1823
  qed
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1824
  have "frontier (cball 0 1) \<inter> (\<Inter>(?k ` S)) \<noteq> {}"
62381
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
  1825
    apply (rule compact_imp_fip)
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
  1826
    apply (rule compact_frontier[OF compact_cball])
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
  1827
    using * closed_halfspace_ge
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
  1828
    by auto
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1829
  then obtain x where "norm x = 1" "\<forall>y\<in>S. x\<in>?k y"
62381
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
  1830
    unfolding frontier_cball dist_norm sphere_def by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1831
  then show ?thesis
62381
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
  1832
    by (metis inner_commute mem_Collect_eq norm_eq_zero zero_neq_one)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1833
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1834
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1835
lemma separating_hyperplane_sets:
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1836
  fixes S T :: "'a::euclidean_space set"
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1837
  assumes "convex S"
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1838
    and "convex T"
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1839
    and "S \<noteq> {}"
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1840
    and "T \<noteq> {}"
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1841
    and "S \<inter> T = {}"
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1842
  shows "\<exists>a b. a \<noteq> 0 \<and> (\<forall>x\<in>S. inner a x \<le> b) \<and> (\<forall>x\<in>T. inner a x \<ge> b)"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1843
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1844
  from separating_hyperplane_set_0[OF convex_differences[OF assms(2,1)]]
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1845
  obtain a where "a \<noteq> 0" "\<forall>x\<in>{x - y |x y. x \<in> T \<and> y \<in> S}. 0 \<le> inner a x"
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1846
    using assms(3-5) by force
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1847
  then have *: "\<And>x y. x \<in> T \<Longrightarrow> y \<in> S \<Longrightarrow> inner a y \<le> inner a x"
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1848
    by (force simp: inner_diff)
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1849
  then have bdd: "bdd_above (((\<bullet>) a)`S)"
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1850
    using \<open>T \<noteq> {}\<close> by (auto intro: bdd_aboveI2[OF *])
54263
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54258
diff changeset
  1851
  show ?thesis
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  1852
    using \<open>a\<noteq>0\<close>
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69064
diff changeset
  1853
    by (intro exI[of _ a] exI[of _ "SUP x\<in>S. a \<bullet> x"])
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1854
       (auto intro!: cSUP_upper bdd cSUP_least \<open>a \<noteq> 0\<close> \<open>S \<noteq> {}\<close> *)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  1855
qed
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  1856
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  1857
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70097
diff changeset
  1858
subsection\<^marker>\<open>tag unimportant\<close> \<open>More convexity generalities\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1859
62948
7700f467892b lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 62843
diff changeset
  1860
lemma convex_closure [intro,simp]:
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1861
  fixes S :: "'a::real_normed_vector set"
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1862
  assumes "convex S"
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1863
  shows "convex (closure S)"
53676
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  1864
  apply (rule convexI)
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  1865
  apply (unfold closure_sequential, elim exE)
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  1866
  apply (rule_tac x="\<lambda>n. u *\<^sub>R xa n + v *\<^sub>R xb n" in exI)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1867
  apply (rule,rule)
53676
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  1868
  apply (rule convexD [OF assms])
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1869
  apply (auto del: tendsto_const intro!: tendsto_intros)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1870
  done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1871
62948
7700f467892b lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 62843
diff changeset
  1872
lemma convex_interior [intro,simp]:
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1873
  fixes S :: "'a::real_normed_vector set"
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1874
  assumes "convex S"
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1875
  shows "convex (interior S)"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1876
  unfolding convex_alt Ball_def mem_interior
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1877
proof clarify
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1878
  fix x y u
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1879
  assume u: "0 \<le> u" "u \<le> (1::real)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1880
  fix e d
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1881
  assume ed: "ball x e \<subseteq> S" "ball y d \<subseteq> S" "0<d" "0<e"
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1882
  show "\<exists>e>0. ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) e \<subseteq> S"
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1883
  proof (intro exI conjI subsetI)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1884
    fix z
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1885
    assume "z \<in> ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) (min d e)"
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1886
    then have "(1- u) *\<^sub>R (z - u *\<^sub>R (y - x)) + u *\<^sub>R (z + (1 - u) *\<^sub>R (y - x)) \<in> S"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1887
      apply (rule_tac assms[unfolded convex_alt, rule_format])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1888
      using ed(1,2) and u
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1889
      unfolding subset_eq mem_ball Ball_def dist_norm
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1890
      apply (auto simp: algebra_simps)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1891
      done
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1892
    then show "z \<in> S"
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1893
      using u by (auto simp: algebra_simps)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1894
  qed(insert u ed(3-4), auto)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1895
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1896
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1897
lemma convex_hull_eq_empty[simp]: "convex hull S = {} \<longleftrightarrow> S = {}"
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1898
  using hull_subset[of S convex] convex_hull_empty by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1899
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1900
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70097
diff changeset
  1901
subsection\<^marker>\<open>tag unimportant\<close> \<open>Convex set as intersection of halfspaces\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1902
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1903
lemma convex_halfspace_intersection:
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 36844
diff changeset
  1904
  fixes s :: "('a::euclidean_space) set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1905
  assumes "closed s" "convex s"
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60420
diff changeset
  1906
  shows "s = \<Inter>{h. s \<subseteq> h \<and> (\<exists>a b. h = {x. inner a x \<le> b})}"
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1907
  apply (rule set_eqI, rule)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1908
  unfolding Inter_iff Ball_def mem_Collect_eq
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1909
  apply (rule,rule,erule conjE)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1910
proof -
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  1911
  fix x
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1912
  assume "\<forall>xa. s \<subseteq> xa \<and> (\<exists>a b. xa = {x. inner a x \<le> b}) \<longrightarrow> x \<in> xa"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1913
  then have "\<forall>a b. s \<subseteq> {x. inner a x \<le> b} \<longrightarrow> x \<in> {x. inner a x \<le> b}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1914
    by blast
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1915
  then show "x \<in> s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1916
    apply (rule_tac ccontr)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1917
    apply (drule separating_hyperplane_closed_point[OF assms(2,1)])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1918
    apply (erule exE)+
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1919
    apply (erule_tac x="-a" in allE)
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1920
    apply (erule_tac x="-b" in allE, auto)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1921
    done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1922
qed auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1923
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1924
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70097
diff changeset
  1925
subsection\<^marker>\<open>tag unimportant\<close> \<open>Convexity of general and special intervals\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1926
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1927
lemma is_interval_convex:
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1928
  fixes S :: "'a::euclidean_space set"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1929
  assumes "is_interval S"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1930
  shows "convex S"
37732
6432bf0d7191 generalize type of is_interval to class euclidean_space
huffman
parents: 37673
diff changeset
  1931
proof (rule convexI)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  1932
  fix x y and u v :: real
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1933
  assume as: "x \<in> S" "y \<in> S" "0 \<le> u" "0 \<le> v" "u + v = 1"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  1934
  then have *: "u = 1 - v" "1 - v \<ge> 0" and **: "v = 1 - u" "1 - u \<ge> 0"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  1935
    by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  1936
  {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  1937
    fix a b
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  1938
    assume "\<not> b \<le> u * a + v * b"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  1939
    then have "u * a < (1 - v) * b"
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1940
      unfolding not_le using as(4) by (auto simp: field_simps)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  1941
    then have "a < b"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  1942
      unfolding * using as(4) *(2)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  1943
      apply (rule_tac mult_left_less_imp_less[of "1 - v"])
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1944
      apply (auto simp: field_simps)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  1945
      done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  1946
    then have "a \<le> u * a + v * b"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  1947
      unfolding * using as(4)
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1948
      by (auto simp: field_simps intro!:mult_right_mono)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  1949
  }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  1950
  moreover
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  1951
  {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  1952
    fix a b
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  1953
    assume "\<not> u * a + v * b \<le> a"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  1954
    then have "v * b > (1 - u) * a"
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1955
      unfolding not_le using as(4) by (auto simp: field_simps)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  1956
    then have "a < b"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  1957
      unfolding * using as(4)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  1958
      apply (rule_tac mult_left_less_imp_less)
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1959
      apply (auto simp: field_simps)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  1960
      done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  1961
    then have "u * a + v * b \<le> b"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  1962
      unfolding **
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  1963
      using **(2) as(3)
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  1964
      by (auto simp: field_simps intro!:mult_right_mono)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  1965
  }
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1966
  ultimately show "u *\<^sub>R x + v *\<^sub>R y \<in> S"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  1967
    apply -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  1968
    apply (rule assms[unfolded is_interval_def, rule_format, OF as(1,2)])
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  1969
    using as(3-) DIM_positive[where 'a='a]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  1970
    apply (auto simp: inner_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  1971
    done
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50104
diff changeset
  1972
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1973
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1974
lemma is_interval_connected:
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1975
  fixes S :: "'a::euclidean_space set"
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  1976
  shows "is_interval S \<Longrightarrow> connected S"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1977
  using is_interval_convex convex_connected by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1978
62618
f7f2467ab854 Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1979
lemma convex_box [simp]: "convex (cbox a b)" "convex (box a (b::'a::euclidean_space))"
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  1980
  apply (rule_tac[!] is_interval_convex)+
56189
c4daa97ac57a removed dependencies on theory Ordered_Euclidean_Space
immler
parents: 56188
diff changeset
  1981
  using is_interval_box is_interval_cbox
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  1982
  apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  1983
  done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1984
63928
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63918
diff changeset
  1985
text\<open>A non-singleton connected set is perfect (i.e. has no isolated points). \<close>
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63918
diff changeset
  1986
lemma connected_imp_perfect:
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63918
diff changeset
  1987
  fixes a :: "'a::metric_space"
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63918
diff changeset
  1988
  assumes "connected S" "a \<in> S" and S: "\<And>x. S \<noteq> {x}"
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63918
diff changeset
  1989
  shows "a islimpt S"
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63918
diff changeset
  1990
proof -
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63918
diff changeset
  1991
  have False if "a \<in> T" "open T" "\<And>y. \<lbrakk>y \<in> S; y \<in> T\<rbrakk> \<Longrightarrow> y = a" for T
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63918
diff changeset
  1992
  proof -
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63918
diff changeset
  1993
    obtain e where "e > 0" and e: "cball a e \<subseteq> T"
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63918
diff changeset
  1994
      using \<open>open T\<close> \<open>a \<in> T\<close> by (auto simp: open_contains_cball)
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69710
diff changeset
  1995
    have "openin (top_of_set S) {a}"
63928
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63918
diff changeset
  1996
      unfolding openin_open using that \<open>a \<in> S\<close> by blast
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69710
diff changeset
  1997
    moreover have "closedin (top_of_set S) {a}"
63928
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63918
diff changeset
  1998
      by (simp add: assms)
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63918
diff changeset
  1999
    ultimately show "False"
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63918
diff changeset
  2000
      using \<open>connected S\<close> connected_clopen S by blast
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63918
diff changeset
  2001
  qed
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63918
diff changeset
  2002
  then show ?thesis
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63918
diff changeset
  2003
    unfolding islimpt_def by blast
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63918
diff changeset
  2004
qed
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63918
diff changeset
  2005
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63918
diff changeset
  2006
lemma connected_imp_perfect_aff_dim:
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63918
diff changeset
  2007
     "\<lbrakk>connected S; aff_dim S \<noteq> 0; a \<in> S\<rbrakk> \<Longrightarrow> a islimpt S"
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63918
diff changeset
  2008
  using aff_dim_sing connected_imp_perfect by blast
d81fb5b46a5c new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents: 63918
diff changeset
  2009
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70097
diff changeset
  2010
subsection\<^marker>\<open>tag unimportant\<close> \<open>On \<open>real\<close>, \<open>is_interval\<close>, \<open>convex\<close> and \<open>connected\<close> are all equivalent\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2011
67685
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67613
diff changeset
  2012
lemma mem_is_interval_1_I:
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67613
diff changeset
  2013
  fixes a b c::real
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67613
diff changeset
  2014
  assumes "is_interval S"
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67613
diff changeset
  2015
  assumes "a \<in> S" "c \<in> S"
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67613
diff changeset
  2016
  assumes "a \<le> b" "b \<le> c"
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67613
diff changeset
  2017
  shows "b \<in> S"
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67613
diff changeset
  2018
  using assms is_interval_1 by blast
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2019
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2020
lemma is_interval_connected_1:
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2021
  fixes s :: "real set"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2022
  shows "is_interval s \<longleftrightarrow> connected s"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2023
  apply rule
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2024
  apply (rule is_interval_connected, assumption)
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2025
  unfolding is_interval_1
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2026
  apply rule
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2027
  apply rule
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2028
  apply rule
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2029
  apply rule
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2030
  apply (erule conjE)
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64394
diff changeset
  2031
  apply (rule ccontr)       
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2032
proof -
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2033
  fix a b x
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2034
  assume as: "connected s" "a \<in> s" "b \<in> s" "a \<le> x" "x \<le> b" "x \<notin> s"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2035
  then have *: "a < x" "x < b"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2036
    unfolding not_le [symmetric] by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2037
  let ?halfl = "{..<x} "
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2038
  let ?halfr = "{x<..}"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2039
  {
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2040
    fix y
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2041
    assume "y \<in> s"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2042
    with \<open>x \<notin> s\<close> have "x \<noteq> y" by auto
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2043
    then have "y \<in> ?halfr \<union> ?halfl" by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2044
  }
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2045
  moreover have "a \<in> ?halfl" "b \<in> ?halfr" using * by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2046
  then have "?halfl \<inter> s \<noteq> {}" "?halfr \<inter> s \<noteq> {}"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2047
    using as(2-3) by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2048
  ultimately show False
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2049
    apply (rule_tac notE[OF as(1)[unfolded connected_def]])
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2050
    apply (rule_tac x = ?halfl in exI)
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  2051
    apply (rule_tac x = ?halfr in exI, rule)
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  2052
    apply (rule open_lessThan, rule)
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  2053
    apply (rule open_greaterThan, auto)
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2054
    done
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2055
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2056
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2057
lemma is_interval_convex_1:
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2058
  fixes s :: "real set"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2059
  shows "is_interval s \<longleftrightarrow> convex s"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2060
  by (metis is_interval_convex convex_connected is_interval_connected_1)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2061
67685
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67613
diff changeset
  2062
lemma is_interval_ball_real: "is_interval (ball a b)" for a b::real
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67613
diff changeset
  2063
  by (metis connected_ball is_interval_connected_1)
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67613
diff changeset
  2064
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64394
diff changeset
  2065
lemma connected_compact_interval_1:
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64394
diff changeset
  2066
     "connected S \<and> compact S \<longleftrightarrow> (\<exists>a b. S = {a..b::real})"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64394
diff changeset
  2067
  by (auto simp: is_interval_connected_1 [symmetric] is_interval_compact)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64394
diff changeset
  2068
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  2069
lemma connected_convex_1:
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2070
  fixes s :: "real set"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2071
  shows "connected s \<longleftrightarrow> convex s"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2072
  by (metis is_interval_convex convex_connected is_interval_connected_1)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2073
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  2074
lemma connected_convex_1_gen:
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  2075
  fixes s :: "'a :: euclidean_space set"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  2076
  assumes "DIM('a) = 1"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  2077
  shows "connected s \<longleftrightarrow> convex s"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  2078
proof -
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  2079
  obtain f:: "'a \<Rightarrow> real" where linf: "linear f" and "inj f"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  2080
    using subspace_isomorphism[OF subspace_UNIV subspace_UNIV,
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  2081
        where 'a='a and 'b=real]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  2082
    unfolding Euclidean_Space.dim_UNIV
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  2083
    by (auto simp: assms)
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  2084
  then have "f -` (f ` s) = s"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  2085
    by (simp add: inj_vimage_image_eq)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  2086
  then show ?thesis
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  2087
    by (metis connected_convex_1 convex_linear_vimage linf convex_connected connected_linear_image)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  2088
qed
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2089
67685
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67613
diff changeset
  2090
lemma is_interval_cball_1[intro, simp]: "is_interval (cball a b)" for a b::real
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  2091
  by (simp add: is_interval_convex_1)
67685
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67613
diff changeset
  2092
70097
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  2093
lemma [simp]:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  2094
  fixes r s::real
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  2095
  shows is_interval_io: "is_interval {..<r}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  2096
    and is_interval_oi: "is_interval {r<..}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  2097
    and is_interval_oo: "is_interval {r<..<s}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  2098
    and is_interval_oc: "is_interval {r<..s}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  2099
    and is_interval_co: "is_interval {r..<s}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents: 69922
diff changeset
  2100
  by (simp_all add: is_interval_convex_1)
67685
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67613
diff changeset
  2101
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70097
diff changeset
  2102
subsection\<^marker>\<open>tag unimportant\<close> \<open>Another intermediate value theorem formulation\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2103
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2104
lemma ivt_increasing_component_on_1:
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61531
diff changeset
  2105
  fixes f :: "real \<Rightarrow> 'a::euclidean_space"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2106
  assumes "a \<le> b"
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  2107
    and "continuous_on {a..b} f"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2108
    and "(f a)\<bullet>k \<le> y" "y \<le> (f b)\<bullet>k"
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  2109
  shows "\<exists>x\<in>{a..b}. (f x)\<bullet>k = y"
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  2110
proof -
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  2111
  have "f a \<in> f ` cbox a b" "f b \<in> f ` cbox a b"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2112
    apply (rule_tac[!] imageI)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2113
    using assms(1)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2114
    apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2115
    done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2116
  then show ?thesis
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  2117
    using connected_ivt_component[of "f ` cbox a b" "f a" "f b" k y]
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents: 66793
diff changeset
  2118
    by (simp add: connected_continuous_image assms)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2119
qed
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2120
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2121
lemma ivt_increasing_component_1:
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2122
  fixes f :: "real \<Rightarrow> 'a::euclidean_space"
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  2123
  shows "a \<le> b \<Longrightarrow> \<forall>x\<in>{a..b}. continuous (at x) f \<Longrightarrow>
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  2124
    f a\<bullet>k \<le> y \<Longrightarrow> y \<le> f b\<bullet>k \<Longrightarrow> \<exists>x\<in>{a..b}. (f x)\<bullet>k = y"
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  2125
  by (rule ivt_increasing_component_on_1) (auto simp: continuous_at_imp_continuous_on)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2126
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2127
lemma ivt_decreasing_component_on_1:
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2128
  fixes f :: "real \<Rightarrow> 'a::euclidean_space"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2129
  assumes "a \<le> b"
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  2130
    and "continuous_on {a..b} f"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2131
    and "(f b)\<bullet>k \<le> y"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2132
    and "y \<le> (f a)\<bullet>k"
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  2133
  shows "\<exists>x\<in>{a..b}. (f x)\<bullet>k = y"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2134
  apply (subst neg_equal_iff_equal[symmetric])
44531
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44525
diff changeset
  2135
  using ivt_increasing_component_on_1[of a b "\<lambda>x. - f x" k "- y"]
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2136
  using assms using continuous_on_minus
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2137
  apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2138
  done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2139
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2140
lemma ivt_decreasing_component_1:
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2141
  fixes f :: "real \<Rightarrow> 'a::euclidean_space"
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  2142
  shows "a \<le> b \<Longrightarrow> \<forall>x\<in>{a..b}. continuous (at x) f \<Longrightarrow>
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  2143
    f b\<bullet>k \<le> y \<Longrightarrow> y \<le> f a\<bullet>k \<Longrightarrow> \<exists>x\<in>{a..b}. (f x)\<bullet>k = y"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2144
  by (rule ivt_decreasing_component_on_1) (auto simp: continuous_at_imp_continuous_on)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2145
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2146
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70097
diff changeset
  2147
subsection\<^marker>\<open>tag unimportant\<close> \<open>A bound within an interval\<close>
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2148
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  2149
lemma convex_hull_eq_real_cbox:
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2150
  fixes x y :: real assumes "x \<le> y"
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  2151
  shows "convex hull {x, y} = cbox x y"
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2152
proof (rule hull_unique)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2153
  show "{x, y} \<subseteq> cbox x y" using \<open>x \<le> y\<close> by auto
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  2154
  show "convex (cbox x y)"
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  2155
    by (rule convex_box)
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2156
next
68058
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2157
  fix S assume "{x, y} \<subseteq> S" and "convex S"
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2158
  then show "cbox x y \<subseteq> S"
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2159
    unfolding is_interval_convex_1 [symmetric] is_interval_def Basis_real_def
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2160
    by - (clarify, simp (no_asm_use), fast)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2161
qed
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50104
diff changeset
  2162
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2163
lemma unit_interval_convex_hull:
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57418
diff changeset
  2164
  "cbox (0::'a::euclidean_space) One = convex hull {x. \<forall>i\<in>Basis. (x\<bullet>i = 0) \<or> (x\<bullet>i = 1)}"
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 36844
diff changeset
  2165
  (is "?int = convex hull ?points")
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50104
diff changeset
  2166
proof -
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50104
diff changeset
  2167
  have One[simp]: "\<And>i. i \<in> Basis \<Longrightarrow> One \<bullet> i = 1"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
  2168
    by (simp add: inner_sum_left sum.If_cases inner_Basis)
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  2169
  have "?int = {x. \<forall>i\<in>Basis. x \<bullet> i \<in> cbox 0 1}"
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  2170
    by (auto simp: cbox_def)
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  2171
  also have "\<dots> = (\<Sum>i\<in>Basis. (\<lambda>x. x *\<^sub>R i) ` cbox 0 1)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
  2172
    by (simp only: box_eq_set_sum_Basis)
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2173
  also have "\<dots> = (\<Sum>i\<in>Basis. (\<lambda>x. x *\<^sub>R i) ` (convex hull {0, 1}))"
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  2174
    by (simp only: convex_hull_eq_real_cbox zero_le_one)
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2175
  also have "\<dots> = (\<Sum>i\<in>Basis. convex hull ((\<lambda>x. x *\<^sub>R i) ` {0, 1}))"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67982
diff changeset
  2176
    by (simp add: convex_hull_linear_image)
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2177
  also have "\<dots> = convex hull (\<Sum>i\<in>Basis. (\<lambda>x. x *\<^sub>R i) ` {0, 1})"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
  2178
    by (simp only: convex_hull_set_sum)
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2179
  also have "\<dots> = convex hull {x. \<forall>i\<in>Basis. x\<bullet>i \<in> {0, 1}}"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
  2180
    by (simp only: box_eq_set_sum_Basis)
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2181
  also have "convex hull {x. \<forall>i\<in>Basis. x\<bullet>i \<in> {0, 1}} = convex hull ?points"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2182
    by simp
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2183
  finally show ?thesis .
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50104
diff changeset
  2184
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2185
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2186
text \<open>And this is a finite set of vertices.\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2187
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50104
diff changeset
  2188
lemma unit_cube_convex_hull:
68058
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2189
  obtains S :: "'a::euclidean_space set"
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2190
  where "finite S" and "cbox 0 (\<Sum>Basis) = convex hull S"
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2191
proof
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2192
  show "finite {x::'a. \<forall>i\<in>Basis. x \<bullet> i = 0 \<or> x \<bullet> i = 1}"
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2193
  proof (rule finite_subset, clarify)
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2194
    show "finite ((\<lambda>S. \<Sum>i\<in>Basis. (if i \<in> S then 1 else 0) *\<^sub>R i) ` Pow Basis)"
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2195
      using finite_Basis by blast
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2196
    fix x :: 'a
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2197
    assume as: "\<forall>i\<in>Basis. x \<bullet> i = 0 \<or> x \<bullet> i = 1"
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2198
    show "x \<in> (\<lambda>S. \<Sum>i\<in>Basis. (if i\<in>S then 1 else 0) *\<^sub>R i) ` Pow Basis"
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2199
      apply (rule image_eqI[where x="{i. i\<in>Basis \<and> x\<bullet>i = 1}"])
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2200
      using as
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2201
       apply (subst euclidean_eq_iff, auto)
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2202
      done
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2203
  qed
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2204
  show "cbox 0 One = convex hull {x. \<forall>i\<in>Basis. x \<bullet> i = 0 \<or> x \<bullet> i = 1}"
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2205
    using unit_interval_convex_hull by blast
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2206
qed 
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2207
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2208
text \<open>Hence any cube (could do any nonempty interval).\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2209
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2210
lemma cube_convex_hull:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2211
  assumes "d > 0"
68058
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2212
  obtains S :: "'a::euclidean_space set" where
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2213
    "finite S" and "cbox (x - (\<Sum>i\<in>Basis. d*\<^sub>Ri)) (x + (\<Sum>i\<in>Basis. d*\<^sub>Ri)) = convex hull S"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2214
proof -
68058
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2215
  let ?d = "(\<Sum>i\<in>Basis. d *\<^sub>R i)::'a"
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  2216
  have *: "cbox (x - ?d) (x + ?d) = (\<lambda>y. x - ?d + (2 * d) *\<^sub>R y) ` cbox 0 (\<Sum>Basis)"
68058
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2217
  proof (intro set_eqI iffI)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2218
    fix y
68058
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2219
    assume "y \<in> cbox (x - ?d) (x + ?d)"
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  2220
    then have "inverse (2 * d) *\<^sub>R (y - (x - ?d)) \<in> cbox 0 (\<Sum>Basis)"
70802
160eaf566bcb formally augmented corresponding rules for field_simps
haftmann
parents: 70136
diff changeset
  2221
      using assms by (simp add: mem_box inner_simps) (simp add: field_simps)
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68607
diff changeset
  2222
    with \<open>0 < d\<close> show "y \<in> (\<lambda>y. x - sum ((*\<^sub>R) d) Basis + (2 * d) *\<^sub>R y) ` cbox 0 One"
68058
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2223
      by (auto intro: image_eqI[where x= "inverse (2 * d) *\<^sub>R (y - (x - ?d))"])
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2224
  next
68058
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2225
    fix y
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2226
    assume "y \<in> (\<lambda>y. x - ?d + (2 * d) *\<^sub>R y) ` cbox 0 One"
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2227
    then obtain z where z: "z \<in> cbox 0 One" "y = x - ?d + (2*d) *\<^sub>R z"
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  2228
      by auto
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  2229
    then show "y \<in> cbox (x - ?d) (x + ?d)"
68058
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2230
      using z assms by (auto simp: mem_box inner_simps)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2231
  qed
68058
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2232
  obtain S where "finite S" "cbox 0 (\<Sum>Basis::'a) = convex hull S"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2233
    using unit_cube_convex_hull by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2234
  then show ?thesis
68058
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2235
    by (rule_tac that[of "(\<lambda>y. x - ?d + (2 * d) *\<^sub>R y)` S"]) (auto simp: convex_hull_affinity *)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2236
qed
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2237
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70097
diff changeset
  2238
subsection\<^marker>\<open>tag unimportant\<close>\<open>Representation of any interval as a finite convex hull\<close>
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2239
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2240
lemma image_stretch_interval:
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2241
  "(\<lambda>x. \<Sum>k\<in>Basis. (m k * (x\<bullet>k)) *\<^sub>R k) ` cbox a (b::'a::euclidean_space) =
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2242
  (if (cbox a b) = {} then {} else
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2243
    cbox (\<Sum>k\<in>Basis. (min (m k * (a\<bullet>k)) (m k * (b\<bullet>k))) *\<^sub>R k::'a)
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2244
     (\<Sum>k\<in>Basis. (max (m k * (a\<bullet>k)) (m k * (b\<bullet>k))) *\<^sub>R k))"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2245
proof cases
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2246
  assume *: "cbox a b \<noteq> {}"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2247
  show ?thesis
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2248
    unfolding box_ne_empty if_not_P[OF *]
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2249
    apply (simp add: cbox_def image_Collect set_eq_iff euclidean_eq_iff[where 'a='a] ball_conj_distrib[symmetric])
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2250
    apply (subst choice_Basis_iff[symmetric])
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2251
  proof (intro allI ball_cong refl)
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2252
    fix x i :: 'a assume "i \<in> Basis"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2253
    with * have a_le_b: "a \<bullet> i \<le> b \<bullet> i"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2254
      unfolding box_ne_empty by auto
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2255
    show "(\<exists>xa. x \<bullet> i = m i * xa \<and> a \<bullet> i \<le> xa \<and> xa \<le> b \<bullet> i) \<longleftrightarrow>
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2256
        min (m i * (a \<bullet> i)) (m i * (b \<bullet> i)) \<le> x \<bullet> i \<and> x \<bullet> i \<le> max (m i * (a \<bullet> i)) (m i * (b \<bullet> i))"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2257
    proof (cases "m i = 0")
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2258
      case True
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2259
      with a_le_b show ?thesis by auto
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2260
    next
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2261
      case False
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2262
      then have *: "\<And>a b. a = m i * b \<longleftrightarrow> b = a / m i"
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  2263
        by (auto simp: field_simps)
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2264
      from False have
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2265
          "min (m i * (a \<bullet> i)) (m i * (b \<bullet> i)) = (if 0 < m i then m i * (a \<bullet> i) else m i * (b \<bullet> i))"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2266
          "max (m i * (a \<bullet> i)) (m i * (b \<bullet> i)) = (if 0 < m i then m i * (b \<bullet> i) else m i * (a \<bullet> i))"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2267
        using a_le_b by (auto simp: min_def max_def mult_le_cancel_left)
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2268
      with False show ?thesis using a_le_b
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  2269
        unfolding * by (auto simp: le_divide_eq divide_le_eq ac_simps)
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2270
    qed
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2271
  qed
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2272
qed simp
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2273
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2274
lemma interval_image_stretch_interval:
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2275
  "\<exists>u v. (\<lambda>x. \<Sum>k\<in>Basis. (m k * (x\<bullet>k))*\<^sub>R k) ` cbox a (b::'a::euclidean_space) = cbox u (v::'a::euclidean_space)"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2276
  unfolding image_stretch_interval by auto
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2277
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2278
lemma cbox_translation: "cbox (c + a) (c + b) = image (\<lambda>x. c + x) (cbox a b)"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2279
  using image_affinity_cbox [of 1 c a b]
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2280
  using box_ne_empty [of "a+c" "b+c"]  box_ne_empty [of a b]
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  2281
  by (auto simp: inner_left_distrib add.commute)
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2282
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2283
lemma cbox_image_unit_interval:
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2284
  fixes a :: "'a::euclidean_space"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2285
  assumes "cbox a b \<noteq> {}"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2286
    shows "cbox a b =
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67155
diff changeset
  2287
           (+) a ` (\<lambda>x. \<Sum>k\<in>Basis. ((b \<bullet> k - a \<bullet> k) * (x \<bullet> k)) *\<^sub>R k) ` cbox 0 One"
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2288
using assms
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2289
apply (simp add: box_ne_empty image_stretch_interval cbox_translation [symmetric])
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
  2290
apply (simp add: min_def max_def algebra_simps sum_subtractf euclidean_representation)
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2291
done
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2292
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2293
lemma closed_interval_as_convex_hull:
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2294
  fixes a :: "'a::euclidean_space"
68058
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2295
  obtains S where "finite S" "cbox a b = convex hull S"
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2296
proof (cases "cbox a b = {}")
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2297
  case True with convex_hull_empty that show ?thesis
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2298
    by blast
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2299
next
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2300
  case False
68058
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2301
  obtain S::"'a set" where "finite S" and eq: "cbox 0 One = convex hull S"
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2302
    by (blast intro: unit_cube_convex_hull)
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2303
  have lin: "linear (\<lambda>x. \<Sum>k\<in>Basis. ((b \<bullet> k - a \<bullet> k) * (x \<bullet> k)) *\<^sub>R k)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
  2304
    by (rule linear_compose_sum) (auto simp: algebra_simps linearI)
68058
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2305
  have "finite ((+) a ` (\<lambda>x. \<Sum>k\<in>Basis. ((b \<bullet> k - a \<bullet> k) * (x \<bullet> k)) *\<^sub>R k) ` S)"
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2306
    by (rule finite_imageI \<open>finite S\<close>)+
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2307
  then show ?thesis
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2308
    apply (rule that)
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2309
    apply (simp add: convex_hull_translation convex_hull_linear_image [OF lin, symmetric])
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2310
    apply (simp add: eq [symmetric] cbox_image_unit_interval [OF False])
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2311
    done
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2312
qed
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2313
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2314
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70097
diff changeset
  2315
subsection\<^marker>\<open>tag unimportant\<close> \<open>Bounded convex function on open set is continuous\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2316
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2317
lemma convex_on_bounded_continuous:
68058
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2318
  fixes S :: "('a::real_normed_vector) set"
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2319
  assumes "open S"
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2320
    and "convex_on S f"
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2321
    and "\<forall>x\<in>S. \<bar>f x\<bar> \<le> b"
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2322
  shows "continuous_on S f"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2323
  apply (rule continuous_at_imp_continuous_on)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2324
  unfolding continuous_at_real_range
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2325
proof (rule,rule,rule)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2326
  fix x and e :: real
68058
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2327
  assume "x \<in> S" "e > 0"
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  2328
  define B where "B = \<bar>b\<bar> + 1"
68058
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2329
  then have B:  "0 < B""\<And>x. x\<in>S \<Longrightarrow> \<bar>f x\<bar> \<le> B"
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2330
    using assms(3) by auto 
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2331
  obtain k where "k > 0" and k: "cball x k \<subseteq> S"
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2332
    using \<open>x \<in> S\<close> assms(1) open_contains_cball_eq by blast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2333
  show "\<exists>d>0. \<forall>x'. norm (x' - x) < d \<longrightarrow> \<bar>f x' - f x\<bar> < e"
68058
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2334
  proof (intro exI conjI allI impI)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2335
    fix y
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2336
    assume as: "norm (y - x) < min (k / 2) (e / (2 * B) * k)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2337
    show "\<bar>f y - f x\<bar> < e"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2338
    proof (cases "y = x")
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2339
      case False
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  2340
      define t where "t = k / norm (y - x)"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2341
      have "2 < t" "0<t"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2342
        unfolding t_def using as False and \<open>k>0\<close>
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  2343
        by (auto simp:field_simps)
68058
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2344
      have "y \<in> S"
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2345
        apply (rule k[THEN subsetD])
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2346
        unfolding mem_cball dist_norm
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2347
        apply (rule order_trans[of _ "2 * norm (x - y)"])
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2348
        using as
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  2349
        by (auto simp: field_simps norm_minus_commute)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2350
      {
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  2351
        define w where "w = x + t *\<^sub>R (y - x)"
68058
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2352
        have "w \<in> S"
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2353
          using \<open>k>0\<close> by (auto simp: dist_norm t_def w_def k[THEN subsetD])
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2354
        have "(1 / t) *\<^sub>R x + - x + ((t - 1) / t) *\<^sub>R x = (1 / t - 1 + (t - 1) / t) *\<^sub>R x"
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  2355
          by (auto simp: algebra_simps)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2356
        also have "\<dots> = 0"
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  2357
          using \<open>t > 0\<close> by (auto simp:field_simps)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2358
        finally have w: "(1 / t) *\<^sub>R w + ((t - 1) / t) *\<^sub>R x = y"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2359
          unfolding w_def using False and \<open>t > 0\<close>
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  2360
          by (auto simp: algebra_simps)
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  2361
        have 2: "2 * B < e * t"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2362
          unfolding t_def using \<open>0 < e\<close> \<open>0 < k\<close> \<open>B > 0\<close> and as and False
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  2363
          by (auto simp:field_simps)
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  2364
        have "f y - f x \<le> (f w - f x) / t"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2365
          using assms(2)[unfolded convex_on_def,rule_format,of w x "1/t" "(t - 1)/t", unfolded w]
68058
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2366
          using \<open>0 < t\<close> \<open>2 < t\<close> and \<open>x \<in> S\<close> \<open>w \<in> S\<close>
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  2367
          by (auto simp:field_simps)
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  2368
        also have "... < e"
68058
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2369
          using B(2)[OF \<open>w\<in>S\<close>] and B(2)[OF \<open>x\<in>S\<close>] 2 \<open>t > 0\<close> by (auto simp: field_simps)
68052
e98988801fa9 another big cleanup
paulson <lp15@cam.ac.uk>
parents: 68048
diff changeset
  2370
        finally have th1: "f y - f x < e" .
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2371
      }
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  2372
      moreover
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2373
      {
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  2374
        define w where "w = x - t *\<^sub>R (y - x)"
68058
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2375
        have "w \<in> S"
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2376
          using \<open>k > 0\<close> by (auto simp: dist_norm t_def w_def k[THEN subsetD])
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2377
        have "(1 / (1 + t)) *\<^sub>R x + (t / (1 + t)) *\<^sub>R x = (1 / (1 + t) + t / (1 + t)) *\<^sub>R x"
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  2378
          by (auto simp: algebra_simps)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2379
        also have "\<dots> = x"
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  2380
          using \<open>t > 0\<close> by (auto simp:field_simps)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2381
        finally have w: "(1 / (1+t)) *\<^sub>R w + (t / (1 + t)) *\<^sub>R y = x"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2382
          unfolding w_def using False and \<open>t > 0\<close>
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  2383
          by (auto simp: algebra_simps)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2384
        have "2 * B < e * t"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2385
          unfolding t_def
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2386
          using \<open>0 < e\<close> \<open>0 < k\<close> \<open>B > 0\<close> and as and False
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  2387
          by (auto simp:field_simps)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2388
        then have *: "(f w - f y) / t < e"
68058
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2389
          using B(2)[OF \<open>w\<in>S\<close>] and B(2)[OF \<open>y\<in>S\<close>]
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2390
          using \<open>t > 0\<close>
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  2391
          by (auto simp:field_simps)
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  2392
        have "f x \<le> 1 / (1 + t) * f w + (t / (1 + t)) * f y"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2393
          using assms(2)[unfolded convex_on_def,rule_format,of w y "1/(1+t)" "t / (1+t)",unfolded w]
68058
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2394
          using \<open>0 < t\<close> \<open>2 < t\<close> and \<open>y \<in> S\<close> \<open>w \<in> S\<close>
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  2395
          by (auto simp:field_simps)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2396
        also have "\<dots> = (f w + t * f y) / (1 + t)"
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70802
diff changeset
  2397
          using \<open>t > 0\<close> by (simp add: add_divide_distrib) 
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2398
        also have "\<dots> < e + f y"
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  2399
          using \<open>t > 0\<close> * \<open>e > 0\<close> by (auto simp: field_simps)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2400
        finally have "f x - f y < e" by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2401
      }
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  2402
      ultimately show ?thesis by auto
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2403
    qed (insert \<open>0<e\<close>, auto)
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2404
  qed (insert \<open>0<e\<close> \<open>0<k\<close> \<open>0<B\<close>, auto simp: field_simps)
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2405
qed
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2406
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2407
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70097
diff changeset
  2408
subsection\<^marker>\<open>tag unimportant\<close> \<open>Upper bound on a ball implies upper and lower bounds\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2409
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2410
lemma convex_bounds_lemma:
36338
7808fbc9c3b4 generalize more constants and lemmas
huffman
parents: 36337
diff changeset
  2411
  fixes x :: "'a::real_normed_vector"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2412
  assumes "convex_on (cball x e) f"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2413
    and "\<forall>y \<in> cball x e. f y \<le> b"
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61880
diff changeset
  2414
  shows "\<forall>y \<in> cball x e. \<bar>f y\<bar> \<le> b + 2 * \<bar>f x\<bar>"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2415
  apply rule
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2416
proof (cases "0 \<le> e")
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2417
  case True
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2418
  fix y
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2419
  assume y: "y \<in> cball x e"
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  2420
  define z where "z = 2 *\<^sub>R x - y"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2421
  have *: "x - (2 *\<^sub>R x - y) = y - x"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2422
    by (simp add: scaleR_2)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2423
  have z: "z \<in> cball x e"
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  2424
    using y unfolding z_def mem_cball dist_norm * by (auto simp: norm_minus_commute)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2425
  have "(1 / 2) *\<^sub>R y + (1 / 2) *\<^sub>R z = x"
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  2426
    unfolding z_def by (auto simp: algebra_simps)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2427
  then show "\<bar>f y\<bar> \<le> b + 2 * \<bar>f x\<bar>"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2428
    using assms(1)[unfolded convex_on_def,rule_format, OF y z, of "1/2" "1/2"]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2429
    using assms(2)[rule_format,OF y] assms(2)[rule_format,OF z]
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  2430
    by (auto simp:field_simps)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2431
next
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2432
  case False
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2433
  fix y
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2434
  assume "y \<in> cball x e"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2435
  then have "dist x y < 0"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2436
    using False unfolding mem_cball not_le by (auto simp del: dist_not_less_zero)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2437
  then show "\<bar>f y\<bar> \<le> b + 2 * \<bar>f x\<bar>"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2438
    using zero_le_dist[of x y] by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2439
qed
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2440
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2441
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70097
diff changeset
  2442
subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Hence a convex function on an open set is continuous\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2443
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50104
diff changeset
  2444
lemma real_of_nat_ge_one_iff: "1 \<le> real (n::nat) \<longleftrightarrow> 1 \<le> n"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50104
diff changeset
  2445
  by auto
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50104
diff changeset
  2446
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2447
lemma convex_on_continuous:
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2448
  assumes "open (s::('a::euclidean_space) set)" "convex_on s f"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2449
  shows "continuous_on s f"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2450
  unfolding continuous_on_eq_continuous_at[OF assms(1)]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2451
proof
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 36844
diff changeset
  2452
  note dimge1 = DIM_positive[where 'a='a]
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2453
  fix x
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2454
  assume "x \<in> s"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2455
  then obtain e where e: "cball x e \<subseteq> s" "e > 0"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2456
    using assms(1) unfolding open_contains_cball by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  2457
  define d where "d = e / real DIM('a)"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2458
  have "0 < d"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2459
    unfolding d_def using \<open>e > 0\<close> dimge1 by auto
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50104
diff changeset
  2460
  let ?d = "(\<Sum>i\<in>Basis. d *\<^sub>R i)::'a"
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2461
  obtain c
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2462
    where c: "finite c"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2463
    and c1: "convex hull c \<subseteq> cball x e"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2464
    and c2: "cball x d \<subseteq> convex hull c"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2465
  proof
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  2466
    define c where "c = (\<Sum>i\<in>Basis. (\<lambda>a. a *\<^sub>R i) ` {x\<bullet>i - d, x\<bullet>i + d})"
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2467
    show "finite c"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
  2468
      unfolding c_def by (simp add: finite_set_sum)
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  2469
    have 1: "convex hull c = {a. \<forall>i\<in>Basis. a \<bullet> i \<in> cbox (x \<bullet> i - d) (x \<bullet> i + d)}"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
  2470
      unfolding box_eq_set_sum_Basis
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
  2471
      unfolding c_def convex_hull_set_sum
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2472
      apply (subst convex_hull_linear_image [symmetric])
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2473
      apply (simp add: linear_iff scaleR_add_left)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
  2474
      apply (rule sum.cong [OF refl])
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2475
      apply (rule image_cong [OF _ refl])
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  2476
      apply (rule convex_hull_eq_real_cbox)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2477
      apply (cut_tac \<open>0 < d\<close>, simp)
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2478
      done
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2479
    then have 2: "convex hull c = {a. \<forall>i\<in>Basis. a \<bullet> i \<in> cball (x \<bullet> i) d}"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2480
      by (simp add: dist_norm abs_le_iff algebra_simps)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2481
    show "cball x d \<subseteq> convex hull c"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2482
      unfolding 2
68058
69715dfdc286 more general tidying up
paulson <lp15@cam.ac.uk>
parents: 68052
diff changeset
  2483
      by (clarsimp simp: dist_norm) (metis inner_commute inner_diff_right norm_bound_Basis_le)
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2484
    have e': "e = (\<Sum>(i::'a)\<in>Basis. d)"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61531
diff changeset
  2485
      by (simp add: d_def DIM_positive)
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2486
    show "convex hull c \<subseteq> cball x e"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2487
      unfolding 2
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2488
      apply clarsimp
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2489
      apply (subst euclidean_dist_l2)
67155
9e5b05d54f9d canonical name
nipkow
parents: 67135
diff changeset
  2490
      apply (rule order_trans [OF L2_set_le_sum])
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2491
      apply (rule zero_le_dist)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2492
      unfolding e'
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  2493
      apply (rule sum_mono, simp)
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2494
      done
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2495
  qed
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  2496
  define k where "k = Max (f ` c)"
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2497
  have "convex_on (convex hull c) f"
50526
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50104
diff changeset
  2498
    apply(rule convex_on_subset[OF assms(2)])
68069
36209dfb981e tidying up and using real induction methods
paulson <lp15@cam.ac.uk>
parents: 68058
diff changeset
  2499
    apply(rule subset_trans[OF c1 e(1)])
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2500
    done
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2501
  then have k: "\<forall>y\<in>convex hull c. f y \<le> k"
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  2502
    apply (rule_tac convex_on_convex_hull_bound, assumption)
68048
0b4fb9fd91b1 more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68041
diff changeset
  2503
    by (simp add: k_def c)
0b4fb9fd91b1 more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68041
diff changeset
  2504
  have "e \<le> e * real DIM('a)"
0b4fb9fd91b1 more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68041
diff changeset
  2505
    using e(2) real_of_nat_ge_one_iff by auto
0b4fb9fd91b1 more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68041
diff changeset
  2506
  then have "d \<le> e"
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70802
diff changeset
  2507
    by (simp add: d_def field_split_simps)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2508
  then have dsube: "cball x d \<subseteq> cball x e"
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  2509
    by (rule subset_cball)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2510
  have conv: "convex_on (cball x d) f"
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  2511
    using \<open>convex_on (convex hull c) f\<close> c2 convex_on_subset by blast
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61880
diff changeset
  2512
  then have "\<forall>y\<in>cball x d. \<bar>f y\<bar> \<le> k + 2 * \<bar>f x\<bar>"
68048
0b4fb9fd91b1 more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68041
diff changeset
  2513
    by (rule convex_bounds_lemma) (use c2 k in blast)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2514
  then have "continuous_on (ball x d) f"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2515
    apply (rule_tac convex_on_bounded_continuous)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2516
    apply (rule open_ball, rule convex_on_subset[OF conv])
68031
eda52f4cd4e4 fixing more messy proofs
paulson <lp15@cam.ac.uk>
parents: 68024
diff changeset
  2517
    apply (rule ball_subset_cball, force)
33270
paulson
parents: 33175
diff changeset
  2518
    done
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2519
  then show "continuous (at x) f"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  2520
    unfolding continuous_on_eq_continuous_at[OF open_ball]
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2521
    using \<open>d > 0\<close> by auto
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2522
qed
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2523
71008
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2524
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2525
section \<open>Line Segments\<close>
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2526
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2527
subsection \<open>Midpoint\<close>
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2528
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2529
definition\<^marker>\<open>tag important\<close> midpoint :: "'a::real_vector \<Rightarrow> 'a \<Rightarrow> 'a"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2530
  where "midpoint a b = (inverse (2::real)) *\<^sub>R (a + b)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2531
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2532
lemma midpoint_idem [simp]: "midpoint x x = x"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2533
  unfolding midpoint_def  by simp
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2534
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2535
lemma midpoint_sym: "midpoint a b = midpoint b a"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2536
  unfolding midpoint_def by (auto simp add: scaleR_right_distrib)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2537
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2538
lemma midpoint_eq_iff: "midpoint a b = c \<longleftrightarrow> a + b = c + c"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2539
proof -
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2540
  have "midpoint a b = c \<longleftrightarrow> scaleR 2 (midpoint a b) = scaleR 2 c"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2541
    by simp
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2542
  then show ?thesis
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2543
    unfolding midpoint_def scaleR_2 [symmetric] by simp
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2544
qed
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2545
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2546
lemma
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2547
  fixes a::real
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2548
  assumes "a \<le> b" shows ge_midpoint_1: "a \<le> midpoint a b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2549
                    and le_midpoint_1: "midpoint a b \<le> b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2550
  by (simp_all add: midpoint_def assms)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2551
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2552
lemma dist_midpoint:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2553
  fixes a b :: "'a::real_normed_vector" shows
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2554
  "dist a (midpoint a b) = (dist a b) / 2" (is ?t1)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2555
  "dist b (midpoint a b) = (dist a b) / 2" (is ?t2)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2556
  "dist (midpoint a b) a = (dist a b) / 2" (is ?t3)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2557
  "dist (midpoint a b) b = (dist a b) / 2" (is ?t4)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2558
proof -
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2559
  have *: "\<And>x y::'a. 2 *\<^sub>R x = - y \<Longrightarrow> norm x = (norm y) / 2"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2560
    unfolding equation_minus_iff by auto
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2561
  have **: "\<And>x y::'a. 2 *\<^sub>R x =   y \<Longrightarrow> norm x = (norm y) / 2"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2562
    by auto
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2563
  note scaleR_right_distrib [simp]
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2564
  show ?t1
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2565
    unfolding midpoint_def dist_norm
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2566
    apply (rule **)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2567
    apply (simp add: scaleR_right_diff_distrib)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2568
    apply (simp add: scaleR_2)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2569
    done
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2570
  show ?t2
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2571
    unfolding midpoint_def dist_norm
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2572
    apply (rule *)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2573
    apply (simp add: scaleR_right_diff_distrib)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2574
    apply (simp add: scaleR_2)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2575
    done
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2576
  show ?t3
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2577
    unfolding midpoint_def dist_norm
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2578
    apply (rule *)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2579
    apply (simp add: scaleR_right_diff_distrib)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2580
    apply (simp add: scaleR_2)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2581
    done
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2582
  show ?t4
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2583
    unfolding midpoint_def dist_norm
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2584
    apply (rule **)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2585
    apply (simp add: scaleR_right_diff_distrib)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2586
    apply (simp add: scaleR_2)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2587
    done
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2588
qed
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2589
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2590
lemma midpoint_eq_endpoint [simp]:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2591
  "midpoint a b = a \<longleftrightarrow> a = b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2592
  "midpoint a b = b \<longleftrightarrow> a = b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2593
  unfolding midpoint_eq_iff by auto
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2594
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2595
lemma midpoint_plus_self [simp]: "midpoint a b + midpoint a b = a + b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2596
  using midpoint_eq_iff by metis
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2597
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2598
lemma midpoint_linear_image:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2599
   "linear f \<Longrightarrow> midpoint(f a)(f b) = f(midpoint a b)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2600
by (simp add: linear_iff midpoint_def)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2601
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2602
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2603
subsection \<open>Line segments\<close>
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2604
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2605
definition\<^marker>\<open>tag important\<close> closed_segment :: "'a::real_vector \<Rightarrow> 'a \<Rightarrow> 'a set"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2606
  where "closed_segment a b = {(1 - u) *\<^sub>R a + u *\<^sub>R b | u::real. 0 \<le> u \<and> u \<le> 1}"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2607
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2608
definition\<^marker>\<open>tag important\<close> open_segment :: "'a::real_vector \<Rightarrow> 'a \<Rightarrow> 'a set" where
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2609
  "open_segment a b \<equiv> closed_segment a b - {a,b}"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2610
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2611
lemmas segment = open_segment_def closed_segment_def
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2612
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2613
lemma in_segment:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2614
    "x \<in> closed_segment a b \<longleftrightarrow> (\<exists>u. 0 \<le> u \<and> u \<le> 1 \<and> x = (1 - u) *\<^sub>R a + u *\<^sub>R b)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2615
    "x \<in> open_segment a b \<longleftrightarrow> a \<noteq> b \<and> (\<exists>u. 0 < u \<and> u < 1 \<and> x = (1 - u) *\<^sub>R a + u *\<^sub>R b)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2616
  using less_eq_real_def by (auto simp: segment algebra_simps)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2617
71029
934e0044e94b Moved or deleted some out of place material, also eliminating obsolete naming conventions
paulson <lp15@cam.ac.uk>
parents: 71008
diff changeset
  2618
lemma closed_segmentI:
934e0044e94b Moved or deleted some out of place material, also eliminating obsolete naming conventions
paulson <lp15@cam.ac.uk>
parents: 71008
diff changeset
  2619
  "u \<in> {0..1} \<Longrightarrow> z = (1 - u) *\<^sub>R a + u *\<^sub>R b \<Longrightarrow> z \<in> closed_segment a b"
934e0044e94b Moved or deleted some out of place material, also eliminating obsolete naming conventions
paulson <lp15@cam.ac.uk>
parents: 71008
diff changeset
  2620
  by (auto simp: closed_segment_def)
934e0044e94b Moved or deleted some out of place material, also eliminating obsolete naming conventions
paulson <lp15@cam.ac.uk>
parents: 71008
diff changeset
  2621
71008
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2622
lemma closed_segment_linear_image:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2623
  "closed_segment (f a) (f b) = f ` (closed_segment a b)" if "linear f"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2624
proof -
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2625
  interpret linear f by fact
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2626
  show ?thesis
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2627
    by (force simp add: in_segment add scale)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2628
qed
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2629
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2630
lemma open_segment_linear_image:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2631
    "\<lbrakk>linear f; inj f\<rbrakk> \<Longrightarrow> open_segment (f a) (f b) = f ` (open_segment a b)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2632
  by (force simp: open_segment_def closed_segment_linear_image inj_on_def)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2633
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2634
lemma closed_segment_translation:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2635
    "closed_segment (c + a) (c + b) = image (\<lambda>x. c + x) (closed_segment a b)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2636
apply safe
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2637
apply (rule_tac x="x-c" in image_eqI)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2638
apply (auto simp: in_segment algebra_simps)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2639
done
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2640
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2641
lemma open_segment_translation:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2642
    "open_segment (c + a) (c + b) = image (\<lambda>x. c + x) (open_segment a b)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2643
by (simp add: open_segment_def closed_segment_translation translation_diff)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2644
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2645
lemma closed_segment_of_real:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2646
    "closed_segment (of_real x) (of_real y) = of_real ` closed_segment x y"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2647
  apply (auto simp: image_iff in_segment scaleR_conv_of_real)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2648
    apply (rule_tac x="(1-u)*x + u*y" in bexI)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2649
  apply (auto simp: in_segment)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2650
  done
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2651
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2652
lemma open_segment_of_real:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2653
    "open_segment (of_real x) (of_real y) = of_real ` open_segment x y"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2654
  apply (auto simp: image_iff in_segment scaleR_conv_of_real)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2655
    apply (rule_tac x="(1-u)*x + u*y" in bexI)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2656
  apply (auto simp: in_segment)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2657
  done
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2658
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2659
lemma closed_segment_Reals:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2660
    "\<lbrakk>x \<in> Reals; y \<in> Reals\<rbrakk> \<Longrightarrow> closed_segment x y = of_real ` closed_segment (Re x) (Re y)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2661
  by (metis closed_segment_of_real of_real_Re)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2662
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2663
lemma open_segment_Reals:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2664
    "\<lbrakk>x \<in> Reals; y \<in> Reals\<rbrakk> \<Longrightarrow> open_segment x y = of_real ` open_segment (Re x) (Re y)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2665
  by (metis open_segment_of_real of_real_Re)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2666
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2667
lemma open_segment_PairD:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2668
    "(x, x') \<in> open_segment (a, a') (b, b')
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2669
     \<Longrightarrow> (x \<in> open_segment a b \<or> a = b) \<and> (x' \<in> open_segment a' b' \<or> a' = b')"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2670
  by (auto simp: in_segment)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2671
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2672
lemma closed_segment_PairD:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2673
  "(x, x') \<in> closed_segment (a, a') (b, b') \<Longrightarrow> x \<in> closed_segment a b \<and> x' \<in> closed_segment a' b'"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2674
  by (auto simp: closed_segment_def)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2675
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2676
lemma closed_segment_translation_eq [simp]:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2677
    "d + x \<in> closed_segment (d + a) (d + b) \<longleftrightarrow> x \<in> closed_segment a b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2678
proof -
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2679
  have *: "\<And>d x a b. x \<in> closed_segment a b \<Longrightarrow> d + x \<in> closed_segment (d + a) (d + b)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2680
    apply (simp add: closed_segment_def)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2681
    apply (erule ex_forward)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2682
    apply (simp add: algebra_simps)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2683
    done
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2684
  show ?thesis
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2685
  using * [where d = "-d"] *
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2686
  by (fastforce simp add:)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2687
qed
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2688
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2689
lemma open_segment_translation_eq [simp]:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2690
    "d + x \<in> open_segment (d + a) (d + b) \<longleftrightarrow> x \<in> open_segment a b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2691
  by (simp add: open_segment_def)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2692
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2693
lemma of_real_closed_segment [simp]:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2694
  "of_real x \<in> closed_segment (of_real a) (of_real b) \<longleftrightarrow> x \<in> closed_segment a b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2695
  apply (auto simp: in_segment scaleR_conv_of_real elim!: ex_forward)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2696
  using of_real_eq_iff by fastforce
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2697
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2698
lemma of_real_open_segment [simp]:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2699
  "of_real x \<in> open_segment (of_real a) (of_real b) \<longleftrightarrow> x \<in> open_segment a b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2700
  apply (auto simp: in_segment scaleR_conv_of_real elim!: ex_forward del: exE)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2701
  using of_real_eq_iff by fastforce
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2702
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2703
lemma convex_contains_segment:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2704
  "convex S \<longleftrightarrow> (\<forall>a\<in>S. \<forall>b\<in>S. closed_segment a b \<subseteq> S)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2705
  unfolding convex_alt closed_segment_def by auto
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2706
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2707
lemma closed_segment_in_Reals:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2708
   "\<lbrakk>x \<in> closed_segment a b; a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> x \<in> Reals"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2709
  by (meson subsetD convex_Reals convex_contains_segment)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2710
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2711
lemma open_segment_in_Reals:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2712
   "\<lbrakk>x \<in> open_segment a b; a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> x \<in> Reals"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2713
  by (metis Diff_iff closed_segment_in_Reals open_segment_def)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2714
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2715
lemma closed_segment_subset: "\<lbrakk>x \<in> S; y \<in> S; convex S\<rbrakk> \<Longrightarrow> closed_segment x y \<subseteq> S"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2716
  by (simp add: convex_contains_segment)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2717
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2718
lemma closed_segment_subset_convex_hull:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2719
    "\<lbrakk>x \<in> convex hull S; y \<in> convex hull S\<rbrakk> \<Longrightarrow> closed_segment x y \<subseteq> convex hull S"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2720
  using convex_contains_segment by blast
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2721
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2722
lemma segment_convex_hull:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2723
  "closed_segment a b = convex hull {a,b}"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2724
proof -
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2725
  have *: "\<And>x. {x} \<noteq> {}" by auto
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2726
  show ?thesis
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2727
    unfolding segment convex_hull_insert[OF *] convex_hull_singleton
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2728
    by (safe; rule_tac x="1 - u" in exI; force)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2729
qed
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2730
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2731
lemma open_closed_segment: "u \<in> open_segment w z \<Longrightarrow> u \<in> closed_segment w z"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2732
  by (auto simp add: closed_segment_def open_segment_def)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2733
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2734
lemma segment_open_subset_closed:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2735
   "open_segment a b \<subseteq> closed_segment a b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2736
  by (auto simp: closed_segment_def open_segment_def)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2737
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2738
lemma bounded_closed_segment:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2739
    fixes a :: "'a::euclidean_space" shows "bounded (closed_segment a b)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2740
  by (simp add: segment_convex_hull compact_convex_hull compact_imp_bounded)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2741
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2742
lemma bounded_open_segment:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2743
    fixes a :: "'a::euclidean_space" shows "bounded (open_segment a b)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2744
  by (rule bounded_subset [OF bounded_closed_segment segment_open_subset_closed])
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2745
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2746
lemmas bounded_segment = bounded_closed_segment open_closed_segment
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2747
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2748
lemma ends_in_segment [iff]: "a \<in> closed_segment a b" "b \<in> closed_segment a b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2749
  unfolding segment_convex_hull
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2750
  by (auto intro!: hull_subset[unfolded subset_eq, rule_format])
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2751
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2752
lemma eventually_closed_segment:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2753
  fixes x0::"'a::real_normed_vector"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2754
  assumes "open X0" "x0 \<in> X0"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2755
  shows "\<forall>\<^sub>F x in at x0 within U. closed_segment x0 x \<subseteq> X0"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2756
proof -
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2757
  from openE[OF assms]
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2758
  obtain e where e: "0 < e" "ball x0 e \<subseteq> X0" .
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2759
  then have "\<forall>\<^sub>F x in at x0 within U. x \<in> ball x0 e"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2760
    by (auto simp: dist_commute eventually_at)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2761
  then show ?thesis
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2762
  proof eventually_elim
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2763
    case (elim x)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2764
    have "x0 \<in> ball x0 e" using \<open>e > 0\<close> by simp
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2765
    from convex_ball[unfolded convex_contains_segment, rule_format, OF this elim]
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2766
    have "closed_segment x0 x \<subseteq> ball x0 e" .
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2767
    also note \<open>\<dots> \<subseteq> X0\<close>
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2768
    finally show ?case .
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2769
  qed
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2770
qed
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2771
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2772
lemma segment_furthest_le:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2773
  fixes a b x y :: "'a::euclidean_space"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2774
  assumes "x \<in> closed_segment a b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2775
  shows "norm (y - x) \<le> norm (y - a) \<or>  norm (y - x) \<le> norm (y - b)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2776
proof -
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2777
  obtain z where "z \<in> {a, b}" "norm (x - y) \<le> norm (z - y)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2778
    using simplex_furthest_le[of "{a, b}" y]
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2779
    using assms[unfolded segment_convex_hull]
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2780
    by auto
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2781
  then show ?thesis
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2782
    by (auto simp add:norm_minus_commute)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2783
qed
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2784
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2785
lemma closed_segment_commute: "closed_segment a b = closed_segment b a"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2786
proof -
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2787
  have "{a, b} = {b, a}" by auto
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2788
  thus ?thesis
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2789
    by (simp add: segment_convex_hull)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2790
qed
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2791
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2792
lemma segment_bound1:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2793
  assumes "x \<in> closed_segment a b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2794
  shows "norm (x - a) \<le> norm (b - a)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2795
proof -
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2796
  obtain u where "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 \<le> u" "u \<le> 1"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2797
    using assms by (auto simp add: closed_segment_def)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2798
  then show "norm (x - a) \<le> norm (b - a)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2799
    apply clarify
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2800
    apply (auto simp: algebra_simps)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2801
    apply (simp add: scaleR_diff_right [symmetric] mult_left_le_one_le)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2802
    done
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2803
qed
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2804
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2805
lemma segment_bound:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2806
  assumes "x \<in> closed_segment a b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2807
  shows "norm (x - a) \<le> norm (b - a)" "norm (x - b) \<le> norm (b - a)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2808
apply (simp add: assms segment_bound1)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2809
by (metis assms closed_segment_commute dist_commute dist_norm segment_bound1)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2810
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2811
lemma open_segment_commute: "open_segment a b = open_segment b a"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2812
proof -
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2813
  have "{a, b} = {b, a}" by auto
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2814
  thus ?thesis
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2815
    by (simp add: closed_segment_commute open_segment_def)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2816
qed
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2817
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2818
lemma closed_segment_idem [simp]: "closed_segment a a = {a}"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2819
  unfolding segment by (auto simp add: algebra_simps)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2820
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2821
lemma open_segment_idem [simp]: "open_segment a a = {}"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2822
  by (simp add: open_segment_def)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2823
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2824
lemma closed_segment_eq_open: "closed_segment a b = open_segment a b \<union> {a,b}"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2825
  using open_segment_def by auto
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2826
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2827
lemma convex_contains_open_segment:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2828
  "convex s \<longleftrightarrow> (\<forall>a\<in>s. \<forall>b\<in>s. open_segment a b \<subseteq> s)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2829
  by (simp add: convex_contains_segment closed_segment_eq_open)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2830
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2831
lemma closed_segment_eq_real_ivl:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2832
  fixes a b::real
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2833
  shows "closed_segment a b = (if a \<le> b then {a .. b} else {b .. a})"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2834
proof -
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2835
  have "b \<le> a \<Longrightarrow> closed_segment b a = {b .. a}"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2836
    and "a \<le> b \<Longrightarrow> closed_segment a b = {a .. b}"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2837
    by (auto simp: convex_hull_eq_real_cbox segment_convex_hull)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2838
  thus ?thesis
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2839
    by (auto simp: closed_segment_commute)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2840
qed
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2841
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2842
lemma open_segment_eq_real_ivl:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2843
  fixes a b::real
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2844
  shows "open_segment a b = (if a \<le> b then {a<..<b} else {b<..<a})"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2845
by (auto simp: closed_segment_eq_real_ivl open_segment_def split: if_split_asm)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2846
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2847
lemma closed_segment_real_eq:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2848
  fixes u::real shows "closed_segment u v = (\<lambda>x. (v - u) * x + u) ` {0..1}"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2849
  by (simp add: add.commute [of u] image_affinity_atLeastAtMost [where c=u] closed_segment_eq_real_ivl)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2850
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2851
lemma dist_in_closed_segment:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2852
  fixes a :: "'a :: euclidean_space"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2853
  assumes "x \<in> closed_segment a b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2854
    shows "dist x a \<le> dist a b \<and> dist x b \<le> dist a b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2855
proof (intro conjI)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2856
  obtain u where u: "0 \<le> u" "u \<le> 1" and x: "x = (1 - u) *\<^sub>R a + u *\<^sub>R b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2857
    using assms by (force simp: in_segment algebra_simps)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2858
  have "dist x a = u * dist a b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2859
    apply (simp add: dist_norm algebra_simps x)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2860
    by (metis \<open>0 \<le> u\<close> abs_of_nonneg norm_minus_commute norm_scaleR real_vector.scale_right_diff_distrib)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2861
  also have "...  \<le> dist a b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2862
    by (simp add: mult_left_le_one_le u)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2863
  finally show "dist x a \<le> dist a b" .
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2864
  have "dist x b = norm ((1-u) *\<^sub>R a - (1-u) *\<^sub>R b)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2865
    by (simp add: dist_norm algebra_simps x)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2866
  also have "... = (1-u) * dist a b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2867
  proof -
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2868
    have "norm ((1 - 1 * u) *\<^sub>R (a - b)) = (1 - 1 * u) * norm (a - b)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2869
      using \<open>u \<le> 1\<close> by force
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2870
    then show ?thesis
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2871
      by (simp add: dist_norm real_vector.scale_right_diff_distrib)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2872
  qed
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2873
  also have "... \<le> dist a b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2874
    by (simp add: mult_left_le_one_le u)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2875
  finally show "dist x b \<le> dist a b" .
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2876
qed
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2877
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2878
lemma dist_in_open_segment:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2879
  fixes a :: "'a :: euclidean_space"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2880
  assumes "x \<in> open_segment a b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2881
    shows "dist x a < dist a b \<and> dist x b < dist a b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2882
proof (intro conjI)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2883
  obtain u where u: "0 < u" "u < 1" and x: "x = (1 - u) *\<^sub>R a + u *\<^sub>R b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2884
    using assms by (force simp: in_segment algebra_simps)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2885
  have "dist x a = u * dist a b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2886
    apply (simp add: dist_norm algebra_simps x)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2887
    by (metis abs_of_nonneg less_eq_real_def norm_minus_commute norm_scaleR real_vector.scale_right_diff_distrib \<open>0 < u\<close>)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2888
  also have *: "...  < dist a b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2889
    by (metis (no_types) assms dist_eq_0_iff dist_not_less_zero in_segment(2) linorder_neqE_linordered_idom mult.left_neutral real_mult_less_iff1 \<open>u < 1\<close>)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2890
  finally show "dist x a < dist a b" .
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2891
  have ab_ne0: "dist a b \<noteq> 0"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2892
    using * by fastforce
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2893
  have "dist x b = norm ((1-u) *\<^sub>R a - (1-u) *\<^sub>R b)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2894
    by (simp add: dist_norm algebra_simps x)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2895
  also have "... = (1-u) * dist a b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2896
  proof -
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2897
    have "norm ((1 - 1 * u) *\<^sub>R (a - b)) = (1 - 1 * u) * norm (a - b)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2898
      using \<open>u < 1\<close> by force
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2899
    then show ?thesis
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2900
      by (simp add: dist_norm real_vector.scale_right_diff_distrib)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2901
  qed
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2902
  also have "... < dist a b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2903
    using ab_ne0 \<open>0 < u\<close> by simp
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2904
  finally show "dist x b < dist a b" .
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2905
qed
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2906
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2907
lemma dist_decreases_open_segment_0:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2908
  fixes x :: "'a :: euclidean_space"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2909
  assumes "x \<in> open_segment 0 b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2910
    shows "dist c x < dist c 0 \<or> dist c x < dist c b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2911
proof (rule ccontr, clarsimp simp: not_less)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2912
  obtain u where u: "0 \<noteq> b" "0 < u" "u < 1" and x: "x = u *\<^sub>R b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2913
    using assms by (auto simp: in_segment)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2914
  have xb: "x \<bullet> b < b \<bullet> b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2915
    using u x by auto
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2916
  assume "norm c \<le> dist c x"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2917
  then have "c \<bullet> c \<le> (c - x) \<bullet> (c - x)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2918
    by (simp add: dist_norm norm_le)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2919
  moreover have "0 < x \<bullet> b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2920
    using u x by auto
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2921
  ultimately have less: "c \<bullet> b < x \<bullet> b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2922
    by (simp add: x algebra_simps inner_commute u)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2923
  assume "dist c b \<le> dist c x"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2924
  then have "(c - b) \<bullet> (c - b) \<le> (c - x) \<bullet> (c - x)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2925
    by (simp add: dist_norm norm_le)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2926
  then have "(b \<bullet> b) * (1 - u*u) \<le> 2 * (b \<bullet> c) * (1-u)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2927
    by (simp add: x algebra_simps inner_commute)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2928
  then have "(1+u) * (b \<bullet> b) * (1-u) \<le> 2 * (b \<bullet> c) * (1-u)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2929
    by (simp add: algebra_simps)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2930
  then have "(1+u) * (b \<bullet> b) \<le> 2 * (b \<bullet> c)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2931
    using \<open>u < 1\<close> by auto
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2932
  with xb have "c \<bullet> b \<ge> x \<bullet> b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2933
    by (auto simp: x algebra_simps inner_commute)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2934
  with less show False by auto
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2935
qed
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2936
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2937
proposition dist_decreases_open_segment:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2938
  fixes a :: "'a :: euclidean_space"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2939
  assumes "x \<in> open_segment a b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2940
    shows "dist c x < dist c a \<or> dist c x < dist c b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2941
proof -
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2942
  have *: "x - a \<in> open_segment 0 (b - a)" using assms
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2943
    by (metis diff_self open_segment_translation_eq uminus_add_conv_diff)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2944
  show ?thesis
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2945
    using dist_decreases_open_segment_0 [OF *, of "c-a"] assms
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2946
    by (simp add: dist_norm)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2947
qed
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2948
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2949
corollary open_segment_furthest_le:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2950
  fixes a b x y :: "'a::euclidean_space"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2951
  assumes "x \<in> open_segment a b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2952
  shows "norm (y - x) < norm (y - a) \<or>  norm (y - x) < norm (y - b)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2953
  by (metis assms dist_decreases_open_segment dist_norm)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2954
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2955
corollary dist_decreases_closed_segment:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2956
  fixes a :: "'a :: euclidean_space"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2957
  assumes "x \<in> closed_segment a b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2958
    shows "dist c x \<le> dist c a \<or> dist c x \<le> dist c b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2959
apply (cases "x \<in> open_segment a b")
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2960
 using dist_decreases_open_segment less_eq_real_def apply blast
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2961
by (metis DiffI assms empty_iff insertE open_segment_def order_refl)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2962
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2963
lemma convex_intermediate_ball:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2964
  fixes a :: "'a :: euclidean_space"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2965
  shows "\<lbrakk>ball a r \<subseteq> T; T \<subseteq> cball a r\<rbrakk> \<Longrightarrow> convex T"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2966
apply (simp add: convex_contains_open_segment, clarify)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2967
by (metis (no_types, hide_lams) less_le_trans mem_ball mem_cball subsetCE dist_decreases_open_segment)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2968
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2969
lemma csegment_midpoint_subset: "closed_segment (midpoint a b) b \<subseteq> closed_segment a b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2970
  apply (clarsimp simp: midpoint_def in_segment)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2971
  apply (rule_tac x="(1 + u) / 2" in exI)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2972
  apply (auto simp: algebra_simps add_divide_distrib diff_divide_distrib)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2973
  by (metis field_sum_of_halves scaleR_left.add)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2974
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2975
lemma notin_segment_midpoint:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2976
  fixes a :: "'a :: euclidean_space"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2977
  shows "a \<noteq> b \<Longrightarrow> a \<notin> closed_segment (midpoint a b) b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2978
by (auto simp: dist_midpoint dest!: dist_in_closed_segment)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2979
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2980
lemma segment_to_closest_point:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2981
  fixes S :: "'a :: euclidean_space set"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2982
  shows "\<lbrakk>closed S; S \<noteq> {}\<rbrakk> \<Longrightarrow> open_segment a (closest_point S a) \<inter> S = {}"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2983
  apply (subst disjoint_iff_not_equal)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2984
  apply (clarify dest!: dist_in_open_segment)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2985
  by (metis closest_point_le dist_commute le_less_trans less_irrefl)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2986
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2987
lemma segment_to_point_exists:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2988
  fixes S :: "'a :: euclidean_space set"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2989
    assumes "closed S" "S \<noteq> {}"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2990
    obtains b where "b \<in> S" "open_segment a b \<inter> S = {}"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2991
  by (metis assms segment_to_closest_point closest_point_exists that)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2992
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2993
subsubsection\<open>More lemmas, especially for working with the underlying formula\<close>
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2994
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2995
lemma segment_eq_compose:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2996
  fixes a :: "'a :: real_vector"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2997
  shows "(\<lambda>u. (1 - u) *\<^sub>R a + u *\<^sub>R b) = (\<lambda>x. a + x) o (\<lambda>u. u *\<^sub>R (b - a))"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2998
    by (simp add: o_def algebra_simps)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  2999
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3000
lemma segment_degen_1:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3001
  fixes a :: "'a :: real_vector"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3002
  shows "(1 - u) *\<^sub>R a + u *\<^sub>R b = b \<longleftrightarrow> a=b \<or> u=1"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3003
proof -
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3004
  { assume "(1 - u) *\<^sub>R a + u *\<^sub>R b = b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3005
    then have "(1 - u) *\<^sub>R a = (1 - u) *\<^sub>R b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3006
      by (simp add: algebra_simps)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3007
    then have "a=b \<or> u=1"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3008
      by simp
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3009
  } then show ?thesis
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3010
      by (auto simp: algebra_simps)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3011
qed
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3012
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3013
lemma segment_degen_0:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3014
    fixes a :: "'a :: real_vector"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3015
    shows "(1 - u) *\<^sub>R a + u *\<^sub>R b = a \<longleftrightarrow> a=b \<or> u=0"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3016
  using segment_degen_1 [of "1-u" b a]
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3017
  by (auto simp: algebra_simps)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3018
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3019
lemma add_scaleR_degen:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3020
  fixes a b ::"'a::real_vector"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3021
  assumes  "(u *\<^sub>R b + v *\<^sub>R a) = (u *\<^sub>R a + v *\<^sub>R b)"  "u \<noteq> v"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3022
  shows "a=b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3023
  by (metis (no_types, hide_lams) add.commute add_diff_eq diff_add_cancel real_vector.scale_cancel_left real_vector.scale_left_diff_distrib assms)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3024
  
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3025
lemma closed_segment_image_interval:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3026
     "closed_segment a b = (\<lambda>u. (1 - u) *\<^sub>R a + u *\<^sub>R b) ` {0..1}"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3027
  by (auto simp: set_eq_iff image_iff closed_segment_def)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3028
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3029
lemma open_segment_image_interval:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3030
     "open_segment a b = (if a=b then {} else (\<lambda>u. (1 - u) *\<^sub>R a + u *\<^sub>R b) ` {0<..<1})"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3031
  by (auto simp:  open_segment_def closed_segment_def segment_degen_0 segment_degen_1)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3032
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3033
lemmas segment_image_interval = closed_segment_image_interval open_segment_image_interval
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3034
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3035
lemma open_segment_bound1:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3036
  assumes "x \<in> open_segment a b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3037
  shows "norm (x - a) < norm (b - a)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3038
proof -
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3039
  obtain u where "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 < u" "u < 1" "a \<noteq> b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3040
    using assms by (auto simp add: open_segment_image_interval split: if_split_asm)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3041
  then show "norm (x - a) < norm (b - a)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3042
    apply clarify
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3043
    apply (auto simp: algebra_simps)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3044
    apply (simp add: scaleR_diff_right [symmetric])
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3045
    done
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3046
qed
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3047
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3048
lemma compact_segment [simp]:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3049
  fixes a :: "'a::real_normed_vector"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3050
  shows "compact (closed_segment a b)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3051
  by (auto simp: segment_image_interval intro!: compact_continuous_image continuous_intros)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3052
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3053
lemma closed_segment [simp]:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3054
  fixes a :: "'a::real_normed_vector"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3055
  shows "closed (closed_segment a b)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3056
  by (simp add: compact_imp_closed)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3057
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3058
lemma closure_closed_segment [simp]:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3059
  fixes a :: "'a::real_normed_vector"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3060
  shows "closure(closed_segment a b) = closed_segment a b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3061
  by simp
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3062
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3063
lemma open_segment_bound:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3064
  assumes "x \<in> open_segment a b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3065
  shows "norm (x - a) < norm (b - a)" "norm (x - b) < norm (b - a)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3066
apply (simp add: assms open_segment_bound1)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3067
by (metis assms norm_minus_commute open_segment_bound1 open_segment_commute)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3068
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3069
lemma closure_open_segment [simp]:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3070
  "closure (open_segment a b) = (if a = b then {} else closed_segment a b)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3071
    for a :: "'a::euclidean_space"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3072
proof (cases "a = b")
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3073
  case True
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3074
  then show ?thesis
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3075
    by simp
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3076
next
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3077
  case False
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3078
  have "closure ((\<lambda>u. u *\<^sub>R (b - a)) ` {0<..<1}) = (\<lambda>u. u *\<^sub>R (b - a)) ` closure {0<..<1}"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3079
    apply (rule closure_injective_linear_image [symmetric])
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3080
     apply (use False in \<open>auto intro!: injI\<close>)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3081
    done
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3082
  then have "closure
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3083
     ((\<lambda>u. (1 - u) *\<^sub>R a + u *\<^sub>R b) ` {0<..<1}) =
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3084
    (\<lambda>x. (1 - x) *\<^sub>R a + x *\<^sub>R b) ` closure {0<..<1}"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3085
    using closure_translation [of a "((\<lambda>x. x *\<^sub>R b - x *\<^sub>R a) ` {0<..<1})"]
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3086
    by (simp add: segment_eq_compose field_simps scaleR_diff_left scaleR_diff_right image_image)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3087
  then show ?thesis
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3088
    by (simp add: segment_image_interval closure_greaterThanLessThan [symmetric] del: closure_greaterThanLessThan)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3089
qed
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3090
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3091
lemma closed_open_segment_iff [simp]:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3092
    fixes a :: "'a::euclidean_space"  shows "closed(open_segment a b) \<longleftrightarrow> a = b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3093
  by (metis open_segment_def DiffE closure_eq closure_open_segment ends_in_segment(1) insert_iff segment_image_interval(2))
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3094
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3095
lemma compact_open_segment_iff [simp]:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3096
    fixes a :: "'a::euclidean_space"  shows "compact(open_segment a b) \<longleftrightarrow> a = b"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3097
  by (simp add: bounded_open_segment compact_eq_bounded_closed)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3098
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3099
lemma convex_closed_segment [iff]: "convex (closed_segment a b)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3100
  unfolding segment_convex_hull by(rule convex_convex_hull)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3101
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3102
lemma convex_open_segment [iff]: "convex (open_segment a b)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3103
proof -
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3104
  have "convex ((\<lambda>u. u *\<^sub>R (b - a)) ` {0<..<1})"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3105
    by (rule convex_linear_image) auto
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3106
  then have "convex ((+) a ` (\<lambda>u. u *\<^sub>R (b - a)) ` {0<..<1})"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3107
    by (rule convex_translation)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3108
  then show ?thesis
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3109
    by (simp add: image_image open_segment_image_interval segment_eq_compose field_simps scaleR_diff_left scaleR_diff_right)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3110
qed
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3111
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3112
lemmas convex_segment = convex_closed_segment convex_open_segment
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3113
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3114
lemma connected_segment [iff]:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3115
  fixes x :: "'a :: real_normed_vector"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3116
  shows "connected (closed_segment x y)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3117
  by (simp add: convex_connected)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3118
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3119
lemma is_interval_closed_segment_1[intro, simp]: "is_interval (closed_segment a b)" for a b::real
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3120
  by (auto simp: is_interval_convex_1)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3121
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3122
lemma IVT'_closed_segment_real:
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3123
  fixes f :: "real \<Rightarrow> real"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3124
  assumes "y \<in> closed_segment (f a) (f b)"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3125
  assumes "continuous_on (closed_segment a b) f"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3126
  shows "\<exists>x \<in> closed_segment a b. f x = y"
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3127
  using IVT'[of f a y b]
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3128
    IVT'[of "-f" a "-y" b]
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3129
    IVT'[of f b y a]
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3130
    IVT'[of "-f" b "-y" a] assms
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3131
  by (cases "a \<le> b"; cases "f b \<ge> f a") (auto simp: closed_segment_eq_real_ivl continuous_on_minus)
e892f7a1fd83 moved line segments to Convex_Euclidean_Space
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  3132
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3133
end