src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy
author wenzelm
Thu, 28 Apr 2016 11:47:01 +0200
changeset 63067 0a8a75e400da
parent 63040 eb4ddd18d635
child 63072 eb5d493a9e03
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
41959
b460124855b8 tuned headers;
wenzelm
parents: 41413
diff changeset
     1
(*  Title:      HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
     2
    Author:     Robert Himmelmann, TU Muenchen
40887
ee8d0548c148 Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents: 40719
diff changeset
     3
    Author:     Bogdan Grechuk, University of Edinburgh
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
     4
*)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
     5
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
     6
section \<open>Convex sets, functions and related things.\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
     7
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
     8
theory Convex_Euclidean_Space
44132
0f35a870ecf1 full import paths
huffman
parents: 44125
diff changeset
     9
imports
56189
c4daa97ac57a removed dependencies on theory Ordered_Euclidean_Space
immler
parents: 56188
diff changeset
    10
  Topology_Euclidean_Space
44132
0f35a870ecf1 full import paths
huffman
parents: 44125
diff changeset
    11
  "~~/src/HOL/Library/Convex"
0f35a870ecf1 full import paths
huffman
parents: 44125
diff changeset
    12
  "~~/src/HOL/Library/Set_Algebras"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
    13
begin
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
    14
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
    15
lemma independent_injective_on_span_image:
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
    16
  assumes iS: "independent S"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
    17
    and lf: "linear f"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
    18
    and fi: "inj_on f (span S)"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
    19
  shows "independent (f ` S)"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
    20
proof -
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
    21
  {
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
    22
    fix a
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
    23
    assume a: "a \<in> S" "f a \<in> span (f ` S - {f a})"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
    24
    have eq: "f ` S - {f a} = f ` (S - {a})"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
    25
      using fi a span_inc by (auto simp add: inj_on_def)
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
    26
    from a have "f a \<in> f ` span (S -{a})"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
    27
      unfolding eq span_linear_image [OF lf, of "S - {a}"] by blast
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
    28
    moreover have "span (S - {a}) \<subseteq> span S"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
    29
      using span_mono[of "S - {a}" S] by auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
    30
    ultimately have "a \<in> span (S - {a})"
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
    31
      using fi a span_inc by (auto simp add: inj_on_def)
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
    32
    with a(1) iS have False
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
    33
      by (simp add: dependent_def)
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
    34
  }
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
    35
  then show ?thesis
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
    36
    unfolding dependent_def by blast
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
    37
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
    38
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
    39
lemma dim_image_eq:
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
    40
  fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
    41
  assumes lf: "linear f"
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
    42
    and fi: "inj_on f (span S)"
53406
d4374a69ddff tuned proofs;
wenzelm
parents: 53374
diff changeset
    43
  shows "dim (f ` S) = dim (S::'n::euclidean_space set)"
d4374a69ddff tuned proofs;
wenzelm
parents: 53374
diff changeset
    44
proof -
d4374a69ddff tuned proofs;
wenzelm
parents: 53374
diff changeset
    45
  obtain B where B: "B \<subseteq> S" "independent B" "S \<subseteq> span B" "card B = dim S"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
    46
    using basis_exists[of S] by auto
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
    47
  then have "span S = span B"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
    48
    using span_mono[of B S] span_mono[of S "span B"] span_span[of B] by auto
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
    49
  then have "independent (f ` B)"
53406
d4374a69ddff tuned proofs;
wenzelm
parents: 53374
diff changeset
    50
    using independent_injective_on_span_image[of B f] B assms by auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
    51
  moreover have "card (f ` B) = card B"
53406
d4374a69ddff tuned proofs;
wenzelm
parents: 53374
diff changeset
    52
    using assms card_image[of f B] subset_inj_on[of f "span S" B] B span_inc by auto
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
    53
  moreover have "(f ` B) \<subseteq> (f ` S)"
53406
d4374a69ddff tuned proofs;
wenzelm
parents: 53374
diff changeset
    54
    using B by auto
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
    55
  ultimately have "dim (f ` S) \<ge> dim S"
53406
d4374a69ddff tuned proofs;
wenzelm
parents: 53374
diff changeset
    56
    using independent_card_le_dim[of "f ` B" "f ` S"] B by auto
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
    57
  then show ?thesis
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
    58
    using dim_image_le[of f S] assms by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
    59
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
    60
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
    61
lemma linear_injective_on_subspace_0:
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
    62
  assumes lf: "linear f"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
    63
    and "subspace S"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
    64
  shows "inj_on f S \<longleftrightarrow> (\<forall>x \<in> S. f x = 0 \<longrightarrow> x = 0)"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
    65
proof -
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
    66
  have "inj_on f S \<longleftrightarrow> (\<forall>x \<in> S. \<forall>y \<in> S. f x = f y \<longrightarrow> x = y)"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
    67
    by (simp add: inj_on_def)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
    68
  also have "\<dots> \<longleftrightarrow> (\<forall>x \<in> S. \<forall>y \<in> S. f x - f y = 0 \<longrightarrow> x - y = 0)"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
    69
    by simp
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
    70
  also have "\<dots> \<longleftrightarrow> (\<forall>x \<in> S. \<forall>y \<in> S. f (x - y) = 0 \<longrightarrow> x - y = 0)"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
    71
    by (simp add: linear_sub[OF lf])
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
    72
  also have "\<dots> \<longleftrightarrow> (\<forall>x \<in> S. f x = 0 \<longrightarrow> x = 0)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
    73
    using \<open>subspace S\<close> subspace_def[of S] subspace_sub[of S] by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
    74
  finally show ?thesis .
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
    75
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
    76
61952
546958347e05 prefer symbols for "Union", "Inter";
wenzelm
parents: 61945
diff changeset
    77
lemma subspace_Inter: "\<forall>s \<in> f. subspace s \<Longrightarrow> subspace (\<Inter>f)"
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
    78
  unfolding subspace_def by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
    79
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
    80
lemma span_eq[simp]: "span s = s \<longleftrightarrow> subspace s"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
    81
  unfolding span_def by (rule hull_eq) (rule subspace_Inter)
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
    82
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
    83
lemma substdbasis_expansion_unique:
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
    84
  assumes d: "d \<subseteq> Basis"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
    85
  shows "(\<Sum>i\<in>d. f i *\<^sub>R i) = (x::'a::euclidean_space) \<longleftrightarrow>
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
    86
    (\<forall>i\<in>Basis. (i \<in> d \<longrightarrow> f i = x \<bullet> i) \<and> (i \<notin> d \<longrightarrow> x \<bullet> i = 0))"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
    87
proof -
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
    88
  have *: "\<And>x a b P. x * (if P then a else b) = (if P then x * a else x * b)"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
    89
    by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
    90
  have **: "finite d"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
    91
    by (auto intro: finite_subset[OF assms])
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
    92
  have ***: "\<And>i. i \<in> Basis \<Longrightarrow> (\<Sum>i\<in>d. f i *\<^sub>R i) \<bullet> i = (\<Sum>x\<in>d. if x = i then f x else 0)"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50104
diff changeset
    93
    using d
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
    94
    by (auto intro!: setsum.cong simp: inner_Basis inner_setsum_left)
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
    95
  show ?thesis
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
    96
    unfolding euclidean_eq_iff[where 'a='a] by (auto simp: setsum.delta[OF **] ***)
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
    97
qed
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50104
diff changeset
    98
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50104
diff changeset
    99
lemma independent_substdbasis: "d \<subseteq> Basis \<Longrightarrow> independent d"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50104
diff changeset
   100
  by (rule independent_mono[OF independent_Basis])
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   101
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
   102
lemma dim_cball:
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   103
  assumes "e > 0"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   104
  shows "dim (cball (0 :: 'n::euclidean_space) e) = DIM('n)"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   105
proof -
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   106
  {
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   107
    fix x :: "'n::euclidean_space"
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
   108
    define y where "y = (e / norm x) *\<^sub>R x"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   109
    then have "y \<in> cball 0 e"
62397
5ae24f33d343 Substantial new material for multivariate analysis. Also removal of some duplicates.
paulson <lp15@cam.ac.uk>
parents: 62381
diff changeset
   110
      using assms by auto
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   111
    moreover have *: "x = (norm x / e) *\<^sub>R y"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   112
      using y_def assms by simp
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   113
    moreover from * have "x = (norm x/e) *\<^sub>R y"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   114
      by auto
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   115
    ultimately have "x \<in> span (cball 0 e)"
62397
5ae24f33d343 Substantial new material for multivariate analysis. Also removal of some duplicates.
paulson <lp15@cam.ac.uk>
parents: 62381
diff changeset
   116
      using span_mul[of y "cball 0 e" "norm x/e"] span_inc[of "cball 0 e"]
5ae24f33d343 Substantial new material for multivariate analysis. Also removal of some duplicates.
paulson <lp15@cam.ac.uk>
parents: 62381
diff changeset
   117
      by (simp add: span_superset)
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   118
  }
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   119
  then have "span (cball 0 e) = (UNIV :: 'n::euclidean_space set)"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   120
    by auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   121
  then show ?thesis
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   122
    using dim_span[of "cball (0 :: 'n::euclidean_space) e"] by (auto simp add: dim_UNIV)
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   123
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   124
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   125
lemma indep_card_eq_dim_span:
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   126
  fixes B :: "'n::euclidean_space set"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   127
  assumes "independent B"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   128
  shows "finite B \<and> card B = dim (span B)"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   129
  using assms basis_card_eq_dim[of B "span B"] span_inc by auto
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   130
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   131
lemma setsum_not_0: "setsum f A \<noteq> 0 \<Longrightarrow> \<exists>a \<in> A. f a \<noteq> 0"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   132
  by (rule ccontr) auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   133
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61609
diff changeset
   134
lemma subset_translation_eq [simp]:
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61609
diff changeset
   135
    fixes a :: "'a::real_vector" shows "op + a ` s \<subseteq> op + a ` t \<longleftrightarrow> s \<subseteq> t"
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61609
diff changeset
   136
  by auto
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61609
diff changeset
   137
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
   138
lemma translate_inj_on:
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   139
  fixes A :: "'a::ab_group_add set"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   140
  shows "inj_on (\<lambda>x. a + x) A"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   141
  unfolding inj_on_def by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   142
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   143
lemma translation_assoc:
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   144
  fixes a b :: "'a::ab_group_add"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   145
  shows "(\<lambda>x. b + x) ` ((\<lambda>x. a + x) ` S) = (\<lambda>x. (a + b) + x) ` S"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   146
  by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   147
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   148
lemma translation_invert:
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   149
  fixes a :: "'a::ab_group_add"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   150
  assumes "(\<lambda>x. a + x) ` A = (\<lambda>x. a + x) ` B"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   151
  shows "A = B"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   152
proof -
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   153
  have "(\<lambda>x. -a + x) ` ((\<lambda>x. a + x) ` A) = (\<lambda>x. - a + x) ` ((\<lambda>x. a + x) ` B)"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   154
    using assms by auto
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   155
  then show ?thesis
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   156
    using translation_assoc[of "-a" a A] translation_assoc[of "-a" a B] by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   157
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   158
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   159
lemma translation_galois:
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   160
  fixes a :: "'a::ab_group_add"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   161
  shows "T = ((\<lambda>x. a + x) ` S) \<longleftrightarrow> S = ((\<lambda>x. (- a) + x) ` T)"
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   162
  using translation_assoc[of "-a" a S]
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   163
  apply auto
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   164
  using translation_assoc[of a "-a" T]
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   165
  apply auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   166
  done
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   167
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61609
diff changeset
   168
lemma convex_translation_eq [simp]: "convex ((\<lambda>x. a + x) ` s) \<longleftrightarrow> convex s"
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61609
diff changeset
   169
  by (metis convex_translation translation_galois)
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61609
diff changeset
   170
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   171
lemma translation_inverse_subset:
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   172
  assumes "((\<lambda>x. - a + x) ` V) \<le> (S :: 'n::ab_group_add set)"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   173
  shows "V \<le> ((\<lambda>x. a + x) ` S)"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   174
proof -
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   175
  {
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   176
    fix x
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   177
    assume "x \<in> V"
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   178
    then have "x-a \<in> S" using assms by auto
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   179
    then have "x \<in> {a + v |v. v \<in> S}"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   180
      apply auto
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   181
      apply (rule exI[of _ "x-a"])
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   182
      apply simp
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   183
      done
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   184
    then have "x \<in> ((\<lambda>x. a+x) ` S)" by auto
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   185
  }
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   186
  then show ?thesis by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   187
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   188
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61609
diff changeset
   189
lemma convex_linear_image_eq [simp]:
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61609
diff changeset
   190
    fixes f :: "'a::real_vector \<Rightarrow> 'b::real_vector"
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61609
diff changeset
   191
    shows "\<lbrakk>linear f; inj f\<rbrakk> \<Longrightarrow> convex (f ` s) \<longleftrightarrow> convex s"
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61609
diff changeset
   192
    by (metis (no_types) convex_linear_image convex_linear_vimage inj_vimage_image_eq)
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61609
diff changeset
   193
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   194
lemma basis_to_basis_subspace_isomorphism:
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   195
  assumes s: "subspace (S:: ('n::euclidean_space) set)"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   196
    and t: "subspace (T :: ('m::euclidean_space) set)"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   197
    and d: "dim S = dim T"
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   198
    and B: "B \<subseteq> S" "independent B" "S \<subseteq> span B" "card B = dim S"
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   199
    and C: "C \<subseteq> T" "independent C" "T \<subseteq> span C" "card C = dim T"
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   200
  shows "\<exists>f. linear f \<and> f ` B = C \<and> f ` S = T \<and> inj_on f S"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   201
proof -
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   202
  from B independent_bound have fB: "finite B"
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   203
    by blast
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   204
  from C independent_bound have fC: "finite C"
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   205
    by blast
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   206
  from B(4) C(4) card_le_inj[of B C] d obtain f where
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
   207
    f: "f ` B \<subseteq> C" "inj_on f B" using \<open>finite B\<close> \<open>finite C\<close> by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   208
  from linear_independent_extend[OF B(2)] obtain g where
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   209
    g: "linear g" "\<forall>x \<in> B. g x = f x" by blast
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   210
  from inj_on_iff_eq_card[OF fB, of f] f(2)
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   211
  have "card (f ` B) = card B" by simp
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   212
  with B(4) C(4) have ceq: "card (f ` B) = card C" using d
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   213
    by simp
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   214
  have "g ` B = f ` B" using g(2)
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   215
    by (auto simp add: image_iff)
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   216
  also have "\<dots> = C" using card_subset_eq[OF fC f(1) ceq] .
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   217
  finally have gBC: "g ` B = C" .
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   218
  have gi: "inj_on g B" using f(2) g(2)
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   219
    by (auto simp add: inj_on_def)
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   220
  note g0 = linear_indep_image_lemma[OF g(1) fB, unfolded gBC, OF C(2) gi]
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   221
  {
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   222
    fix x y
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   223
    assume x: "x \<in> S" and y: "y \<in> S" and gxy: "g x = g y"
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   224
    from B(3) x y have x': "x \<in> span B" and y': "y \<in> span B"
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   225
      by blast+
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   226
    from gxy have th0: "g (x - y) = 0"
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   227
      by (simp add: linear_sub[OF g(1)])
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   228
    have th1: "x - y \<in> span B" using x' y'
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   229
      by (metis span_sub)
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   230
    have "x = y" using g0[OF th1 th0] by simp
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   231
  }
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   232
  then have giS: "inj_on g S" unfolding inj_on_def by blast
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   233
  from span_subspace[OF B(1,3) s]
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   234
  have "g ` S = span (g ` B)"
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   235
    by (simp add: span_linear_image[OF g(1)])
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   236
  also have "\<dots> = span C"
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   237
    unfolding gBC ..
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   238
  also have "\<dots> = T"
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   239
    using span_subspace[OF C(1,3) t] .
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   240
  finally have gS: "g ` S = T" .
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   241
  from g(1) gS giS gBC show ?thesis
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   242
    by blast
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   243
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   244
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   245
lemma closure_bounded_linear_image_subset:
44524
04ad69081646 generalize some lemmas
huffman
parents: 44523
diff changeset
   246
  assumes f: "bounded_linear f"
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   247
  shows "f ` closure S \<subseteq> closure (f ` S)"
44524
04ad69081646 generalize some lemmas
huffman
parents: 44523
diff changeset
   248
  using linear_continuous_on [OF f] closed_closure closure_subset
04ad69081646 generalize some lemmas
huffman
parents: 44523
diff changeset
   249
  by (rule image_closure_subset)
04ad69081646 generalize some lemmas
huffman
parents: 44523
diff changeset
   250
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   251
lemma closure_linear_image_subset:
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   252
  fixes f :: "'m::euclidean_space \<Rightarrow> 'n::real_normed_vector"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   253
  assumes "linear f"
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   254
  shows "f ` (closure S) \<subseteq> closure (f ` S)"
44524
04ad69081646 generalize some lemmas
huffman
parents: 44523
diff changeset
   255
  using assms unfolding linear_conv_bounded_linear
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   256
  by (rule closure_bounded_linear_image_subset)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   257
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   258
lemma closed_injective_linear_image:
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   259
    fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   260
    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
   261
    shows "closed (f ` S)"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   262
proof -
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   263
  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
   264
    using linear_injective_left_inverse [OF f] by blast
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   265
  then have confg: "continuous_on (range f) g"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   266
    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
   267
  have [simp]: "g ` f ` S = S"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   268
    using g by (simp add: image_comp)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   269
  have cgf: "closed (g ` f ` S)"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61762
diff changeset
   270
    by (simp add: \<open>g \<circ> f = id\<close> S image_comp)
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   271
  have [simp]: "{x \<in> range f. g x \<in> S} = f ` S"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   272
    using g by (simp add: o_def id_def image_def) metis
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   273
  show ?thesis
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   274
    apply (rule closedin_closed_trans [of "range f"])
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   275
    apply (rule continuous_closedin_preimage [OF confg cgf, simplified])
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   276
    apply (rule closed_injective_image_subspace)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   277
    using f
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   278
    apply (auto simp: linear_linear linear_injective_0)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   279
    done
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   280
qed
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   281
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   282
lemma closed_injective_linear_image_eq:
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   283
    fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   284
    assumes f: "linear f" "inj f"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   285
      shows "(closed(image f s) \<longleftrightarrow> closed s)"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   286
  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
   287
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   288
lemma closure_injective_linear_image:
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   289
    fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   290
    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
   291
  apply (rule subset_antisym)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   292
  apply (simp add: closure_linear_image_subset)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   293
  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
   294
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   295
lemma closure_bounded_linear_image:
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   296
    fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   297
    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
   298
  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
   299
  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
   300
  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
   301
44524
04ad69081646 generalize some lemmas
huffman
parents: 44523
diff changeset
   302
lemma closure_scaleR:
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   303
  fixes S :: "'a::real_normed_vector set"
44524
04ad69081646 generalize some lemmas
huffman
parents: 44523
diff changeset
   304
  shows "(op *\<^sub>R c) ` (closure S) = closure ((op *\<^sub>R c) ` S)"
04ad69081646 generalize some lemmas
huffman
parents: 44523
diff changeset
   305
proof
04ad69081646 generalize some lemmas
huffman
parents: 44523
diff changeset
   306
  show "(op *\<^sub>R c) ` (closure S) \<subseteq> closure ((op *\<^sub>R c) ` S)"
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   307
    using bounded_linear_scaleR_right
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   308
    by (rule closure_bounded_linear_image_subset)
44524
04ad69081646 generalize some lemmas
huffman
parents: 44523
diff changeset
   309
  show "closure ((op *\<^sub>R c) ` S) \<subseteq> (op *\<^sub>R c) ` (closure S)"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   310
    by (intro closure_minimal image_mono closure_subset closed_scaling closed_closure)
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   311
qed
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   312
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   313
lemma fst_linear: "linear fst"
53600
8fda7ad57466 make 'linear' into a sublocale of 'bounded_linear';
huffman
parents: 53406
diff changeset
   314
  unfolding linear_iff by (simp add: algebra_simps)
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   315
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   316
lemma snd_linear: "linear snd"
53600
8fda7ad57466 make 'linear' into a sublocale of 'bounded_linear';
huffman
parents: 53406
diff changeset
   317
  unfolding linear_iff by (simp add: algebra_simps)
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   318
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
   319
lemma fst_snd_linear: "linear (\<lambda>(x,y). x + y)"
53600
8fda7ad57466 make 'linear' into a sublocale of 'bounded_linear';
huffman
parents: 53406
diff changeset
   320
  unfolding linear_iff by (simp add: algebra_simps)
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   321
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 36844
diff changeset
   322
lemma scaleR_2:
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 36844
diff changeset
   323
  fixes x :: "'a::real_vector"
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 36844
diff changeset
   324
  shows "scaleR 2 x = x + x"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   325
  unfolding one_add_one [symmetric] scaleR_left_distrib by simp
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   326
62618
f7f2467ab854 Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   327
lemma scaleR_half_double [simp]:
f7f2467ab854 Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   328
  fixes a :: "'a::real_normed_vector"
f7f2467ab854 Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   329
  shows "(1 / 2) *\<^sub>R (a + a) = a"
f7f2467ab854 Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   330
proof -
f7f2467ab854 Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   331
  have "\<And>r. r *\<^sub>R (a + a) = (r * 2) *\<^sub>R a"
f7f2467ab854 Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   332
    by (metis scaleR_2 scaleR_scaleR)
f7f2467ab854 Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   333
  then show ?thesis
f7f2467ab854 Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   334
    by simp
f7f2467ab854 Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   335
qed
f7f2467ab854 Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   336
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   337
lemma vector_choose_size:
62381
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
   338
  assumes "0 \<le> c"
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
   339
  obtains x :: "'a::{real_normed_vector, perfect_space}" where "norm x = c"
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
   340
proof -
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
   341
  obtain a::'a where "a \<noteq> 0"
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
   342
    using UNIV_not_singleton UNIV_eq_I set_zero singletonI by fastforce
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
   343
  then show ?thesis
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
   344
    by (rule_tac x="scaleR (c / norm a) a" in that) (simp add: assms)
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
   345
qed
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
   346
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
   347
lemma vector_choose_dist:
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
   348
  assumes "0 \<le> c"
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
   349
  obtains y :: "'a::{real_normed_vector, perfect_space}" where "dist x y = c"
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
   350
by (metis add_diff_cancel_left' assms dist_commute dist_norm vector_choose_size)
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
   351
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
   352
lemma sphere_eq_empty [simp]:
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
   353
  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
   354
  shows "sphere a r = {} \<longleftrightarrow> r < 0"
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
   355
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
   356
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   357
lemma setsum_delta_notmem:
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   358
  assumes "x \<notin> s"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   359
  shows "setsum (\<lambda>y. if (y = x) then P x else Q y) s = setsum Q s"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   360
    and "setsum (\<lambda>y. if (x = y) then P x else Q y) s = setsum Q s"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   361
    and "setsum (\<lambda>y. if (y = x) then P y else Q y) s = setsum Q s"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   362
    and "setsum (\<lambda>y. if (x = y) then P y else Q y) s = setsum Q s"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
   363
  apply (rule_tac [!] setsum.cong)
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   364
  using assms
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   365
  apply auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   366
  done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   367
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   368
lemma setsum_delta'':
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   369
  fixes s::"'a::real_vector set"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   370
  assumes "finite s"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   371
  shows "(\<Sum>x\<in>s. (if y = x then f x else 0) *\<^sub>R x) = (if y\<in>s then (f y) *\<^sub>R y else 0)"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   372
proof -
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   373
  have *: "\<And>x y. (if y = x then f x else (0::real)) *\<^sub>R x = (if x=y then (f x) *\<^sub>R x else 0)"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   374
    by auto
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   375
  show ?thesis
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
   376
    unfolding * using setsum.delta[OF assms, of y "\<lambda>x. f x *\<^sub>R x"] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   377
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   378
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   379
lemma if_smult: "(if P then x else (y::real)) *\<^sub>R v = (if P then x *\<^sub>R v else y *\<^sub>R v)"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
   380
  by (fact if_distrib)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   381
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   382
lemma dist_triangle_eq:
44361
75ec83d45303 remove unnecessary euclidean_space class constraints
huffman
parents: 44349
diff changeset
   383
  fixes x y z :: "'a::real_inner"
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   384
  shows "dist x z = dist x y + dist y z \<longleftrightarrow>
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   385
    norm (x - y) *\<^sub>R (y - z) = norm (y - z) *\<^sub>R (x - y)"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   386
proof -
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   387
  have *: "x - y + (y - z) = x - z" by auto
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 36844
diff changeset
   388
  show ?thesis unfolding dist_norm norm_triangle_eq[of "x - y" "y - z", unfolded *]
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   389
    by (auto simp add:norm_minus_commute)
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   390
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   391
53406
d4374a69ddff tuned proofs;
wenzelm
parents: 53374
diff changeset
   392
lemma norm_minus_eqI: "x = - y \<Longrightarrow> norm x = norm y" by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   393
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   394
lemma Min_grI:
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   395
  assumes "finite A" "A \<noteq> {}" "\<forall>a\<in>A. x < a"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   396
  shows "x < Min A"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   397
  unfolding Min_gr_iff[OF assms(1,2)] using assms(3) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   398
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 36844
diff changeset
   399
lemma norm_lt: "norm x < norm y \<longleftrightarrow> inner x x < inner y y"
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 36844
diff changeset
   400
  unfolding norm_eq_sqrt_inner by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   401
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 36844
diff changeset
   402
lemma norm_le: "norm x \<le> norm y \<longleftrightarrow> inner x x \<le> inner y y"
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 36844
diff changeset
   403
  unfolding norm_eq_sqrt_inner by simp
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 36844
diff changeset
   404
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 36844
diff changeset
   405
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
   406
subsection \<open>Affine set and affine hull\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   407
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   408
definition affine :: "'a::real_vector set \<Rightarrow> bool"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   409
  where "affine s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> s)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   410
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   411
lemma affine_alt: "affine s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. \<forall>u::real. (1 - u) *\<^sub>R x + u *\<^sub>R y \<in> s)"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   412
  unfolding affine_def by (metis eq_diff_eq')
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   413
62948
7700f467892b lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 62843
diff changeset
   414
lemma affine_empty [iff]: "affine {}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   415
  unfolding affine_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   416
62948
7700f467892b lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 62843
diff changeset
   417
lemma affine_sing [iff]: "affine {x}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   418
  unfolding affine_alt by (auto simp add: scaleR_left_distrib [symmetric])
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   419
62948
7700f467892b lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 62843
diff changeset
   420
lemma affine_UNIV [iff]: "affine UNIV"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   421
  unfolding affine_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   422
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
   423
lemma affine_Inter [intro]: "(\<And>s. s\<in>f \<Longrightarrow> affine s) \<Longrightarrow> affine (\<Inter>f)"
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
   424
  unfolding affine_def by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   425
60303
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60176
diff changeset
   426
lemma affine_Int[intro]: "affine s \<Longrightarrow> affine t \<Longrightarrow> affine (s \<inter> t)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   427
  unfolding affine_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   428
60303
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60176
diff changeset
   429
lemma affine_affine_hull [simp]: "affine(affine hull s)"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   430
  unfolding hull_def
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   431
  using affine_Inter[of "{t. affine t \<and> s \<subseteq> t}"] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   432
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   433
lemma affine_hull_eq[simp]: "(affine hull s = s) \<longleftrightarrow> affine s"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   434
  by (metis affine_affine_hull hull_same)
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   435
62948
7700f467892b lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 62843
diff changeset
   436
lemma affine_hyperplane: "affine {x. a \<bullet> x = b}"
7700f467892b lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 62843
diff changeset
   437
  by (simp add: affine_def algebra_simps) (metis distrib_right mult.left_neutral)
7700f467892b lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 62843
diff changeset
   438
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   439
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
   440
subsubsection \<open>Some explicit formulations (from Lars Schewe)\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   441
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   442
lemma affine:
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   443
  fixes V::"'a::real_vector set"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   444
  shows "affine V \<longleftrightarrow>
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   445
    (\<forall>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> V \<and> setsum u s = 1 \<longrightarrow> (setsum (\<lambda>x. (u x) *\<^sub>R x)) s \<in> V)"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   446
  unfolding affine_def
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   447
  apply rule
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   448
  apply(rule, rule, rule)
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
   449
  apply(erule conjE)+
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   450
  defer
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   451
  apply (rule, rule, rule, rule, rule)
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   452
proof -
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   453
  fix x y u v
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   454
  assume as: "x \<in> V" "y \<in> V" "u + v = (1::real)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   455
    "\<forall>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> V \<and> setsum u s = 1 \<longrightarrow> (\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   456
  then show "u *\<^sub>R x + v *\<^sub>R y \<in> V"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   457
    apply (cases "x = y")
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   458
    using as(4)[THEN spec[where x="{x,y}"], THEN spec[where x="\<lambda>w. if w = x then u else v"]]
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   459
      and as(1-3)
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   460
    apply (auto simp add: scaleR_left_distrib[symmetric])
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   461
    done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   462
next
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   463
  fix s u
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   464
  assume as: "\<forall>x\<in>V. \<forall>y\<in>V. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> V"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   465
    "finite s" "s \<noteq> {}" "s \<subseteq> V" "setsum u s = (1::real)"
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
   466
  define n where "n = card s"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   467
  have "card s = 0 \<or> card s = 1 \<or> card s = 2 \<or> card s > 2" by auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   468
  then show "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   469
  proof (auto simp only: disjE)
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   470
    assume "card s = 2"
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   471
    then have "card s = Suc (Suc 0)"
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   472
      by auto
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   473
    then obtain a b where "s = {a, b}"
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   474
      unfolding card_Suc_eq by auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   475
    then show ?thesis
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   476
      using as(1)[THEN bspec[where x=a], THEN bspec[where x=b]] using as(4,5)
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   477
      by (auto simp add: setsum_clauses(2))
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   478
  next
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   479
    assume "card s > 2"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   480
    then show ?thesis using as and n_def
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   481
    proof (induct n arbitrary: u s)
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   482
      case 0
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   483
      then show ?case by auto
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   484
    next
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   485
      case (Suc n)
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   486
      fix s :: "'a set" and u :: "'a \<Rightarrow> real"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   487
      assume IA:
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   488
        "\<And>u s.  \<lbrakk>2 < card s; \<forall>x\<in>V. \<forall>y\<in>V. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> V; finite s;
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   489
          s \<noteq> {}; s \<subseteq> V; setsum u s = 1; n = card s \<rbrakk> \<Longrightarrow> (\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   490
        and as:
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   491
          "Suc n = card s" "2 < card s" "\<forall>x\<in>V. \<forall>y\<in>V. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> V"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   492
           "finite s" "s \<noteq> {}" "s \<subseteq> V" "setsum u s = 1"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   493
      have "\<exists>x\<in>s. u x \<noteq> 1"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   494
      proof (rule ccontr)
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   495
        assume "\<not> ?thesis"
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   496
        then have "setsum u s = real_of_nat (card s)"
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   497
          unfolding card_eq_setsum by auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   498
        then show False
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
   499
          using as(7) and \<open>card s > 2\<close>
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   500
          by (metis One_nat_def less_Suc0 Zero_not_Suc of_nat_1 of_nat_eq_iff numeral_2_eq_2)
45498
2dc373f1867a avoid numeral-representation-specific rules in metis proof
huffman
parents: 45051
diff changeset
   501
      qed
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   502
      then obtain x where x:"x \<in> s" "u x \<noteq> 1" by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   503
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   504
      have c: "card (s - {x}) = card s - 1"
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   505
        apply (rule card_Diff_singleton)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
   506
        using \<open>x\<in>s\<close> as(4)
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   507
        apply auto
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   508
        done
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   509
      have *: "s = insert x (s - {x})" "finite (s - {x})"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
   510
        using \<open>x\<in>s\<close> and as(4) by auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   511
      have **: "setsum u (s - {x}) = 1 - u x"
49530
wenzelm
parents: 49529
diff changeset
   512
        using setsum_clauses(2)[OF *(2), of u x, unfolded *(1)[symmetric] as(7)] by auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   513
      have ***: "inverse (1 - u x) * setsum u (s - {x}) = 1"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
   514
        unfolding ** using \<open>u x \<noteq> 1\<close> by auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   515
      have "(\<Sum>xa\<in>s - {x}. (inverse (1 - u x) * u xa) *\<^sub>R xa) \<in> V"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   516
      proof (cases "card (s - {x}) > 2")
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   517
        case True
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   518
        then have "s - {x} \<noteq> {}" "card (s - {x}) = n"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   519
          unfolding c and as(1)[symmetric]
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
   520
        proof (rule_tac ccontr)
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   521
          assume "\<not> s - {x} \<noteq> {}"
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
   522
          then have "card (s - {x}) = 0" unfolding card_0_eq[OF *(2)] by simp
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   523
          then show False using True by auto
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   524
        qed auto
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   525
        then show ?thesis
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   526
          apply (rule_tac IA[of "s - {x}" "\<lambda>y. (inverse (1 - u x) * u y)"])
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   527
          unfolding setsum_right_distrib[symmetric]
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   528
          using as and *** and True
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   529
          apply auto
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   530
          done
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   531
      next
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   532
        case False
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   533
        then have "card (s - {x}) = Suc (Suc 0)"
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   534
          using as(2) and c by auto
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   535
        then obtain a b where "(s - {x}) = {a, b}" "a\<noteq>b"
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   536
          unfolding card_Suc_eq by auto
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   537
        then show ?thesis
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   538
          using as(3)[THEN bspec[where x=a], THEN bspec[where x=b]]
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
   539
          using *** *(2) and \<open>s \<subseteq> V\<close>
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   540
          unfolding setsum_right_distrib
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   541
          by (auto simp add: setsum_clauses(2))
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   542
      qed
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   543
      then have "u x + (1 - u x) = 1 \<Longrightarrow>
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   544
          u x *\<^sub>R x + (1 - u x) *\<^sub>R ((\<Sum>xa\<in>s - {x}. u xa *\<^sub>R xa) /\<^sub>R (1 - u x)) \<in> V"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   545
        apply -
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   546
        apply (rule as(3)[rule_format])
51524
7cb5ac44ca9e rename RealVector.thy to Real_Vector_Spaces.thy
hoelzl
parents: 51480
diff changeset
   547
        unfolding  Real_Vector_Spaces.scaleR_right.setsum
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   548
        using x(1) as(6)
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   549
        apply auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   550
        done
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   551
      then show "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V"
49530
wenzelm
parents: 49529
diff changeset
   552
        unfolding scaleR_scaleR[symmetric] and scaleR_right.setsum [symmetric]
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   553
        apply (subst *)
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   554
        unfolding setsum_clauses(2)[OF *(2)]
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
   555
        using \<open>u x \<noteq> 1\<close>
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   556
        apply auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   557
        done
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   558
    qed
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   559
  next
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   560
    assume "card s = 1"
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   561
    then obtain a where "s={a}"
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   562
      by (auto simp add: card_Suc_eq)
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   563
    then show ?thesis
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   564
      using as(4,5) by simp
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
   565
  qed (insert \<open>s\<noteq>{}\<close> \<open>finite s\<close>, auto)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   566
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   567
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   568
lemma affine_hull_explicit:
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   569
  "affine hull p =
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   570
    {y. \<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> setsum u s = 1 \<and> setsum (\<lambda>v. (u v) *\<^sub>R v) s = y}"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   571
  apply (rule hull_unique)
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   572
  apply (subst subset_eq)
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   573
  prefer 3
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   574
  apply rule
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   575
  unfolding mem_Collect_eq
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   576
  apply (erule exE)+
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   577
  apply (erule conjE)+
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   578
  prefer 2
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   579
  apply rule
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   580
proof -
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   581
  fix x
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   582
  assume "x\<in>p"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   583
  then show "\<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x"
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   584
    apply (rule_tac x="{x}" in exI)
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   585
    apply (rule_tac x="\<lambda>x. 1" in exI)
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   586
    apply auto
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   587
    done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   588
next
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   589
  fix t x s u
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   590
  assume as: "p \<subseteq> t" "affine t" "finite s" "s \<noteq> {}"
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   591
    "s \<subseteq> p" "setsum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   592
  then show "x \<in> t"
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   593
    using as(2)[unfolded affine, THEN spec[where x=s], THEN spec[where x=u]]
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   594
    by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   595
next
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   596
  show "affine {y. \<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y}"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   597
    unfolding affine_def
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   598
    apply (rule, rule, rule, rule, rule)
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   599
    unfolding mem_Collect_eq
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   600
  proof -
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   601
    fix u v :: real
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   602
    assume uv: "u + v = 1"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   603
    fix x
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   604
    assume "\<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   605
    then obtain sx ux where
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   606
      x: "finite sx" "sx \<noteq> {}" "sx \<subseteq> p" "setsum ux sx = 1" "(\<Sum>v\<in>sx. ux v *\<^sub>R v) = x"
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   607
      by auto
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   608
    fix y
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   609
    assume "\<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   610
    then obtain sy uy where
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   611
      y: "finite sy" "sy \<noteq> {}" "sy \<subseteq> p" "setsum uy sy = 1" "(\<Sum>v\<in>sy. uy v *\<^sub>R v) = y" by auto
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   612
    have xy: "finite (sx \<union> sy)"
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   613
      using x(1) y(1) by auto
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   614
    have **: "(sx \<union> sy) \<inter> sx = sx" "(sx \<union> sy) \<inter> sy = sy"
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   615
      by auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   616
    show "\<exists>s ua. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and>
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   617
        setsum ua s = 1 \<and> (\<Sum>v\<in>s. ua v *\<^sub>R v) = u *\<^sub>R x + v *\<^sub>R y"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   618
      apply (rule_tac x="sx \<union> sy" in exI)
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   619
      apply (rule_tac x="\<lambda>a. (if a\<in>sx then u * ux a else 0) + (if a\<in>sy then v * uy a else 0)" in exI)
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
   620
      unfolding scaleR_left_distrib setsum.distrib if_smult scaleR_zero_left
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
   621
        ** setsum.inter_restrict[OF xy, symmetric]
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   622
      unfolding scaleR_scaleR[symmetric] Real_Vector_Spaces.scaleR_right.setsum [symmetric]
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   623
        and setsum_right_distrib[symmetric]
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   624
      unfolding x y
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   625
      using x(1-3) y(1-3) uv
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   626
      apply simp
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   627
      done
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   628
  qed
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   629
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   630
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   631
lemma affine_hull_finite:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   632
  assumes "finite s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   633
  shows "affine hull s = {y. \<exists>u. setsum u s = 1 \<and> setsum (\<lambda>v. u v *\<^sub>R v) s = y}"
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   634
  unfolding affine_hull_explicit and set_eq_iff and mem_Collect_eq
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   635
  apply (rule, rule)
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   636
  apply (erule exE)+
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   637
  apply (erule conjE)+
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   638
  defer
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   639
  apply (erule exE)
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   640
  apply (erule conjE)
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   641
proof -
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   642
  fix x u
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   643
  assume "setsum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   644
  then show "\<exists>sa u. finite sa \<and>
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   645
      \<not> (\<forall>x. (x \<in> sa) = (x \<in> {})) \<and> sa \<subseteq> s \<and> setsum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = x"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   646
    apply (rule_tac x=s in exI, rule_tac x=u in exI)
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   647
    using assms
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   648
    apply auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   649
    done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   650
next
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   651
  fix x t u
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   652
  assume "t \<subseteq> s"
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   653
  then have *: "s \<inter> t = t"
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   654
    by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   655
  assume "finite t" "\<not> (\<forall>x. (x \<in> t) = (x \<in> {}))" "setsum u t = 1" "(\<Sum>v\<in>t. u v *\<^sub>R v) = x"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   656
  then show "\<exists>u. setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   657
    apply (rule_tac x="\<lambda>x. if x\<in>t then u x else 0" in exI)
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
   658
    unfolding if_smult scaleR_zero_left and setsum.inter_restrict[OF assms, symmetric] and *
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   659
    apply auto
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   660
    done
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   661
qed
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   662
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   663
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
   664
subsubsection \<open>Stepping theorems and hence small special cases\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   665
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   666
lemma affine_hull_empty[simp]: "affine hull {} = {}"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   667
  by (rule hull_unique) auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   668
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   669
lemma affine_hull_finite_step:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   670
  fixes y :: "'a::real_vector"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   671
  shows
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   672
    "(\<exists>u. setsum u {} = w \<and> setsum (\<lambda>x. u x *\<^sub>R x) {} = y) \<longleftrightarrow> w = 0 \<and> y = 0" (is ?th1)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   673
    and
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   674
    "finite s \<Longrightarrow>
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   675
      (\<exists>u. setsum u (insert a s) = w \<and> setsum (\<lambda>x. u x *\<^sub>R x) (insert a s) = y) \<longleftrightarrow>
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   676
      (\<exists>v u. setsum u s = w - v \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y - v *\<^sub>R a)" (is "_ \<Longrightarrow> ?lhs = ?rhs")
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   677
proof -
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   678
  show ?th1 by simp
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   679
  assume fin: "finite s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   680
  show "?lhs = ?rhs"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   681
  proof
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   682
    assume ?lhs
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   683
    then obtain u where u: "setsum u (insert a s) = w \<and> (\<Sum>x\<in>insert a s. u x *\<^sub>R x) = y"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   684
      by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   685
    show ?rhs
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   686
    proof (cases "a \<in> s")
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   687
      case True
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   688
      then have *: "insert a s = s" by auto
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   689
      show ?thesis
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   690
        using u[unfolded *]
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   691
        apply(rule_tac x=0 in exI)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   692
        apply auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   693
        done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   694
    next
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   695
      case False
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   696
      then show ?thesis
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   697
        apply (rule_tac x="u a" in exI)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   698
        using u and fin
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   699
        apply auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   700
        done
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   701
    qed
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   702
  next
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   703
    assume ?rhs
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   704
    then obtain v u where vu: "setsum u s = w - v"  "(\<Sum>x\<in>s. u x *\<^sub>R x) = y - v *\<^sub>R a"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   705
      by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   706
    have *: "\<And>x M. (if x = a then v else M) *\<^sub>R x = (if x = a then v *\<^sub>R x else M *\<^sub>R x)"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   707
      by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   708
    show ?lhs
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   709
    proof (cases "a \<in> s")
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   710
      case True
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   711
      then show ?thesis
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   712
        apply (rule_tac x="\<lambda>x. (if x=a then v else 0) + u x" in exI)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   713
        unfolding setsum_clauses(2)[OF fin]
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   714
        apply simp
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
   715
        unfolding scaleR_left_distrib and setsum.distrib
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   716
        unfolding vu and * and scaleR_zero_left
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
   717
        apply (auto simp add: setsum.delta[OF fin])
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   718
        done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   719
    next
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
   720
      case False
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   721
      then have **:
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   722
        "\<And>x. x \<in> s \<Longrightarrow> u x = (if x = a then v else u x)"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   723
        "\<And>x. x \<in> s \<Longrightarrow> u x *\<^sub>R x = (if x = a then v *\<^sub>R x else u x *\<^sub>R x)" by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   724
      from False show ?thesis
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   725
        apply (rule_tac x="\<lambda>x. if x=a then v else u x" in exI)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   726
        unfolding setsum_clauses(2)[OF fin] and * using vu
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
   727
        using setsum.cong [of s _ "\<lambda>x. u x *\<^sub>R x" "\<lambda>x. if x = a then v *\<^sub>R x else u x *\<^sub>R x", OF _ **(2)]
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
   728
        using setsum.cong [of s _ u "\<lambda>x. if x = a then v else u x", OF _ **(1)]
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   729
        apply auto
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   730
        done
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   731
    qed
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   732
  qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   733
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   734
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   735
lemma affine_hull_2:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   736
  fixes a b :: "'a::real_vector"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   737
  shows "affine hull {a,b} = {u *\<^sub>R a + v *\<^sub>R b| u v. (u + v = 1)}"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   738
  (is "?lhs = ?rhs")
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   739
proof -
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   740
  have *:
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
   741
    "\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::real)"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   742
    "\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::'a)" by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   743
  have "?lhs = {y. \<exists>u. setsum u {a, b} = 1 \<and> (\<Sum>v\<in>{a, b}. u v *\<^sub>R v) = y}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   744
    using affine_hull_finite[of "{a,b}"] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   745
  also have "\<dots> = {y. \<exists>v u. u b = 1 - v \<and> u b *\<^sub>R b = y - v *\<^sub>R a}"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   746
    by (simp add: affine_hull_finite_step(2)[of "{b}" a])
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   747
  also have "\<dots> = ?rhs" unfolding * by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   748
  finally show ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   749
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   750
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   751
lemma affine_hull_3:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   752
  fixes a b c :: "'a::real_vector"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   753
  shows "affine hull {a,b,c} = { u *\<^sub>R a + v *\<^sub>R b + w *\<^sub>R c| u v w. u + v + w = 1}"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   754
proof -
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   755
  have *:
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
   756
    "\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::real)"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   757
    "\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::'a)" by auto
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   758
  show ?thesis
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   759
    apply (simp add: affine_hull_finite affine_hull_finite_step)
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   760
    unfolding *
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   761
    apply auto
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   762
    apply (rule_tac x=v in exI)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   763
    apply (rule_tac x=va in exI)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   764
    apply auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   765
    apply (rule_tac x=u in exI)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   766
    apply force
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   767
    done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   768
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   769
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   770
lemma mem_affine:
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   771
  assumes "affine S" "x \<in> S" "y \<in> S" "u + v = 1"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   772
  shows "u *\<^sub>R x + v *\<^sub>R y \<in> S"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   773
  using assms affine_def[of S] by auto
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   774
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   775
lemma mem_affine_3:
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   776
  assumes "affine S" "x \<in> S" "y \<in> S" "z \<in> S" "u + v + w = 1"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   777
  shows "u *\<^sub>R x + v *\<^sub>R y + w *\<^sub>R z \<in> S"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   778
proof -
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   779
  have "u *\<^sub>R x + v *\<^sub>R y + w *\<^sub>R z \<in> affine hull {x, y, z}"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   780
    using affine_hull_3[of x y z] assms by auto
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   781
  moreover
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   782
  have "affine hull {x, y, z} \<subseteq> affine hull S"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   783
    using hull_mono[of "{x, y, z}" "S"] assms by auto
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   784
  moreover
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   785
  have "affine hull S = S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   786
    using assms affine_hull_eq[of S] by auto
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
   787
  ultimately show ?thesis by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   788
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   789
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   790
lemma mem_affine_3_minus:
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   791
  assumes "affine S" "x \<in> S" "y \<in> S" "z \<in> S"
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   792
  shows "x + v *\<^sub>R (y-z) \<in> S"
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   793
  using mem_affine_3[of S x y z 1 v "-v"] assms
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   794
  by (simp add: algebra_simps)
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   795
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
   796
corollary mem_affine_3_minus2:
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
   797
    "\<lbrakk>affine S; x \<in> S; y \<in> S; z \<in> S\<rbrakk> \<Longrightarrow> x - v *\<^sub>R (y-z) \<in> S"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
   798
  by (metis add_uminus_conv_diff mem_affine_3_minus real_vector.scale_minus_left)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
   799
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   800
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
   801
subsubsection \<open>Some relations between affine hull and subspaces\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   802
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   803
lemma affine_hull_insert_subset_span:
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   804
  "affine hull (insert a s) \<subseteq> {a + v| v . v \<in> span {x - a | x . x \<in> s}}"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   805
  unfolding subset_eq Ball_def
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   806
  unfolding affine_hull_explicit span_explicit mem_Collect_eq
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
   807
  apply (rule, rule)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
   808
  apply (erule exE)+
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
   809
  apply (erule conjE)+
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   810
proof -
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   811
  fix x t u
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   812
  assume as: "finite t" "t \<noteq> {}" "t \<subseteq> insert a s" "setsum u t = 1" "(\<Sum>v\<in>t. u v *\<^sub>R v) = x"
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   813
  have "(\<lambda>x. x - a) ` (t - {a}) \<subseteq> {x - a |x. x \<in> s}"
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   814
    using as(3) by auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   815
  then show "\<exists>v. x = a + v \<and> (\<exists>S u. finite S \<and> S \<subseteq> {x - a |x. x \<in> s} \<and> (\<Sum>v\<in>S. u v *\<^sub>R v) = v)"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   816
    apply (rule_tac x="x - a" in exI)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   817
    apply (rule conjI, simp)
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   818
    apply (rule_tac x="(\<lambda>x. x - a) ` (t - {a})" in exI)
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   819
    apply (rule_tac x="\<lambda>x. u (x + a)" in exI)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   820
    apply (rule conjI) using as(1) apply simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   821
    apply (erule conjI)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   822
    using as(1)
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
   823
    apply (simp add: setsum.reindex[unfolded inj_on_def] scaleR_right_diff_distrib
49530
wenzelm
parents: 49529
diff changeset
   824
      setsum_subtractf scaleR_left.setsum[symmetric] setsum_diff1 scaleR_left_diff_distrib)
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   825
    unfolding as
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   826
    apply simp
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   827
    done
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   828
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   829
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   830
lemma affine_hull_insert_span:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   831
  assumes "a \<notin> s"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   832
  shows "affine hull (insert a s) = {a + v | v . v \<in> span {x - a | x.  x \<in> s}}"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   833
  apply (rule, rule affine_hull_insert_subset_span)
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   834
  unfolding subset_eq Ball_def
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   835
  unfolding affine_hull_explicit and mem_Collect_eq
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   836
proof (rule, rule, erule exE, erule conjE)
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
   837
  fix y v
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   838
  assume "y = a + v" "v \<in> span {x - a |x. x \<in> s}"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   839
  then obtain t u where obt: "finite t" "t \<subseteq> {x - a |x. x \<in> s}" "a + (\<Sum>v\<in>t. u v *\<^sub>R v) = y"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   840
    unfolding span_explicit by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
   841
  define f where "f = (\<lambda>x. x + a) ` t"
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   842
  have f: "finite f" "f \<subseteq> s" "(\<Sum>v\<in>f. u (v - a) *\<^sub>R (v - a)) = y - a"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
   843
    unfolding f_def using obt by (auto simp add: setsum.reindex[unfolded inj_on_def])
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   844
  have *: "f \<inter> {a} = {}" "f \<inter> - {a} = f"
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   845
    using f(2) assms by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   846
  show "\<exists>sa u. finite sa \<and> sa \<noteq> {} \<and> sa \<subseteq> insert a s \<and> setsum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = y"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   847
    apply (rule_tac x = "insert a f" in exI)
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   848
    apply (rule_tac x = "\<lambda>x. if x=a then 1 - setsum (\<lambda>x. u (x - a)) f else u (x - a)" in exI)
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   849
    using assms and f
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   850
    unfolding setsum_clauses(2)[OF f(1)] and if_smult
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
   851
    unfolding setsum.If_cases[OF f(1), of "\<lambda>x. x = a"]
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   852
    apply (auto simp add: setsum_subtractf scaleR_left.setsum algebra_simps *)
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   853
    done
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   854
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   855
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   856
lemma affine_hull_span:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   857
  assumes "a \<in> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   858
  shows "affine hull s = {a + v | v. v \<in> span {x - a | x. x \<in> s - {a}}}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   859
  using affine_hull_insert_span[of a "s - {a}", unfolded insert_Diff[OF assms]] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   860
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   861
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
   862
subsubsection \<open>Parallel affine sets\<close>
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   863
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
   864
definition affine_parallel :: "'a::real_vector set \<Rightarrow> 'a::real_vector set \<Rightarrow> bool"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   865
  where "affine_parallel S T \<longleftrightarrow> (\<exists>a. T = (\<lambda>x. a + x) ` S)"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   866
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   867
lemma affine_parallel_expl_aux:
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   868
  fixes S T :: "'a::real_vector set"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   869
  assumes "\<forall>x. x \<in> S \<longleftrightarrow> a + x \<in> T"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   870
  shows "T = (\<lambda>x. a + x) ` S"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   871
proof -
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   872
  {
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   873
    fix x
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   874
    assume "x \<in> T"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   875
    then have "( - a) + x \<in> S"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   876
      using assms by auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   877
    then have "x \<in> ((\<lambda>x. a + x) ` S)"
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   878
      using imageI[of "-a+x" S "(\<lambda>x. a+x)"] by auto
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   879
  }
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   880
  moreover have "T \<ge> (\<lambda>x. a + x) ` S"
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   881
    using assms by auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   882
  ultimately show ?thesis by auto
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   883
qed
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   884
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   885
lemma affine_parallel_expl: "affine_parallel S T \<longleftrightarrow> (\<exists>a. \<forall>x. x \<in> S \<longleftrightarrow> a + x \<in> T)"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   886
  unfolding affine_parallel_def
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   887
  using affine_parallel_expl_aux[of S _ T] by auto
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   888
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   889
lemma affine_parallel_reflex: "affine_parallel S S"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   890
  unfolding affine_parallel_def
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   891
  apply (rule exI[of _ "0"])
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   892
  apply auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   893
  done
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   894
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   895
lemma affine_parallel_commut:
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   896
  assumes "affine_parallel A B"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   897
  shows "affine_parallel B A"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   898
proof -
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53676
diff changeset
   899
  from assms obtain a where B: "B = (\<lambda>x. a + x) ` A"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   900
    unfolding affine_parallel_def by auto
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53676
diff changeset
   901
  have [simp]: "(\<lambda>x. x - a) = plus (- a)" by (simp add: fun_eq_iff)
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53676
diff changeset
   902
  from B show ?thesis
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   903
    using translation_galois [of B a A]
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   904
    unfolding affine_parallel_def by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   905
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   906
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   907
lemma affine_parallel_assoc:
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   908
  assumes "affine_parallel A B"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   909
    and "affine_parallel B C"
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
   910
  shows "affine_parallel A C"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   911
proof -
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   912
  from assms obtain ab where "B = (\<lambda>x. ab + x) ` A"
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
   913
    unfolding affine_parallel_def by auto
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
   914
  moreover
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   915
  from assms obtain bc where "C = (\<lambda>x. bc + x) ` B"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   916
    unfolding affine_parallel_def by auto
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   917
  ultimately show ?thesis
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   918
    using translation_assoc[of bc ab A] unfolding affine_parallel_def by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   919
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   920
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   921
lemma affine_translation_aux:
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   922
  fixes a :: "'a::real_vector"
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   923
  assumes "affine ((\<lambda>x. a + x) ` S)"
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   924
  shows "affine S"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   925
proof -
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   926
  {
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   927
    fix x y u v
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   928
    assume xy: "x \<in> S" "y \<in> S" "(u :: real) + v = 1"
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   929
    then have "(a + x) \<in> ((\<lambda>x. a + x) ` S)" "(a + y) \<in> ((\<lambda>x. a + x) ` S)"
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   930
      by auto
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   931
    then have h1: "u *\<^sub>R  (a + x) + v *\<^sub>R (a + y) \<in> (\<lambda>x. a + x) ` S"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   932
      using xy assms unfolding affine_def by auto
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   933
    have "u *\<^sub>R (a + x) + v *\<^sub>R (a + y) = (u + v) *\<^sub>R a + (u *\<^sub>R x + v *\<^sub>R y)"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   934
      by (simp add: algebra_simps)
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   935
    also have "\<dots> = a + (u *\<^sub>R x + v *\<^sub>R y)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
   936
      using \<open>u + v = 1\<close> by auto
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   937
    ultimately have "a + (u *\<^sub>R x + v *\<^sub>R y) \<in> (\<lambda>x. a + x) ` S"
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   938
      using h1 by auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   939
    then have "u *\<^sub>R x + v *\<^sub>R y : S" by auto
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   940
  }
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   941
  then show ?thesis unfolding affine_def by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   942
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   943
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   944
lemma affine_translation:
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   945
  fixes a :: "'a::real_vector"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   946
  shows "affine S \<longleftrightarrow> affine ((\<lambda>x. a + x) ` S)"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   947
proof -
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   948
  have "affine S \<Longrightarrow> affine ((\<lambda>x. a + x) ` S)"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   949
    using affine_translation_aux[of "-a" "((\<lambda>x. a + x) ` S)"]
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   950
    using translation_assoc[of "-a" a S] by auto
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   951
  then show ?thesis using affine_translation_aux by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   952
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   953
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   954
lemma parallel_is_affine:
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   955
  fixes S T :: "'a::real_vector set"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   956
  assumes "affine S" "affine_parallel S T"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   957
  shows "affine T"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   958
proof -
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   959
  from assms obtain a where "T = (\<lambda>x. a + x) ` S"
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
   960
    unfolding affine_parallel_def by auto
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   961
  then show ?thesis
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   962
    using affine_translation assms by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   963
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   964
44361
75ec83d45303 remove unnecessary euclidean_space class constraints
huffman
parents: 44349
diff changeset
   965
lemma subspace_imp_affine: "subspace s \<Longrightarrow> affine s"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   966
  unfolding subspace_def affine_def by auto
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   967
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   968
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
   969
subsubsection \<open>Subspace parallel to an affine set\<close>
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
   970
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   971
lemma subspace_affine: "subspace S \<longleftrightarrow> affine S \<and> 0 \<in> S"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   972
proof -
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   973
  have h0: "subspace S \<Longrightarrow> affine S \<and> 0 \<in> S"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   974
    using subspace_imp_affine[of S] subspace_0 by auto
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   975
  {
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   976
    assume assm: "affine S \<and> 0 \<in> S"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   977
    {
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   978
      fix c :: real
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
   979
      fix x
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
   980
      assume x: "x \<in> S"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   981
      have "c *\<^sub>R x = (1-c) *\<^sub>R 0 + c *\<^sub>R x" by auto
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   982
      moreover
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
   983
      have "(1 - c) *\<^sub>R 0 + c *\<^sub>R x \<in> S"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
   984
        using affine_alt[of S] assm x by auto
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   985
      ultimately have "c *\<^sub>R x \<in> S" by auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   986
    }
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
   987
    then have h1: "\<forall>c. \<forall>x \<in> S. c *\<^sub>R x \<in> S" by auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   988
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   989
    {
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   990
      fix x y
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
   991
      assume xy: "x \<in> S" "y \<in> S"
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
   992
      define u where "u = (1 :: real)/2"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   993
      have "(1/2) *\<^sub>R (x+y) = (1/2) *\<^sub>R (x+y)"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   994
        by auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   995
      moreover
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   996
      have "(1/2) *\<^sub>R (x+y)=(1/2) *\<^sub>R x + (1-(1/2)) *\<^sub>R y"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
   997
        by (simp add: algebra_simps)
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
   998
      moreover
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
   999
      have "(1 - u) *\<^sub>R x + u *\<^sub>R y \<in> S"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  1000
        using affine_alt[of S] assm xy by auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1001
      ultimately
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
  1002
      have "(1/2) *\<^sub>R (x+y) \<in> S"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1003
        using u_def by auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1004
      moreover
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  1005
      have "x + y = 2 *\<^sub>R ((1/2) *\<^sub>R (x+y))"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1006
        by auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1007
      ultimately
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  1008
      have "x + y \<in> S"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1009
        using h1[rule_format, of "(1/2) *\<^sub>R (x+y)" "2"] by auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1010
    }
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1011
    then have "\<forall>x \<in> S. \<forall>y \<in> S. x + y \<in> S"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1012
      by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1013
    then have "subspace S"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1014
      using h1 assm unfolding subspace_def by auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1015
  }
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1016
  then show ?thesis using h0 by metis
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1017
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1018
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1019
lemma affine_diffs_subspace:
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
  1020
  assumes "affine S" "a \<in> S"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1021
  shows "subspace ((\<lambda>x. (-a)+x) ` S)"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1022
proof -
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53676
diff changeset
  1023
  have [simp]: "(\<lambda>x. x - a) = plus (- a)" by (simp add: fun_eq_iff)
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1024
  have "affine ((\<lambda>x. (-a)+x) ` S)"
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  1025
    using  affine_translation assms by auto
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1026
  moreover have "0 : ((\<lambda>x. (-a)+x) ` S)"
53333
e9dba6602a84 tuned proofs;
wenzelm
parents: 53302
diff changeset
  1027
    using assms exI[of "(\<lambda>x. x\<in>S \<and> -a+x = 0)" a] by auto
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  1028
  ultimately show ?thesis using subspace_affine by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1029
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1030
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1031
lemma parallel_subspace_explicit:
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  1032
  assumes "affine S"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  1033
    and "a \<in> S"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  1034
  assumes "L \<equiv> {y. \<exists>x \<in> S. (-a) + x = y}"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  1035
  shows "subspace L \<and> affine_parallel S L"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1036
proof -
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53676
diff changeset
  1037
  from assms have "L = plus (- a) ` S" by auto
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53676
diff changeset
  1038
  then have par: "affine_parallel S L"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  1039
    unfolding affine_parallel_def ..
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  1040
  then have "affine L" using assms parallel_is_affine by auto
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1041
  moreover have "0 \<in> L"
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53676
diff changeset
  1042
    using assms by auto
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1043
  ultimately show ?thesis
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1044
    using subspace_affine par by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1045
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1046
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1047
lemma parallel_subspace_aux:
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1048
  assumes "subspace A"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1049
    and "subspace B"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1050
    and "affine_parallel A B"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1051
  shows "A \<supseteq> B"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1052
proof -
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  1053
  from assms obtain a where a: "\<forall>x. x \<in> A \<longleftrightarrow> a + x \<in> B"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1054
    using affine_parallel_expl[of A B] by auto
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1055
  then have "-a \<in> A"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1056
    using assms subspace_0[of B] by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1057
  then have "a \<in> A"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1058
    using assms subspace_neg[of A "-a"] by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1059
  then show ?thesis
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  1060
    using assms a unfolding subspace_def by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1061
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1062
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1063
lemma parallel_subspace:
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1064
  assumes "subspace A"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1065
    and "subspace B"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1066
    and "affine_parallel A B"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1067
  shows "A = B"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1068
proof
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1069
  show "A \<supseteq> B"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1070
    using assms parallel_subspace_aux by auto
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1071
  show "A \<subseteq> B"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1072
    using assms parallel_subspace_aux[of B A] affine_parallel_commut by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1073
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1074
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1075
lemma affine_parallel_subspace:
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1076
  assumes "affine S" "S \<noteq> {}"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1077
  shows "\<exists>!L. subspace L \<and> affine_parallel S L"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1078
proof -
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1079
  have ex: "\<exists>L. subspace L \<and> affine_parallel S L"
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  1080
    using assms parallel_subspace_explicit by auto
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1081
  {
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1082
    fix L1 L2
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1083
    assume ass: "subspace L1 \<and> affine_parallel S L1" "subspace L2 \<and> affine_parallel S L2"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1084
    then have "affine_parallel L1 L2"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1085
      using affine_parallel_commut[of S L1] affine_parallel_assoc[of L1 S L2] by auto
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1086
    then have "L1 = L2"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1087
      using ass parallel_subspace by auto
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1088
  }
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1089
  then show ?thesis using ex by auto
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1090
qed
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1091
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1092
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  1093
subsection \<open>Cones\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1094
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1095
definition cone :: "'a::real_vector set \<Rightarrow> bool"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1096
  where "cone s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>c\<ge>0. c *\<^sub>R x \<in> s)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1097
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1098
lemma cone_empty[intro, simp]: "cone {}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1099
  unfolding cone_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1100
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1101
lemma cone_univ[intro, simp]: "cone UNIV"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1102
  unfolding cone_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1103
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1104
lemma cone_Inter[intro]: "\<forall>s\<in>f. cone s \<Longrightarrow> cone (\<Inter>f)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1105
  unfolding cone_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1106
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1107
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  1108
subsubsection \<open>Conic hull\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1109
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1110
lemma cone_cone_hull: "cone (cone hull s)"
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44142
diff changeset
  1111
  unfolding hull_def by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1112
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1113
lemma cone_hull_eq: "cone hull s = s \<longleftrightarrow> cone s"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1114
  apply (rule hull_eq)
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1115
  using cone_Inter
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1116
  unfolding subset_eq
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1117
  apply auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1118
  done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1119
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1120
lemma mem_cone:
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1121
  assumes "cone S" "x \<in> S" "c \<ge> 0"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1122
  shows "c *\<^sub>R x : S"
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1123
  using assms cone_def[of S] by auto
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1124
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1125
lemma cone_contains_0:
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1126
  assumes "cone S"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1127
  shows "S \<noteq> {} \<longleftrightarrow> 0 \<in> S"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1128
proof -
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1129
  {
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1130
    assume "S \<noteq> {}"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1131
    then obtain a where "a \<in> S" by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1132
    then have "0 \<in> S"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1133
      using assms mem_cone[of S a 0] by auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1134
  }
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1135
  then show ?thesis by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1136
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1137
44361
75ec83d45303 remove unnecessary euclidean_space class constraints
huffman
parents: 44349
diff changeset
  1138
lemma cone_0: "cone {0}"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1139
  unfolding cone_def by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1140
61952
546958347e05 prefer symbols for "Union", "Inter";
wenzelm
parents: 61945
diff changeset
  1141
lemma cone_Union[intro]: "(\<forall>s\<in>f. cone s) \<longrightarrow> cone (\<Union>f)"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1142
  unfolding cone_def by blast
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1143
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1144
lemma cone_iff:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1145
  assumes "S \<noteq> {}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1146
  shows "cone S \<longleftrightarrow> 0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> (op *\<^sub>R c) ` S = S)"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1147
proof -
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1148
  {
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1149
    assume "cone S"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1150
    {
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1151
      fix c :: real
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1152
      assume "c > 0"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1153
      {
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1154
        fix x
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1155
        assume "x \<in> S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1156
        then have "x \<in> (op *\<^sub>R c) ` S"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1157
          unfolding image_def
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  1158
          using \<open>cone S\<close> \<open>c>0\<close> mem_cone[of S x "1/c"]
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  1159
            exI[of "(\<lambda>t. t \<in> S \<and> x = c *\<^sub>R t)" "(1 / c) *\<^sub>R x"]
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1160
          by auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1161
      }
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1162
      moreover
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1163
      {
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1164
        fix x
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1165
        assume "x \<in> (op *\<^sub>R c) ` S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1166
        then have "x \<in> S"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  1167
          using \<open>cone S\<close> \<open>c > 0\<close>
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  1168
          unfolding cone_def image_def \<open>c > 0\<close> by auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1169
      }
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1170
      ultimately have "(op *\<^sub>R c) ` S = S" by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1171
    }
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1172
    then have "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> (op *\<^sub>R c) ` S = S)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  1173
      using \<open>cone S\<close> cone_contains_0[of S] assms by auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1174
  }
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1175
  moreover
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1176
  {
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1177
    assume a: "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> (op *\<^sub>R c) ` S = S)"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1178
    {
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1179
      fix x
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1180
      assume "x \<in> S"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1181
      fix c1 :: real
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1182
      assume "c1 \<ge> 0"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1183
      then have "c1 = 0 \<or> c1 > 0" by auto
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  1184
      then have "c1 *\<^sub>R x \<in> S" using a \<open>x \<in> S\<close> by auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1185
    }
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1186
    then have "cone S" unfolding cone_def by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1187
  }
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1188
  ultimately show ?thesis by blast
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1189
qed
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1190
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1191
lemma cone_hull_empty: "cone hull {} = {}"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1192
  by (metis cone_empty cone_hull_eq)
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1193
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1194
lemma cone_hull_empty_iff: "S = {} \<longleftrightarrow> cone hull S = {}"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1195
  by (metis bot_least cone_hull_empty hull_subset xtrans(5))
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1196
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1197
lemma cone_hull_contains_0: "S \<noteq> {} \<longleftrightarrow> 0 \<in> cone hull S"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1198
  using cone_cone_hull[of S] cone_contains_0[of "cone hull S"] cone_hull_empty_iff[of S]
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1199
  by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1200
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1201
lemma mem_cone_hull:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1202
  assumes "x \<in> S" "c \<ge> 0"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1203
  shows "c *\<^sub>R x \<in> cone hull S"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1204
  by (metis assms cone_cone_hull hull_inc mem_cone)
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1205
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1206
lemma cone_hull_expl: "cone hull S = {c *\<^sub>R x | c x. c \<ge> 0 \<and> x \<in> S}"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1207
  (is "?lhs = ?rhs")
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1208
proof -
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1209
  {
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1210
    fix x
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1211
    assume "x \<in> ?rhs"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  1212
    then obtain cx :: real and xx where x: "x = cx *\<^sub>R xx" "cx \<ge> 0" "xx \<in> S"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1213
      by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1214
    fix c :: real
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1215
    assume c: "c \<ge> 0"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1216
    then have "c *\<^sub>R x = (c * cx) *\<^sub>R xx"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  1217
      using x by (simp add: algebra_simps)
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1218
    moreover
56536
aefb4a8da31f made mult_nonneg_nonneg a simp rule
nipkow
parents: 56480
diff changeset
  1219
    have "c * cx \<ge> 0" using c x by auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1220
    ultimately
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  1221
    have "c *\<^sub>R x \<in> ?rhs" using x by auto
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1222
  }
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1223
  then have "cone ?rhs"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1224
    unfolding cone_def by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1225
  then have "?rhs \<in> Collect cone"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1226
    unfolding mem_Collect_eq by auto
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1227
  {
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1228
    fix x
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1229
    assume "x \<in> S"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1230
    then have "1 *\<^sub>R x \<in> ?rhs"
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  1231
      apply auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1232
      apply (rule_tac x = 1 in exI)
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1233
      apply auto
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1234
      done
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1235
    then have "x \<in> ?rhs" by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1236
  }
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1237
  then have "S \<subseteq> ?rhs" by auto
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1238
  then have "?lhs \<subseteq> ?rhs"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  1239
    using \<open>?rhs \<in> Collect cone\<close> hull_minimal[of S "?rhs" "cone"] by auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1240
  moreover
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1241
  {
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1242
    fix x
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1243
    assume "x \<in> ?rhs"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  1244
    then obtain cx :: real and xx where x: "x = cx *\<^sub>R xx" "cx \<ge> 0" "xx \<in> S"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1245
      by auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1246
    then have "xx \<in> cone hull S"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1247
      using hull_subset[of S] by auto
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1248
    then have "x \<in> ?lhs"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  1249
      using x cone_cone_hull[of S] cone_def[of "cone hull S"] by auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1250
  }
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1251
  ultimately show ?thesis by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1252
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1253
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1254
lemma cone_closure:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1255
  fixes S :: "'a::real_normed_vector set"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1256
  assumes "cone S"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1257
  shows "cone (closure S)"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1258
proof (cases "S = {}")
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1259
  case True
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1260
  then show ?thesis by auto
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1261
next
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1262
  case False
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1263
  then have "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> op *\<^sub>R c ` S = S)"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1264
    using cone_iff[of S] assms by auto
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1265
  then have "0 \<in> closure S \<and> (\<forall>c. c > 0 \<longrightarrow> op *\<^sub>R c ` closure S = closure S)"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1266
    using closure_subset by (auto simp add: closure_scaleR)
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1267
  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
  1268
    using False cone_iff[of "closure S"] by auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1269
qed
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1270
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1271
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  1272
subsection \<open>Affine dependence and consequential theorems (from Lars Schewe)\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1273
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1274
definition affine_dependent :: "'a::real_vector set \<Rightarrow> bool"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1275
  where "affine_dependent s \<longleftrightarrow> (\<exists>x\<in>s. x \<in> affine hull (s - {x}))"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1276
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  1277
lemma affine_dependent_subset:
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  1278
   "\<lbrakk>affine_dependent s; s \<subseteq> t\<rbrakk> \<Longrightarrow> affine_dependent t"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  1279
apply (simp add: affine_dependent_def Bex_def)
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  1280
apply (blast dest: hull_mono [OF Diff_mono [OF _ subset_refl]])
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  1281
done
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  1282
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  1283
lemma affine_independent_subset:
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  1284
  shows "\<lbrakk>~ affine_dependent t; s \<subseteq> t\<rbrakk> \<Longrightarrow> ~ affine_dependent s"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  1285
by (metis affine_dependent_subset)
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  1286
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  1287
lemma affine_independent_Diff:
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  1288
   "~ affine_dependent s \<Longrightarrow> ~ affine_dependent(s - t)"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  1289
by (meson Diff_subset affine_dependent_subset)
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  1290
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1291
lemma affine_dependent_explicit:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1292
  "affine_dependent p \<longleftrightarrow>
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1293
    (\<exists>s u. finite s \<and> s \<subseteq> p \<and> setsum u s = 0 \<and>
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1294
      (\<exists>v\<in>s. u v \<noteq> 0) \<and> setsum (\<lambda>v. u v *\<^sub>R v) s = 0)"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1295
  unfolding affine_dependent_def affine_hull_explicit mem_Collect_eq
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1296
  apply rule
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1297
  apply (erule bexE, erule exE, erule exE)
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1298
  apply (erule conjE)+
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1299
  defer
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1300
  apply (erule exE, erule exE)
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1301
  apply (erule conjE)+
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1302
  apply (erule bexE)
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1303
proof -
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1304
  fix x s u
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1305
  assume as: "x \<in> p" "finite s" "s \<noteq> {}" "s \<subseteq> p - {x}" "setsum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1306
  have "x \<notin> s" using as(1,4) by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1307
  show "\<exists>s u. finite s \<and> s \<subseteq> p \<and> setsum u s = 0 \<and> (\<exists>v\<in>s. u v \<noteq> 0) \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = 0"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1308
    apply (rule_tac x="insert x s" in exI, rule_tac x="\<lambda>v. if v = x then - 1 else u v" in exI)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  1309
    unfolding if_smult and setsum_clauses(2)[OF as(2)] and setsum_delta_notmem[OF \<open>x\<notin>s\<close>] and as
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1310
    using as
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1311
    apply auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1312
    done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1313
next
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1314
  fix s u v
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1315
  assume as: "finite s" "s \<subseteq> p" "setsum u s = 0" "(\<Sum>v\<in>s. u v *\<^sub>R v) = 0" "v \<in> s" "u v \<noteq> 0"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1316
  have "s \<noteq> {v}"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1317
    using as(3,6) by auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1318
  then show "\<exists>x\<in>p. \<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p - {x} \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1319
    apply (rule_tac x=v in bexI)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1320
    apply (rule_tac x="s - {v}" in exI)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1321
    apply (rule_tac x="\<lambda>x. - (1 / u v) * u x" in exI)
49530
wenzelm
parents: 49529
diff changeset
  1322
    unfolding scaleR_scaleR[symmetric] and scaleR_right.setsum [symmetric]
wenzelm
parents: 49529
diff changeset
  1323
    unfolding setsum_right_distrib[symmetric] and setsum_diff1[OF as(1)]
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1324
    using as
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1325
    apply auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1326
    done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1327
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1328
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1329
lemma affine_dependent_explicit_finite:
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1330
  fixes s :: "'a::real_vector set"
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1331
  assumes "finite s"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1332
  shows "affine_dependent s \<longleftrightarrow>
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1333
    (\<exists>u. setsum u s = 0 \<and> (\<exists>v\<in>s. u v \<noteq> 0) \<and> setsum (\<lambda>v. u v *\<^sub>R v) s = 0)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1334
  (is "?lhs = ?rhs")
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1335
proof
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1336
  have *: "\<And>vt u v. (if vt then u v else 0) *\<^sub>R v = (if vt then (u v) *\<^sub>R v else 0::'a)"
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1337
    by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1338
  assume ?lhs
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1339
  then obtain t u v where
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1340
    "finite t" "t \<subseteq> s" "setsum u t = 0" "v\<in>t" "u v \<noteq> 0"  "(\<Sum>v\<in>t. u v *\<^sub>R v) = 0"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1341
    unfolding affine_dependent_explicit by auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1342
  then show ?rhs
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1343
    apply (rule_tac x="\<lambda>x. if x\<in>t then u x else 0" in exI)
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  1344
    apply auto unfolding * and setsum.inter_restrict[OF assms, symmetric]
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  1345
    unfolding Int_absorb1[OF \<open>t\<subseteq>s\<close>]
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1346
    apply auto
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1347
    done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1348
next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1349
  assume ?rhs
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1350
  then obtain u v where "setsum u s = 0"  "v\<in>s" "u v \<noteq> 0" "(\<Sum>v\<in>s. u v *\<^sub>R v) = 0"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1351
    by auto
49529
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1352
  then show ?lhs unfolding affine_dependent_explicit
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1353
    using assms by auto
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1354
qed
d523702bdae7 tuned proofs;
wenzelm
parents: 47445
diff changeset
  1355
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1356
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  1357
subsection \<open>Connectedness of convex sets\<close>
44465
fa56622bb7bc move connected_real_lemma to the one place it is used
huffman
parents: 44457
diff changeset
  1358
51480
3793c3a11378 move connected to HOL image; used to show intermediate value theorem
hoelzl
parents: 51475
diff changeset
  1359
lemma connectedD:
3793c3a11378 move connected to HOL image; used to show intermediate value theorem
hoelzl
parents: 51475
diff changeset
  1360
  "connected S \<Longrightarrow> open A \<Longrightarrow> open B \<Longrightarrow> S \<subseteq> A \<union> B \<Longrightarrow> A \<inter> B \<inter> S = {} \<Longrightarrow> A \<inter> S = {} \<or> B \<inter> 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
  1361
  by (rule Topological_Spaces.topological_space_class.connectedD)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1362
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1363
lemma convex_connected:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1364
  fixes s :: "'a::real_normed_vector set"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1365
  assumes "convex s"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1366
  shows "connected s"
51480
3793c3a11378 move connected to HOL image; used to show intermediate value theorem
hoelzl
parents: 51475
diff changeset
  1367
proof (rule connectedI)
3793c3a11378 move connected to HOL image; used to show intermediate value theorem
hoelzl
parents: 51475
diff changeset
  1368
  fix A B
3793c3a11378 move connected to HOL image; used to show intermediate value theorem
hoelzl
parents: 51475
diff changeset
  1369
  assume "open A" "open B" "A \<inter> B \<inter> s = {}" "s \<subseteq> A \<union> B"
3793c3a11378 move connected to HOL image; used to show intermediate value theorem
hoelzl
parents: 51475
diff changeset
  1370
  moreover
3793c3a11378 move connected to HOL image; used to show intermediate value theorem
hoelzl
parents: 51475
diff changeset
  1371
  assume "A \<inter> s \<noteq> {}" "B \<inter> s \<noteq> {}"
3793c3a11378 move connected to HOL image; used to show intermediate value theorem
hoelzl
parents: 51475
diff changeset
  1372
  then obtain a b where a: "a \<in> A" "a \<in> s" and b: "b \<in> B" "b \<in> s" by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  1373
  define f where [abs_def]: "f u = u *\<^sub>R a + (1 - u) *\<^sub>R b" for u
51480
3793c3a11378 move connected to HOL image; used to show intermediate value theorem
hoelzl
parents: 51475
diff changeset
  1374
  then have "continuous_on {0 .. 1} f"
56371
fb9ae0727548 extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents: 56369
diff changeset
  1375
    by (auto intro!: continuous_intros)
51480
3793c3a11378 move connected to HOL image; used to show intermediate value theorem
hoelzl
parents: 51475
diff changeset
  1376
  then have "connected (f ` {0 .. 1})"
3793c3a11378 move connected to HOL image; used to show intermediate value theorem
hoelzl
parents: 51475
diff changeset
  1377
    by (auto intro!: connected_continuous_image)
3793c3a11378 move connected to HOL image; used to show intermediate value theorem
hoelzl
parents: 51475
diff changeset
  1378
  note connectedD[OF this, of A B]
3793c3a11378 move connected to HOL image; used to show intermediate value theorem
hoelzl
parents: 51475
diff changeset
  1379
  moreover have "a \<in> A \<inter> f ` {0 .. 1}"
3793c3a11378 move connected to HOL image; used to show intermediate value theorem
hoelzl
parents: 51475
diff changeset
  1380
    using a by (auto intro!: image_eqI[of _ _ 1] simp: f_def)
3793c3a11378 move connected to HOL image; used to show intermediate value theorem
hoelzl
parents: 51475
diff changeset
  1381
  moreover have "b \<in> B \<inter> f ` {0 .. 1}"
3793c3a11378 move connected to HOL image; used to show intermediate value theorem
hoelzl
parents: 51475
diff changeset
  1382
    using b by (auto intro!: image_eqI[of _ _ 0] simp: f_def)
3793c3a11378 move connected to HOL image; used to show intermediate value theorem
hoelzl
parents: 51475
diff changeset
  1383
  moreover have "f ` {0 .. 1} \<subseteq> s"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  1384
    using \<open>convex s\<close> a b unfolding convex_def f_def by auto
51480
3793c3a11378 move connected to HOL image; used to show intermediate value theorem
hoelzl
parents: 51475
diff changeset
  1385
  ultimately show False by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1386
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1387
61426
d53db136e8fd new material on path_component_sets, inside, outside, etc. And more default simprules
paulson <lp15@cam.ac.uk>
parents: 61222
diff changeset
  1388
corollary connected_UNIV[intro]: "connected (UNIV :: 'a::real_normed_vector set)"
d53db136e8fd new material on path_component_sets, inside, outside, etc. And more default simprules
paulson <lp15@cam.ac.uk>
parents: 61222
diff changeset
  1389
  by(simp add: convex_connected)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1390
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62097
diff changeset
  1391
proposition clopen:
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62097
diff changeset
  1392
  fixes s :: "'a :: real_normed_vector set"
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62097
diff changeset
  1393
  shows "closed s \<and> open s \<longleftrightarrow> s = {} \<or> s = UNIV"
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62097
diff changeset
  1394
apply (rule iffI)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62097
diff changeset
  1395
 apply (rule connected_UNIV [unfolded connected_clopen, rule_format])
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62097
diff changeset
  1396
 apply (force simp add: open_openin closed_closedin, force)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62097
diff changeset
  1397
done
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62097
diff changeset
  1398
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62097
diff changeset
  1399
corollary compact_open:
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62097
diff changeset
  1400
  fixes s :: "'a :: euclidean_space set"
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62097
diff changeset
  1401
  shows "compact s \<and> open s \<longleftrightarrow> s = {}"
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62097
diff changeset
  1402
  by (auto simp: compact_eq_bounded_closed clopen)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62097
diff changeset
  1403
62948
7700f467892b lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 62843
diff changeset
  1404
corollary finite_imp_not_open:
7700f467892b lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 62843
diff changeset
  1405
    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
  1406
    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
  1407
  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
  1408
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  1409
corollary empty_interior_finite:
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  1410
    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
  1411
    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
  1412
  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
  1413
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  1414
text \<open>Balls, being convex, are connected.\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1415
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  1416
lemma convex_prod:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1417
  assumes "\<And>i. i \<in> Basis \<Longrightarrow> convex {x. P i x}"
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
  1418
  shows "convex {x. \<forall>i\<in>Basis. P i (x\<bullet>i)}"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50104
diff changeset
  1419
  using assms unfolding convex_def
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50104
diff changeset
  1420
  by (auto simp: inner_add_left)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50104
diff changeset
  1421
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50104
diff changeset
  1422
lemma convex_positive_orthant: "convex {x::'a::euclidean_space. (\<forall>i\<in>Basis. 0 \<le> x\<bullet>i)}"
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  1423
  by (rule convex_prod) (simp add: atLeast_def[symmetric] convex_real_interval)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1424
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1425
lemma convex_local_global_minimum:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1426
  fixes s :: "'a::real_normed_vector set"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1427
  assumes "e > 0"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1428
    and "convex_on s f"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1429
    and "ball x e \<subseteq> s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1430
    and "\<forall>y\<in>ball x e. f x \<le> f y"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1431
  shows "\<forall>y\<in>s. f x \<le> f y"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1432
proof (rule ccontr)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1433
  have "x \<in> s" using assms(1,3) by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1434
  assume "\<not> ?thesis"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1435
  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
  1436
  then have xy: "0 < dist x y"  by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1437
  then obtain u where "0 < u" "u \<le> 1" and u: "u < e / dist x y"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  1438
    using real_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
  1439
  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
  1440
    using \<open>x\<in>s\<close> \<open>y\<in>s\<close>
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1441
    using assms(2)[unfolded convex_on_def,
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1442
      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
  1443
    by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1444
  moreover
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1445
  have *: "x - ((1 - u) *\<^sub>R x + u *\<^sub>R y) = u *\<^sub>R (x - y)"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1446
    by (simp add: algebra_simps)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1447
  have "(1 - u) *\<^sub>R x + u *\<^sub>R y \<in> ball x e"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1448
    unfolding mem_ball dist_norm
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  1449
    unfolding * and norm_scaleR and abs_of_pos[OF \<open>0<u\<close>]
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1450
    unfolding dist_norm[symmetric]
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1451
    using u
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1452
    unfolding pos_less_divide_eq[OF xy]
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1453
    by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1454
  then have "f x \<le> f ((1 - u) *\<^sub>R x + u *\<^sub>R y)"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1455
    using assms(4) by auto
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1456
  ultimately show False
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  1457
    using mult_strict_left_mono[OF y \<open>u>0\<close>]
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1458
    unfolding left_diff_distrib
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1459
    by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1460
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1461
60800
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
  1462
lemma convex_ball [iff]:
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1463
  fixes x :: "'a::real_normed_vector"
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  1464
  shows "convex (ball x e)"
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1465
proof (auto simp add: convex_def)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1466
  fix y z
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1467
  assume yz: "dist x y < e" "dist x z < e"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1468
  fix u v :: real
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1469
  assume uv: "0 \<le> u" "0 \<le> v" "u + v = 1"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1470
  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
  1471
    using uv yz
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1472
    using convex_on_dist [of "ball x e" x, unfolded convex_on_def,
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1473
      THEN bspec[where x=y], THEN bspec[where x=z]]
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1474
    by auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1475
  then show "dist x (u *\<^sub>R y + v *\<^sub>R z) < e"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1476
    using convex_bound_lt[OF yz uv] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1477
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1478
60800
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
  1479
lemma convex_cball [iff]:
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1480
  fixes x :: "'a::real_normed_vector"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1481
  shows "convex (cball x e)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1482
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1483
  {
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1484
    fix y z
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1485
    assume yz: "dist x y \<le> e" "dist x z \<le> e"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1486
    fix u v :: real
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1487
    assume uv: "0 \<le> u" "0 \<le> v" "u + v = 1"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1488
    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
  1489
      using uv yz
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1490
      using convex_on_dist [of "cball x e" x, unfolded convex_on_def,
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1491
        THEN bspec[where x=y], THEN bspec[where x=z]]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1492
      by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1493
    then have "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> e"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1494
      using convex_bound_le[OF yz uv] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1495
  }
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1496
  then show ?thesis by (auto simp add: convex_def Ball_def)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1497
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1498
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  1499
lemma connected_ball [iff]:
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1500
  fixes x :: "'a::real_normed_vector"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1501
  shows "connected (ball x e)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1502
  using convex_connected convex_ball by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1503
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  1504
lemma connected_cball [iff]:
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1505
  fixes x :: "'a::real_normed_vector"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1506
  shows "connected (cball x e)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1507
  using convex_connected convex_cball by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1508
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1509
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  1510
subsection \<open>Convex hull\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1511
60762
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
  1512
lemma convex_convex_hull [iff]: "convex (convex hull s)"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1513
  unfolding hull_def
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1514
  using convex_Inter[of "{t. convex t \<and> s \<subseteq> t}"]
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44142
diff changeset
  1515
  by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1516
63016
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  1517
lemma convex_hull_subset:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  1518
    "s \<subseteq> convex hull t \<Longrightarrow> convex hull s \<subseteq> convex hull t"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  1519
  by (simp add: convex_convex_hull subset_hull)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  1520
34064
eee04bbbae7e avoid dependency on implicit dest rule predicate1D in proofs
haftmann
parents: 33758
diff changeset
  1521
lemma convex_hull_eq: "convex hull s = s \<longleftrightarrow> convex s"
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1522
  by (metis convex_convex_hull hull_same)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1523
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1524
lemma bounded_convex_hull:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1525
  fixes s :: "'a::real_normed_vector set"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1526
  assumes "bounded s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1527
  shows "bounded (convex hull s)"
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1528
proof -
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1529
  from assms obtain B where B: "\<forall>x\<in>s. norm x \<le> B"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1530
    unfolding bounded_iff by auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1531
  show ?thesis
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1532
    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
  1533
    unfolding subset_hull[of convex, OF convex_cball]
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1534
    unfolding subset_eq mem_cball dist_norm using B
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1535
    apply auto
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1536
    done
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1537
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1538
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1539
lemma finite_imp_bounded_convex_hull:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1540
  fixes s :: "'a::real_normed_vector set"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1541
  shows "finite s \<Longrightarrow> bounded (convex hull s)"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1542
  using bounded_convex_hull finite_imp_bounded
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1543
  by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1544
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1545
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  1546
subsubsection \<open>Convex hull is "preserved" by a linear function\<close>
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1547
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1548
lemma convex_hull_linear_image:
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1549
  assumes f: "linear f"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1550
  shows "f ` (convex hull s) = convex hull (f ` s)"
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1551
proof
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1552
  show "convex hull (f ` s) \<subseteq> f ` (convex hull s)"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1553
    by (intro hull_minimal image_mono hull_subset convex_linear_image assms convex_convex_hull)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1554
  show "f ` (convex hull s) \<subseteq> convex hull (f ` s)"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1555
  proof (unfold image_subset_iff_subset_vimage, rule hull_minimal)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1556
    show "s \<subseteq> f -` (convex hull (f ` s))"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1557
      by (fast intro: hull_inc)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1558
    show "convex (f -` (convex hull (f ` s)))"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1559
      by (intro convex_linear_vimage [OF f] convex_convex_hull)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1560
  qed
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1561
qed
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1562
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1563
lemma in_convex_hull_linear_image:
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1564
  assumes "linear f"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1565
    and "x \<in> convex hull s"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1566
  shows "f x \<in> convex hull (f ` s)"
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1567
  using convex_hull_linear_image[OF assms(1)] assms(2) by auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1568
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1569
lemma convex_hull_Times:
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1570
  "convex hull (s \<times> t) = (convex hull s) \<times> (convex hull t)"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1571
proof
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1572
  show "convex hull (s \<times> t) \<subseteq> (convex hull s) \<times> (convex hull t)"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1573
    by (intro hull_minimal Sigma_mono hull_subset convex_Times convex_convex_hull)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1574
  have "\<forall>x\<in>convex hull s. \<forall>y\<in>convex hull t. (x, y) \<in> convex hull (s \<times> t)"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1575
  proof (intro hull_induct)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1576
    fix x y assume "x \<in> s" and "y \<in> t"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1577
    then show "(x, y) \<in> convex hull (s \<times> t)"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1578
      by (simp add: hull_inc)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1579
  next
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1580
    fix x let ?S = "((\<lambda>y. (0, y)) -` (\<lambda>p. (- x, 0) + p) ` (convex hull s \<times> t))"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1581
    have "convex ?S"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1582
      by (intro convex_linear_vimage convex_translation convex_convex_hull,
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1583
        simp add: linear_iff)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1584
    also have "?S = {y. (x, y) \<in> convex hull (s \<times> t)}"
57865
dcfb33c26f50 tuned proofs -- fewer warnings;
wenzelm
parents: 57512
diff changeset
  1585
      by (auto simp add: image_def Bex_def)
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1586
    finally show "convex {y. (x, y) \<in> convex hull (s \<times> t)}" .
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1587
  next
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1588
    show "convex {x. \<forall>y\<in>convex hull t. (x, y) \<in> convex hull (s \<times> t)}"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1589
    proof (unfold Collect_ball_eq, rule convex_INT [rule_format])
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1590
      fix y let ?S = "((\<lambda>x. (x, 0)) -` (\<lambda>p. (0, - y) + p) ` (convex hull s \<times> t))"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1591
      have "convex ?S"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1592
      by (intro convex_linear_vimage convex_translation convex_convex_hull,
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1593
        simp add: linear_iff)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1594
      also have "?S = {x. (x, y) \<in> convex hull (s \<times> t)}"
57865
dcfb33c26f50 tuned proofs -- fewer warnings;
wenzelm
parents: 57512
diff changeset
  1595
        by (auto simp add: image_def Bex_def)
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1596
      finally show "convex {x. (x, y) \<in> convex hull (s \<times> t)}" .
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1597
    qed
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1598
  qed
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1599
  then show "(convex hull s) \<times> (convex hull t) \<subseteq> convex hull (s \<times> t)"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1600
    unfolding subset_eq split_paired_Ball_Sigma .
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1601
qed
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  1602
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  1603
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  1604
subsubsection \<open>Stepping theorems for convex hulls of finite sets\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1605
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1606
lemma convex_hull_empty[simp]: "convex hull {} = {}"
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1607
  by (rule hull_unique) auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1608
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1609
lemma convex_hull_singleton[simp]: "convex hull {a} = {a}"
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1610
  by (rule hull_unique) auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1611
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1612
lemma convex_hull_insert:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1613
  fixes s :: "'a::real_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1614
  assumes "s \<noteq> {}"
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1615
  shows "convex hull (insert a s) =
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1616
    {x. \<exists>u\<ge>0. \<exists>v\<ge>0. \<exists>b. (u + v = 1) \<and> b \<in> (convex hull s) \<and> (x = u *\<^sub>R a + v *\<^sub>R b)}"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1617
  (is "_ = ?hull")
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1618
  apply (rule, rule hull_minimal, rule)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1619
  unfolding insert_iff
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1620
  prefer 3
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1621
  apply rule
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1622
proof -
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1623
  fix x
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1624
  assume x: "x = a \<or> x \<in> s"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1625
  then show "x \<in> ?hull"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1626
    apply rule
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1627
    unfolding mem_Collect_eq
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1628
    apply (rule_tac x=1 in exI)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1629
    defer
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1630
    apply (rule_tac x=0 in exI)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1631
    using assms hull_subset[of s convex]
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1632
    apply auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1633
    done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1634
next
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1635
  fix x
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1636
  assume "x \<in> ?hull"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1637
  then obtain u v b where obt: "u\<ge>0" "v\<ge>0" "u + v = 1" "b \<in> convex hull s" "x = u *\<^sub>R a + v *\<^sub>R b"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1638
    by auto
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1639
  have "a \<in> convex hull insert a s" "b \<in> convex hull insert a s"
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1640
    using hull_mono[of s "insert a s" convex] hull_mono[of "{a}" "insert a s" convex] and obt(4)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1641
    by auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1642
  then show "x \<in> convex hull insert a s"
53676
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  1643
    unfolding obt(5) using obt(1-3)
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  1644
    by (rule convexD [OF convex_convex_hull])
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1645
next
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1646
  show "convex ?hull"
53676
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  1647
  proof (rule convexI)
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1648
    fix x y u v
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1649
    assume as: "(0::real) \<le> u" "0 \<le> v" "u + v = 1" "x\<in>?hull" "y\<in>?hull"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1650
    from as(4) obtain u1 v1 b1 where
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1651
      obt1: "u1\<ge>0" "v1\<ge>0" "u1 + v1 = 1" "b1 \<in> convex hull s" "x = u1 *\<^sub>R a + v1 *\<^sub>R b1"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1652
      by auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1653
    from as(5) obtain u2 v2 b2 where
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1654
      obt2: "u2\<ge>0" "v2\<ge>0" "u2 + v2 = 1" "b2 \<in> convex hull s" "y = u2 *\<^sub>R a + v2 *\<^sub>R b2"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1655
      by auto
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1656
    have *: "\<And>(x::'a) s1 s2. x - s1 *\<^sub>R x - s2 *\<^sub>R x = ((1::real) - (s1 + s2)) *\<^sub>R x"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1657
      by (auto simp add: algebra_simps)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1658
    have **: "\<exists>b \<in> convex hull s. u *\<^sub>R x + v *\<^sub>R y =
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1659
      (u * u1) *\<^sub>R a + (v * u2) *\<^sub>R a + (b - (u * u1) *\<^sub>R b - (v * u2) *\<^sub>R b)"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1660
    proof (cases "u * v1 + v * v2 = 0")
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1661
      case True
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1662
      have *: "\<And>(x::'a) s1 s2. x - s1 *\<^sub>R x - s2 *\<^sub>R x = ((1::real) - (s1 + s2)) *\<^sub>R x"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1663
        by (auto simp add: algebra_simps)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1664
      from True have ***: "u * v1 = 0" "v * v2 = 0"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  1665
        using mult_nonneg_nonneg[OF \<open>u\<ge>0\<close> \<open>v1\<ge>0\<close>] mult_nonneg_nonneg[OF \<open>v\<ge>0\<close> \<open>v2\<ge>0\<close>]
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1666
        by arith+
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1667
      then have "u * u1 + v * u2 = 1"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1668
        using as(3) obt1(3) obt2(3) by auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1669
      then show ?thesis
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1670
        unfolding obt1(5) obt2(5) *
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1671
        using assms hull_subset[of s convex]
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1672
        by (auto simp add: *** scaleR_right_distrib)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1673
    next
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1674
      case False
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1675
      have "1 - (u * u1 + v * u2) = (u + v) - (u * u1 + v * u2)"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1676
        using as(3) obt1(3) obt2(3) by (auto simp add: field_simps)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1677
      also have "\<dots> = u * (v1 + u1 - u1) + v * (v2 + u2 - u2)"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1678
        using as(3) obt1(3) obt2(3) by (auto simp add: field_simps)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1679
      also have "\<dots> = u * v1 + v * v2"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1680
        by simp
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1681
      finally have **:"1 - (u * u1 + v * u2) = u * v1 + v * v2" by auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1682
      have "0 \<le> u * v1 + v * v2" "0 \<le> u * v1" "0 \<le> u * v1 + v * v2" "0 \<le> v * v2"
56536
aefb4a8da31f made mult_nonneg_nonneg a simp rule
nipkow
parents: 56480
diff changeset
  1683
        using as(1,2) obt1(1,2) obt2(1,2) by auto
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1684
      then show ?thesis
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1685
        unfolding obt1(5) obt2(5)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1686
        unfolding * and **
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1687
        using False
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1688
        apply (rule_tac
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1689
          x = "((u * v1) / (u * v1 + v * v2)) *\<^sub>R b1 + ((v * v2) / (u * v1 + v * v2)) *\<^sub>R b2" in bexI)
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1690
        defer
53676
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  1691
        apply (rule convexD [OF convex_convex_hull])
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1692
        using obt1(4) obt2(4)
49530
wenzelm
parents: 49529
diff changeset
  1693
        unfolding add_divide_distrib[symmetric] and zero_le_divide_iff
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1694
        apply (auto simp add: scaleR_left_distrib scaleR_right_distrib)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1695
        done
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1696
    qed
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1697
    have u1: "u1 \<le> 1"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1698
      unfolding obt1(3)[symmetric] and not_le using obt1(2) by auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1699
    have u2: "u2 \<le> 1"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1700
      unfolding obt2(3)[symmetric] and not_le using obt2(2) by auto
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1701
    have "u1 * u + u2 * v \<le> max u1 u2 * u + max u1 u2 * v"
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1702
      apply (rule add_mono)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1703
      apply (rule_tac [!] mult_right_mono)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1704
      using as(1,2) obt1(1,2) obt2(1,2)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1705
      apply auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1706
      done
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1707
    also have "\<dots> \<le> 1"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1708
      unfolding distrib_left[symmetric] and as(3) using u1 u2 by auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1709
    finally show "u *\<^sub>R x + v *\<^sub>R y \<in> ?hull"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1710
      unfolding mem_Collect_eq
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1711
      apply (rule_tac x="u * u1 + v * u2" in exI)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1712
      apply (rule conjI)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1713
      defer
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1714
      apply (rule_tac x="1 - u * u1 - v * u2" in exI)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1715
      unfolding Bex_def
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1716
      using as(1,2) obt1(1,2) obt2(1,2) **
56536
aefb4a8da31f made mult_nonneg_nonneg a simp rule
nipkow
parents: 56480
diff changeset
  1717
      apply (auto simp add: algebra_simps)
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1718
      done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1719
  qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1720
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1721
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1722
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  1723
subsubsection \<open>Explicit expression for convex hull\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1724
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1725
lemma convex_hull_indexed:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1726
  fixes s :: "'a::real_vector set"
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1727
  shows "convex hull s =
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1728
    {y. \<exists>k u x.
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1729
      (\<forall>i\<in>{1::nat .. k}. 0 \<le> u i \<and> x i \<in> s) \<and>
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1730
      (setsum u {1..k} = 1) \<and> (setsum (\<lambda>i. u i *\<^sub>R x i) {1..k} = y)}"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1731
  (is "?xyz = ?hull")
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1732
  apply (rule hull_unique)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1733
  apply rule
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1734
  defer
53676
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  1735
  apply (rule convexI)
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1736
proof -
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1737
  fix x
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1738
  assume "x\<in>s"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1739
  then show "x \<in> ?hull"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1740
    unfolding mem_Collect_eq
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1741
    apply (rule_tac x=1 in exI, rule_tac x="\<lambda>x. 1" in exI)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1742
    apply auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1743
    done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1744
next
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1745
  fix t
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1746
  assume as: "s \<subseteq> t" "convex t"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1747
  show "?hull \<subseteq> t"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1748
    apply rule
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1749
    unfolding mem_Collect_eq
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1750
    apply (elim exE conjE)
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1751
  proof -
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1752
    fix x k u y
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1753
    assume assm:
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1754
      "\<forall>i\<in>{1::nat..k}. 0 \<le> u i \<and> y i \<in> s"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1755
      "setsum u {1..k} = 1" "(\<Sum>i = 1..k. u i *\<^sub>R y i) = x"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1756
    show "x\<in>t"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1757
      unfolding assm(3) [symmetric]
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1758
      apply (rule as(2)[unfolded convex, rule_format])
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1759
      using assm(1,2) as(1) apply auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1760
      done
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1761
  qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1762
next
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1763
  fix x y u v
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1764
  assume uv: "0 \<le> u" "0 \<le> v" "u + v = (1::real)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1765
  assume xy: "x \<in> ?hull" "y \<in> ?hull"
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1766
  from xy obtain k1 u1 x1 where
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1767
    x: "\<forall>i\<in>{1::nat..k1}. 0\<le>u1 i \<and> x1 i \<in> s" "setsum u1 {Suc 0..k1} = 1" "(\<Sum>i = Suc 0..k1. u1 i *\<^sub>R x1 i) = x"
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1768
    by auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1769
  from xy obtain k2 u2 x2 where
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1770
    y: "\<forall>i\<in>{1::nat..k2}. 0\<le>u2 i \<and> x2 i \<in> s" "setsum u2 {Suc 0..k2} = 1" "(\<Sum>i = Suc 0..k2. u2 i *\<^sub>R x2 i) = y"
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1771
    by auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1772
  have *: "\<And>P (x1::'a) x2 s1 s2 i.
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1773
    (if P i then s1 else s2) *\<^sub>R (if P i then x1 else x2) = (if P i then s1 *\<^sub>R x1 else s2 *\<^sub>R x2)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1774
    "{1..k1 + k2} \<inter> {1..k1} = {1..k1}" "{1..k1 + k2} \<inter> - {1..k1} = (\<lambda>i. i + k1) ` {1..k2}"
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1775
    prefer 3
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1776
    apply (rule, rule)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1777
    unfolding image_iff
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1778
    apply (rule_tac x = "x - k1" in bexI)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1779
    apply (auto simp add: not_le)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1780
    done
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1781
  have inj: "inj_on (\<lambda>i. i + k1) {1..k2}"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1782
    unfolding inj_on_def by auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1783
  show "u *\<^sub>R x + v *\<^sub>R y \<in> ?hull"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1784
    apply rule
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1785
    apply (rule_tac x="k1 + k2" in exI)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1786
    apply (rule_tac x="\<lambda>i. if i \<in> {1..k1} then u * u1 i else v * u2 (i - k1)" in exI)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1787
    apply (rule_tac x="\<lambda>i. if i \<in> {1..k1} then x1 i else x2 (i - k1)" in exI)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1788
    apply (rule, rule)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1789
    defer
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1790
    apply rule
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  1791
    unfolding * and setsum.If_cases[OF finite_atLeastAtMost[of 1 "k1 + k2"]] and
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  1792
      setsum.reindex[OF inj] and o_def Collect_mem_eq
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1793
    unfolding scaleR_scaleR[symmetric] scaleR_right.setsum [symmetric] setsum_right_distrib[symmetric]
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1794
  proof -
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1795
    fix i
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1796
    assume i: "i \<in> {1..k1+k2}"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1797
    show "0 \<le> (if i \<in> {1..k1} then u * u1 i else v * u2 (i - k1)) \<and>
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1798
      (if i \<in> {1..k1} then x1 i else x2 (i - k1)) \<in> s"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1799
    proof (cases "i\<in>{1..k1}")
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1800
      case True
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1801
      then show ?thesis
56536
aefb4a8da31f made mult_nonneg_nonneg a simp rule
nipkow
parents: 56480
diff changeset
  1802
        using uv(1) x(1)[THEN bspec[where x=i]] by auto
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1803
    next
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1804
      case False
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  1805
      define j where "j = i - k1"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1806
      from i False have "j \<in> {1..k2}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1807
        unfolding j_def by auto
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1808
      then show ?thesis
56536
aefb4a8da31f made mult_nonneg_nonneg a simp rule
nipkow
parents: 56480
diff changeset
  1809
        using False uv(2) y(1)[THEN bspec[where x=j]]
aefb4a8da31f made mult_nonneg_nonneg a simp rule
nipkow
parents: 56480
diff changeset
  1810
        by (auto simp: j_def[symmetric])
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1811
    qed
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1812
  qed (auto simp add: not_le x(2,3) y(2,3) uv(3))
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1813
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1814
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1815
lemma convex_hull_finite:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1816
  fixes s :: "'a::real_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1817
  assumes "finite s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1818
  shows "convex hull s = {y. \<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and>
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1819
    setsum u s = 1 \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y}"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1820
  (is "?HULL = ?set")
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1821
proof (rule hull_unique, auto simp add: convex_def[of ?set])
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1822
  fix x
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1823
  assume "x \<in> s"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1824
  then show "\<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> (\<Sum>x\<in>s. u x *\<^sub>R x) = x"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1825
    apply (rule_tac x="\<lambda>y. if x=y then 1 else 0" in exI)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1826
    apply auto
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  1827
    unfolding setsum.delta'[OF assms] and setsum_delta''[OF assms]
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1828
    apply auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1829
    done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1830
next
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1831
  fix u v :: real
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1832
  assume uv: "0 \<le> u" "0 \<le> v" "u + v = 1"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1833
  fix ux assume ux: "\<forall>x\<in>s. 0 \<le> ux x" "setsum ux s = (1::real)"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1834
  fix uy assume uy: "\<forall>x\<in>s. 0 \<le> uy x" "setsum uy s = (1::real)"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1835
  {
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1836
    fix x
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1837
    assume "x\<in>s"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1838
    then have "0 \<le> u * ux x + v * uy x"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1839
      using ux(1)[THEN bspec[where x=x]] uy(1)[THEN bspec[where x=x]] and uv(1,2)
56536
aefb4a8da31f made mult_nonneg_nonneg a simp rule
nipkow
parents: 56480
diff changeset
  1840
      by auto
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1841
  }
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1842
  moreover
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1843
  have "(\<Sum>x\<in>s. u * ux x + v * uy x) = 1"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  1844
    unfolding setsum.distrib and setsum_right_distrib[symmetric] and ux(2) uy(2)
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1845
    using uv(3) by auto
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1846
  moreover
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1847
  have "(\<Sum>x\<in>s. (u * ux x + v * uy x) *\<^sub>R x) = u *\<^sub>R (\<Sum>x\<in>s. ux x *\<^sub>R x) + v *\<^sub>R (\<Sum>x\<in>s. uy x *\<^sub>R x)"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  1848
    unfolding scaleR_left_distrib and setsum.distrib and scaleR_scaleR[symmetric]
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1849
      and scaleR_right.setsum [symmetric]
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1850
    by auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1851
  ultimately
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1852
  show "\<exists>uc. (\<forall>x\<in>s. 0 \<le> uc x) \<and> setsum uc s = 1 \<and>
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1853
      (\<Sum>x\<in>s. uc x *\<^sub>R x) = u *\<^sub>R (\<Sum>x\<in>s. ux x *\<^sub>R x) + v *\<^sub>R (\<Sum>x\<in>s. uy x *\<^sub>R x)"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1854
    apply (rule_tac x="\<lambda>x. u * ux x + v * uy x" in exI)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1855
    apply auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1856
    done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1857
next
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1858
  fix t
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1859
  assume t: "s \<subseteq> t" "convex t"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1860
  fix u
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1861
  assume u: "\<forall>x\<in>s. 0 \<le> u x" "setsum u s = (1::real)"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1862
  then show "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> t"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1863
    using t(2)[unfolded convex_explicit, THEN spec[where x=s], THEN spec[where x=u]]
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1864
    using assms and t(1) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1865
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1866
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1867
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  1868
subsubsection \<open>Another formulation from Lars Schewe\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1869
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1870
lemma convex_hull_explicit:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1871
  fixes p :: "'a::real_vector set"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1872
  shows "convex hull p =
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1873
    {y. \<exists>s u. finite s \<and> s \<subseteq> p \<and> (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> setsum (\<lambda>v. u v *\<^sub>R v) s = y}"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1874
  (is "?lhs = ?rhs")
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1875
proof -
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1876
  {
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1877
    fix x
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1878
    assume "x\<in>?lhs"
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1879
    then obtain k u y where
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1880
        obt: "\<forall>i\<in>{1::nat..k}. 0 \<le> u i \<and> y i \<in> p" "setsum u {1..k} = 1" "(\<Sum>i = 1..k. u i *\<^sub>R y i) = x"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1881
      unfolding convex_hull_indexed by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1882
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1883
    have fin: "finite {1..k}" by auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1884
    have fin': "\<And>v. finite {i \<in> {1..k}. y i = v}" by auto
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1885
    {
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1886
      fix j
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1887
      assume "j\<in>{1..k}"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1888
      then have "y j \<in> p" "0 \<le> setsum u {i. Suc 0 \<le> i \<and> i \<le> k \<and> y i = y j}"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1889
        using obt(1)[THEN bspec[where x=j]] and obt(2)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1890
        apply simp
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1891
        apply (rule setsum_nonneg)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1892
        using obt(1)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1893
        apply auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1894
        done
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1895
    }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1896
    moreover
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  1897
    have "(\<Sum>v\<in>y ` {1..k}. setsum u {i \<in> {1..k}. y i = v}) = 1"
49530
wenzelm
parents: 49529
diff changeset
  1898
      unfolding setsum_image_gen[OF fin, symmetric] using obt(2) by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1899
    moreover have "(\<Sum>v\<in>y ` {1..k}. setsum u {i \<in> {1..k}. y i = v} *\<^sub>R v) = x"
49530
wenzelm
parents: 49529
diff changeset
  1900
      using setsum_image_gen[OF fin, of "\<lambda>i. u i *\<^sub>R y i" y, symmetric]
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1901
      unfolding scaleR_left.setsum using obt(3) by auto
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1902
    ultimately
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1903
    have "\<exists>s u. finite s \<and> s \<subseteq> p \<and> (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1904
      apply (rule_tac x="y ` {1..k}" in exI)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1905
      apply (rule_tac x="\<lambda>v. setsum u {i\<in>{1..k}. y i = v}" in exI)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1906
      apply auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1907
      done
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1908
    then have "x\<in>?rhs" by auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1909
  }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1910
  moreover
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1911
  {
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1912
    fix y
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1913
    assume "y\<in>?rhs"
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1914
    then obtain s u where
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1915
      obt: "finite s" "s \<subseteq> p" "\<forall>x\<in>s. 0 \<le> u x" "setsum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = y"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1916
      by auto
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1917
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1918
    obtain f where f: "inj_on f {1..card s}" "f ` {1..card s} = s"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1919
      using ex_bij_betw_nat_finite_1[OF obt(1)] unfolding bij_betw_def by auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1920
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1921
    {
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1922
      fix i :: nat
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1923
      assume "i\<in>{1..card s}"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1924
      then have "f i \<in> s"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1925
        apply (subst f(2)[symmetric])
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1926
        apply auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1927
        done
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1928
      then have "0 \<le> u (f i)" "f i \<in> p" using obt(2,3) by auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1929
    }
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  1930
    moreover have *: "finite {1..card s}" by auto
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1931
    {
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1932
      fix y
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1933
      assume "y\<in>s"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1934
      then obtain i where "i\<in>{1..card s}" "f i = y"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1935
        using f using image_iff[of y f "{1..card s}"]
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1936
        by auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1937
      then have "{x. Suc 0 \<le> x \<and> x \<le> card s \<and> f x = y} = {i}"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1938
        apply auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1939
        using f(1)[unfolded inj_on_def]
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1940
        apply(erule_tac x=x in ballE)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1941
        apply auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1942
        done
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1943
      then have "card {x. Suc 0 \<le> x \<and> x \<le> card s \<and> f x = y} = 1" by auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1944
      then have "(\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x)) = u y"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1945
          "(\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x) *\<^sub>R f x) = u y *\<^sub>R y"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1946
        by (auto simp add: setsum_constant_scaleR)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1947
    }
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1948
    then have "(\<Sum>x = 1..card s. u (f x)) = 1" "(\<Sum>i = 1..card s. u (f i) *\<^sub>R f i) = y"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1949
      unfolding setsum_image_gen[OF *(1), of "\<lambda>x. u (f x) *\<^sub>R f x" f]
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1950
        and setsum_image_gen[OF *(1), of "\<lambda>x. u (f x)" f]
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1951
      unfolding f
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  1952
      using setsum.cong [of s s "\<lambda>y. (\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x) *\<^sub>R f x)" "\<lambda>v. u v *\<^sub>R v"]
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  1953
      using setsum.cong [of s s "\<lambda>y. (\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x))" u]
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1954
      unfolding obt(4,5)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1955
      by auto
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1956
    ultimately
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1957
    have "\<exists>k u x. (\<forall>i\<in>{1..k}. 0 \<le> u i \<and> x i \<in> p) \<and> setsum u {1..k} = 1 \<and>
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1958
        (\<Sum>i::nat = 1..k. u i *\<^sub>R x i) = y"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1959
      apply (rule_tac x="card s" in exI)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1960
      apply (rule_tac x="u \<circ> f" in exI)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1961
      apply (rule_tac x=f in exI)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1962
      apply fastforce
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1963
      done
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1964
    then have "y \<in> ?lhs"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1965
      unfolding convex_hull_indexed by auto
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1966
  }
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1967
  ultimately show ?thesis
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1968
    unfolding set_eq_iff by blast
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1969
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1970
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1971
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  1972
subsubsection \<open>A stepping theorem for that expansion\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1973
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1974
lemma convex_hull_finite_step:
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1975
  fixes s :: "'a::real_vector set"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1976
  assumes "finite s"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1977
  shows
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1978
    "(\<exists>u. (\<forall>x\<in>insert a s. 0 \<le> u x) \<and> setsum u (insert a s) = w \<and> setsum (\<lambda>x. u x *\<^sub>R x) (insert a s) = y)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1979
      \<longleftrightarrow> (\<exists>v\<ge>0. \<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = w - v \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y - v *\<^sub>R a)"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1980
  (is "?lhs = ?rhs")
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1981
proof (rule, case_tac[!] "a\<in>s")
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1982
  assume "a \<in> s"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  1983
  then have *: "insert a s = s" by auto
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1984
  assume ?lhs
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1985
  then show ?rhs
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1986
    unfolding *
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1987
    apply (rule_tac x=0 in exI)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1988
    apply auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1989
    done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  1990
next
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1991
  assume ?lhs
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1992
  then obtain u where
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1993
      u: "\<forall>x\<in>insert a s. 0 \<le> u x" "setsum u (insert a s) = w" "(\<Sum>x\<in>insert a s. u x *\<^sub>R x) = y"
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1994
    by auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1995
  assume "a \<notin> s"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1996
  then show ?rhs
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1997
    apply (rule_tac x="u a" in exI)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1998
    using u(1)[THEN bspec[where x=a]]
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  1999
    apply simp
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  2000
    apply (rule_tac x=u in exI)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2001
    using u[unfolded setsum_clauses(2)[OF assms]] and \<open>a\<notin>s\<close>
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  2002
    apply auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  2003
    done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2004
next
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  2005
  assume "a \<in> s"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  2006
  then have *: "insert a s = s" by auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  2007
  have fin: "finite (insert a s)" using assms by auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  2008
  assume ?rhs
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  2009
  then obtain v u where uv: "v\<ge>0" "\<forall>x\<in>s. 0 \<le> u x" "setsum u s = w - v" "(\<Sum>x\<in>s. u x *\<^sub>R x) = y - v *\<^sub>R a"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  2010
    by auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  2011
  show ?lhs
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  2012
    apply (rule_tac x = "\<lambda>x. (if a = x then v else 0) + u x" in exI)
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  2013
    unfolding scaleR_left_distrib and setsum.distrib and setsum_delta''[OF fin] and setsum.delta'[OF fin]
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  2014
    unfolding setsum_clauses(2)[OF assms]
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2015
    using uv and uv(2)[THEN bspec[where x=a]] and \<open>a\<in>s\<close>
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  2016
    apply auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  2017
    done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2018
next
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  2019
  assume ?rhs
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2020
  then obtain v u where
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2021
    uv: "v\<ge>0" "\<forall>x\<in>s. 0 \<le> u x" "setsum u s = w - v" "(\<Sum>x\<in>s. u x *\<^sub>R x) = y - v *\<^sub>R a"
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  2022
    by auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  2023
  moreover
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  2024
  assume "a \<notin> s"
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  2025
  moreover
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2026
  have "(\<Sum>x\<in>s. if a = x then v else u x) = setsum u s"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2027
    and "(\<Sum>x\<in>s. (if a = x then v else u x) *\<^sub>R x) = (\<Sum>x\<in>s. u x *\<^sub>R x)"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  2028
    apply (rule_tac setsum.cong) apply rule
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  2029
    defer
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  2030
    apply (rule_tac setsum.cong) apply rule
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2031
    using \<open>a \<notin> s\<close>
50804
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  2032
    apply auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  2033
    done
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  2034
  ultimately show ?lhs
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  2035
    apply (rule_tac x="\<lambda>x. if a = x then v else u x" in exI)
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  2036
    unfolding setsum_clauses(2)[OF assms]
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  2037
    apply auto
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  2038
    done
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  2039
qed
4156a45aeb63 tuned proofs;
wenzelm
parents: 50526
diff changeset
  2040
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2041
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2042
subsubsection \<open>Hence some special cases\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2043
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2044
lemma convex_hull_2:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2045
  "convex hull {a,b} = {u *\<^sub>R a + v *\<^sub>R b | u v. 0 \<le> u \<and> 0 \<le> v \<and> u + v = 1}"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2046
proof -
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2047
  have *: "\<And>u. (\<forall>x\<in>{a, b}. 0 \<le> u x) \<longleftrightarrow> 0 \<le> u a \<and> 0 \<le> u b"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2048
    by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2049
  have **: "finite {b}" by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2050
  show ?thesis
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2051
    apply (simp add: convex_hull_finite)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2052
    unfolding convex_hull_finite_step[OF **, of a 1, unfolded * conj_assoc]
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2053
    apply auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2054
    apply (rule_tac x=v in exI)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2055
    apply (rule_tac x="1 - v" in exI)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2056
    apply simp
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2057
    apply (rule_tac x=u in exI)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2058
    apply simp
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2059
    apply (rule_tac x="\<lambda>x. v" in exI)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2060
    apply simp
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2061
    done
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2062
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2063
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2064
lemma convex_hull_2_alt: "convex hull {a,b} = {a + u *\<^sub>R (b - a) | u.  0 \<le> u \<and> u \<le> 1}"
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44142
diff changeset
  2065
  unfolding convex_hull_2
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2066
proof (rule Collect_cong)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2067
  have *: "\<And>x y ::real. x + y = 1 \<longleftrightarrow> x = 1 - y"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2068
    by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2069
  fix x
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2070
  show "(\<exists>v u. x = v *\<^sub>R a + u *\<^sub>R b \<and> 0 \<le> v \<and> 0 \<le> u \<and> v + u = 1) \<longleftrightarrow>
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2071
    (\<exists>u. x = a + u *\<^sub>R (b - a) \<and> 0 \<le> u \<and> u \<le> 1)"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2072
    unfolding *
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2073
    apply auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2074
    apply (rule_tac[!] x=u in exI)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2075
    apply (auto simp add: algebra_simps)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2076
    done
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2077
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2078
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2079
lemma convex_hull_3:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2080
  "convex hull {a,b,c} = { u *\<^sub>R a + v *\<^sub>R b + w *\<^sub>R c | u v w. 0 \<le> u \<and> 0 \<le> v \<and> 0 \<le> w \<and> u + v + w = 1}"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2081
proof -
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2082
  have fin: "finite {a,b,c}" "finite {b,c}" "finite {c}"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2083
    by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2084
  have *: "\<And>x y z ::real. x + y + z = 1 \<longleftrightarrow> x = 1 - y - z"
44361
75ec83d45303 remove unnecessary euclidean_space class constraints
huffman
parents: 44349
diff changeset
  2085
    by (auto simp add: field_simps)
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2086
  show ?thesis
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2087
    unfolding convex_hull_finite[OF fin(1)] and convex_hull_finite_step[OF fin(2)] and *
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2088
    unfolding convex_hull_finite_step[OF fin(3)]
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2089
    apply (rule Collect_cong)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2090
    apply simp
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2091
    apply auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2092
    apply (rule_tac x=va in exI)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2093
    apply (rule_tac x="u c" in exI)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2094
    apply simp
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2095
    apply (rule_tac x="1 - v - w" in exI)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2096
    apply simp
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2097
    apply (rule_tac x=v in exI)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2098
    apply simp
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2099
    apply (rule_tac x="\<lambda>x. w" in exI)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2100
    apply simp
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2101
    done
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2102
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2103
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2104
lemma convex_hull_3_alt:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2105
  "convex hull {a,b,c} = {a + u *\<^sub>R (b - a) + v *\<^sub>R (c - a) | u v.  0 \<le> u \<and> 0 \<le> v \<and> u + v \<le> 1}"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2106
proof -
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2107
  have *: "\<And>x y z ::real. x + y + z = 1 \<longleftrightarrow> x = 1 - y - z"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2108
    by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2109
  show ?thesis
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2110
    unfolding convex_hull_3
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2111
    apply (auto simp add: *)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2112
    apply (rule_tac x=v in exI)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2113
    apply (rule_tac x=w in exI)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2114
    apply (simp add: algebra_simps)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2115
    apply (rule_tac x=u in exI)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2116
    apply (rule_tac x=v in exI)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2117
    apply (simp add: algebra_simps)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2118
    done
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2119
qed
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2120
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2121
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2122
subsection \<open>Relations among closure notions and corresponding hulls\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2123
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2124
lemma affine_imp_convex: "affine s \<Longrightarrow> convex s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2125
  unfolding affine_def convex_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2126
44361
75ec83d45303 remove unnecessary euclidean_space class constraints
huffman
parents: 44349
diff changeset
  2127
lemma subspace_imp_convex: "subspace s \<Longrightarrow> convex s"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2128
  using subspace_imp_affine affine_imp_convex by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2129
44361
75ec83d45303 remove unnecessary euclidean_space class constraints
huffman
parents: 44349
diff changeset
  2130
lemma affine_hull_subset_span: "(affine hull s) \<subseteq> (span s)"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2131
  by (metis hull_minimal span_inc subspace_imp_affine subspace_span)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2132
44361
75ec83d45303 remove unnecessary euclidean_space class constraints
huffman
parents: 44349
diff changeset
  2133
lemma convex_hull_subset_span: "(convex hull s) \<subseteq> (span s)"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2134
  by (metis hull_minimal span_inc subspace_imp_convex subspace_span)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2135
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2136
lemma convex_hull_subset_affine_hull: "(convex hull s) \<subseteq> (affine hull s)"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2137
  by (metis affine_affine_hull affine_imp_convex hull_minimal hull_subset)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2138
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2139
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2140
lemma affine_dependent_imp_dependent: "affine_dependent s \<Longrightarrow> dependent s"
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  2141
  unfolding affine_dependent_def dependent_def
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2142
  using affine_hull_subset_span by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2143
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2144
lemma dependent_imp_affine_dependent:
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2145
  assumes "dependent {x - a| x . x \<in> s}"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2146
    and "a \<notin> s"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2147
  shows "affine_dependent (insert a s)"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2148
proof -
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  2149
  from assms(1)[unfolded dependent_explicit] obtain S u v
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2150
    where obt: "finite S" "S \<subseteq> {x - a |x. x \<in> s}" "v\<in>S" "u v  \<noteq> 0" "(\<Sum>v\<in>S. u v *\<^sub>R v) = 0"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2151
    by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  2152
  define t where "t = (\<lambda>x. x + a) ` S"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2153
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2154
  have inj: "inj_on (\<lambda>x. x + a) S"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2155
    unfolding inj_on_def by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2156
  have "0 \<notin> S"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2157
    using obt(2) assms(2) unfolding subset_eq by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2158
  have fin: "finite t" and "t \<subseteq> s"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2159
    unfolding t_def using obt(1,2) by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2160
  then have "finite (insert a t)" and "insert a t \<subseteq> insert a s"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2161
    by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2162
  moreover have *: "\<And>P Q. (\<Sum>x\<in>t. (if x = a then P x else Q x)) = (\<Sum>x\<in>t. Q x)"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  2163
    apply (rule setsum.cong)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2164
    using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close>
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2165
    apply auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2166
    done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2167
  have "(\<Sum>x\<in>insert a t. if x = a then - (\<Sum>x\<in>t. u (x - a)) else u (x - a)) = 0"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2168
    unfolding setsum_clauses(2)[OF fin]
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2169
    using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close>
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2170
    apply auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2171
    unfolding *
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2172
    apply auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2173
    done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2174
  moreover have "\<exists>v\<in>insert a t. (if v = a then - (\<Sum>x\<in>t. u (x - a)) else u (v - a)) \<noteq> 0"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2175
    apply (rule_tac x="v + a" in bexI)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2176
    using obt(3,4) and \<open>0\<notin>S\<close>
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2177
    unfolding t_def
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2178
    apply auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2179
    done
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2180
  moreover have *: "\<And>P Q. (\<Sum>x\<in>t. (if x = a then P x else Q x) *\<^sub>R x) = (\<Sum>x\<in>t. Q x *\<^sub>R x)"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  2181
    apply (rule setsum.cong)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2182
    using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close>
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2183
    apply auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2184
    done
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  2185
  have "(\<Sum>x\<in>t. u (x - a)) *\<^sub>R a = (\<Sum>v\<in>t. u (v - a) *\<^sub>R v)"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2186
    unfolding scaleR_left.setsum
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  2187
    unfolding t_def and setsum.reindex[OF inj] and o_def
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2188
    using obt(5)
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  2189
    by (auto simp add: setsum.distrib scaleR_right_distrib)
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2190
  then have "(\<Sum>v\<in>insert a t. (if v = a then - (\<Sum>x\<in>t. u (x - a)) else u (v - a)) *\<^sub>R v) = 0"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2191
    unfolding setsum_clauses(2)[OF fin]
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2192
    using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close>
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2193
    by (auto simp add: *)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2194
  ultimately show ?thesis
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2195
    unfolding affine_dependent_explicit
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2196
    apply (rule_tac x="insert a t" in exI)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2197
    apply auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2198
    done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2199
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2200
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2201
lemma convex_cone:
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2202
  "convex s \<and> cone s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. (x + y) \<in> s) \<and> (\<forall>x\<in>s. \<forall>c\<ge>0. (c *\<^sub>R x) \<in> s)"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2203
  (is "?lhs = ?rhs")
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2204
proof -
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2205
  {
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2206
    fix x y
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2207
    assume "x\<in>s" "y\<in>s" and ?lhs
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2208
    then have "2 *\<^sub>R x \<in>s" "2 *\<^sub>R y \<in> s"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2209
      unfolding cone_def by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2210
    then have "x + y \<in> s"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2211
      using \<open>?lhs\<close>[unfolded convex_def, THEN conjunct1]
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2212
      apply (erule_tac x="2*\<^sub>R x" in ballE)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2213
      apply (erule_tac x="2*\<^sub>R y" in ballE)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2214
      apply (erule_tac x="1/2" in allE)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2215
      apply simp
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2216
      apply (erule_tac x="1/2" in allE)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2217
      apply auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2218
      done
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2219
  }
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2220
  then show ?thesis
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2221
    unfolding convex_def cone_def by blast
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2222
qed
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2223
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2224
lemma affine_dependent_biggerset:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2225
  fixes s :: "'a::euclidean_space set"
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 36844
diff changeset
  2226
  assumes "finite s" "card s \<ge> DIM('a) + 2"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2227
  shows "affine_dependent s"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2228
proof -
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2229
  have "s \<noteq> {}" using assms by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2230
  then obtain a where "a\<in>s" by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2231
  have *: "{x - a |x. x \<in> s - {a}} = (\<lambda>x. x - a) ` (s - {a})"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2232
    by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2233
  have "card {x - a |x. x \<in> s - {a}} = card (s - {a})"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2234
    unfolding *
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2235
    apply (rule card_image)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2236
    unfolding inj_on_def
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2237
    apply auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2238
    done
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 36844
diff changeset
  2239
  also have "\<dots> > DIM('a)" using assms(2)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2240
    unfolding card_Diff_singleton[OF assms(1) \<open>a\<in>s\<close>] by auto
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2241
  finally show ?thesis
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2242
    apply (subst insert_Diff[OF \<open>a\<in>s\<close>, symmetric])
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2243
    apply (rule dependent_imp_affine_dependent)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2244
    apply (rule dependent_biggerset)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2245
    apply auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2246
    done
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2247
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2248
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2249
lemma affine_dependent_biggerset_general:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2250
  assumes "finite (s :: 'a::euclidean_space set)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2251
    and "card s \<ge> dim s + 2"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2252
  shows "affine_dependent s"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2253
proof -
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2254
  from assms(2) have "s \<noteq> {}" by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2255
  then obtain a where "a\<in>s" by auto
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2256
  have *: "{x - a |x. x \<in> s - {a}} = (\<lambda>x. x - a) ` (s - {a})"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2257
    by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2258
  have **: "card {x - a |x. x \<in> s - {a}} = card (s - {a})"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2259
    unfolding *
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2260
    apply (rule card_image)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2261
    unfolding inj_on_def
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2262
    apply auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2263
    done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2264
  have "dim {x - a |x. x \<in> s - {a}} \<le> dim s"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2265
    apply (rule subset_le_dim)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2266
    unfolding subset_eq
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2267
    using \<open>a\<in>s\<close>
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2268
    apply (auto simp add:span_superset span_sub)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2269
    done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2270
  also have "\<dots> < dim s + 1" by auto
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2271
  also have "\<dots> \<le> card (s - {a})"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2272
    using assms
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2273
    using card_Diff_singleton[OF assms(1) \<open>a\<in>s\<close>]
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2274
    by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2275
  finally show ?thesis
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2276
    apply (subst insert_Diff[OF \<open>a\<in>s\<close>, symmetric])
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2277
    apply (rule dependent_imp_affine_dependent)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2278
    apply (rule dependent_biggerset_general)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2279
    unfolding **
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2280
    apply auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2281
    done
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2282
qed
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2283
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  2284
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2285
subsection \<open>Some Properties of Affine Dependent Sets\<close>
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2286
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2287
lemma affine_independent_0: "\<not> affine_dependent {}"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2288
  by (simp add: affine_dependent_def)
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2289
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2290
lemma affine_independent_1: "\<not> affine_dependent {a}"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2291
  by (simp add: affine_dependent_def)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2292
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2293
lemma affine_independent_2: "\<not> affine_dependent {a,b}"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2294
  by (simp add: affine_dependent_def insert_Diff_if hull_same)
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2295
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2296
lemma affine_hull_translation: "affine hull ((\<lambda>x. a + x) `  S) = (\<lambda>x. a + x) ` (affine hull S)"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2297
proof -
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2298
  have "affine ((\<lambda>x. a + x) ` (affine hull S))"
60303
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60176
diff changeset
  2299
    using affine_translation affine_affine_hull by blast
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2300
  moreover have "(\<lambda>x. a + x) `  S \<subseteq> (\<lambda>x. a + x) ` (affine hull S)"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2301
    using hull_subset[of S] by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2302
  ultimately have h1: "affine hull ((\<lambda>x. a + x) `  S) \<subseteq> (\<lambda>x. a + x) ` (affine hull S)"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2303
    by (metis hull_minimal)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2304
  have "affine((\<lambda>x. -a + x) ` (affine hull ((\<lambda>x. a + x) `  S)))"
60303
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60176
diff changeset
  2305
    using affine_translation affine_affine_hull by blast
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2306
  moreover have "(\<lambda>x. -a + x) ` (\<lambda>x. a + x) `  S \<subseteq> (\<lambda>x. -a + x) ` (affine hull ((\<lambda>x. a + x) `  S))"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2307
    using hull_subset[of "(\<lambda>x. a + x) `  S"] by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2308
  moreover have "S = (\<lambda>x. -a + x) ` (\<lambda>x. a + x) `  S"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2309
    using translation_assoc[of "-a" a] by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2310
  ultimately have "(\<lambda>x. -a + x) ` (affine hull ((\<lambda>x. a + x) `  S)) >= (affine hull S)"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2311
    by (metis hull_minimal)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2312
  then have "affine hull ((\<lambda>x. a + x) ` S) >= (\<lambda>x. a + x) ` (affine hull S)"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2313
    by auto
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2314
  then show ?thesis using h1 by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2315
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2316
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2317
lemma affine_dependent_translation:
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2318
  assumes "affine_dependent S"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2319
  shows "affine_dependent ((\<lambda>x. a + x) ` S)"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2320
proof -
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2321
  obtain x where x: "x \<in> S \<and> x \<in> affine hull (S - {x})"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2322
    using assms affine_dependent_def by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2323
  have "op + a ` (S - {x}) = op + a ` S - {a + x}"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2324
    by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2325
  then have "a + x \<in> affine hull ((\<lambda>x. a + x) ` S - {a + x})"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2326
    using affine_hull_translation[of a "S - {x}"] x by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2327
  moreover have "a + x \<in> (\<lambda>x. a + x) ` S"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2328
    using x by auto
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2329
  ultimately show ?thesis
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2330
    unfolding affine_dependent_def by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2331
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2332
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2333
lemma affine_dependent_translation_eq:
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2334
  "affine_dependent S \<longleftrightarrow> affine_dependent ((\<lambda>x. a + x) ` S)"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2335
proof -
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2336
  {
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2337
    assume "affine_dependent ((\<lambda>x. a + x) ` S)"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2338
    then have "affine_dependent S"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2339
      using affine_dependent_translation[of "((\<lambda>x. a + x) ` S)" "-a"] translation_assoc[of "-a" a]
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2340
      by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2341
  }
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2342
  then show ?thesis
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2343
    using affine_dependent_translation by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2344
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2345
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2346
lemma affine_hull_0_dependent:
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2347
  assumes "0 \<in> affine hull S"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2348
  shows "dependent S"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2349
proof -
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2350
  obtain s u where s_u: "finite s \<and> s \<noteq> {} \<and> s \<subseteq> S \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = 0"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2351
    using assms affine_hull_explicit[of S] by auto
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2352
  then have "\<exists>v\<in>s. u v \<noteq> 0"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2353
    using setsum_not_0[of "u" "s"] by auto
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2354
  then have "finite s \<and> s \<subseteq> S \<and> (\<exists>v\<in>s. u v \<noteq> 0 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = 0)"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2355
    using s_u by auto
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2356
  then show ?thesis
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2357
    unfolding dependent_explicit[of S] by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2358
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2359
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2360
lemma affine_dependent_imp_dependent2:
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2361
  assumes "affine_dependent (insert 0 S)"
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2362
  shows "dependent S"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2363
proof -
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2364
  obtain x where x: "x \<in> insert 0 S \<and> x \<in> affine hull (insert 0 S - {x})"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2365
    using affine_dependent_def[of "(insert 0 S)"] assms by blast
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2366
  then have "x \<in> span (insert 0 S - {x})"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2367
    using affine_hull_subset_span by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2368
  moreover have "span (insert 0 S - {x}) = span (S - {x})"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2369
    using insert_Diff_if[of "0" S "{x}"] span_insert_0[of "S-{x}"] by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2370
  ultimately have "x \<in> span (S - {x})" by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2371
  then have "x \<noteq> 0 \<Longrightarrow> dependent S"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2372
    using x dependent_def by auto
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2373
  moreover
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2374
  {
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2375
    assume "x = 0"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2376
    then have "0 \<in> affine hull S"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2377
      using x hull_mono[of "S - {0}" S] by auto
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2378
    then have "dependent S"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2379
      using affine_hull_0_dependent by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2380
  }
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2381
  ultimately show ?thesis by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2382
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2383
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2384
lemma affine_dependent_iff_dependent:
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2385
  assumes "a \<notin> S"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2386
  shows "affine_dependent (insert a S) \<longleftrightarrow> dependent ((\<lambda>x. -a + x) ` S)"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2387
proof -
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2388
  have "(op + (- a) ` S) = {x - a| x . x : S}" by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2389
  then show ?thesis
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2390
    using affine_dependent_translation_eq[of "(insert a S)" "-a"]
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  2391
      affine_dependent_imp_dependent2 assms
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2392
      dependent_imp_affine_dependent[of a S]
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53676
diff changeset
  2393
    by (auto simp del: uminus_add_conv_diff)
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2394
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2395
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2396
lemma affine_dependent_iff_dependent2:
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2397
  assumes "a \<in> S"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2398
  shows "affine_dependent S \<longleftrightarrow> dependent ((\<lambda>x. -a + x) ` (S-{a}))"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2399
proof -
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2400
  have "insert a (S - {a}) = S"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2401
    using assms by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2402
  then show ?thesis
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2403
    using assms affine_dependent_iff_dependent[of a "S-{a}"] by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2404
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2405
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2406
lemma affine_hull_insert_span_gen:
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2407
  "affine hull (insert a s) = (\<lambda>x. a + x) ` span ((\<lambda>x. - a + x) ` s)"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2408
proof -
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2409
  have h1: "{x - a |x. x \<in> s} = ((\<lambda>x. -a+x) ` s)"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2410
    by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2411
  {
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2412
    assume "a \<notin> s"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2413
    then have ?thesis
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2414
      using affine_hull_insert_span[of a s] h1 by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2415
  }
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2416
  moreover
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2417
  {
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2418
    assume a1: "a \<in> s"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2419
    have "\<exists>x. x \<in> s \<and> -a+x=0"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2420
      apply (rule exI[of _ a])
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2421
      using a1
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2422
      apply auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2423
      done
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2424
    then have "insert 0 ((\<lambda>x. -a+x) ` (s - {a})) = (\<lambda>x. -a+x) ` s"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2425
      by auto
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2426
    then have "span ((\<lambda>x. -a+x) ` (s - {a}))=span ((\<lambda>x. -a+x) ` s)"
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53676
diff changeset
  2427
      using span_insert_0[of "op + (- a) ` (s - {a})"] by (auto simp del: uminus_add_conv_diff)
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2428
    moreover have "{x - a |x. x \<in> (s - {a})} = ((\<lambda>x. -a+x) ` (s - {a}))"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2429
      by auto
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2430
    moreover have "insert a (s - {a}) = insert a s"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2431
      using assms by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2432
    ultimately have ?thesis
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2433
      using assms affine_hull_insert_span[of "a" "s-{a}"] by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2434
  }
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2435
  ultimately show ?thesis by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2436
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2437
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2438
lemma affine_hull_span2:
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2439
  assumes "a \<in> s"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2440
  shows "affine hull s = (\<lambda>x. a+x) ` span ((\<lambda>x. -a+x) ` (s-{a}))"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2441
  using affine_hull_insert_span_gen[of a "s - {a}", unfolded insert_Diff[OF assms]]
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2442
  by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2443
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2444
lemma affine_hull_span_gen:
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2445
  assumes "a \<in> affine hull s"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2446
  shows "affine hull s = (\<lambda>x. a+x) ` span ((\<lambda>x. -a+x) ` s)"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2447
proof -
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2448
  have "affine hull (insert a s) = affine hull s"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2449
    using hull_redundant[of a affine s] assms by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2450
  then show ?thesis
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2451
    using affine_hull_insert_span_gen[of a "s"] by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2452
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2453
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2454
lemma affine_hull_span_0:
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2455
  assumes "0 \<in> affine hull S"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2456
  shows "affine hull S = span S"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2457
  using affine_hull_span_gen[of "0" S] assms by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2458
63016
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  2459
lemma extend_to_affine_basis_nonempty:
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2460
  fixes S V :: "'n::euclidean_space set"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2461
  assumes "\<not> affine_dependent S" "S \<subseteq> V" "S \<noteq> {}"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2462
  shows "\<exists>T. \<not> affine_dependent T \<and> S \<subseteq> T \<and> T \<subseteq> V \<and> affine hull T = affine hull V"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2463
proof -
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2464
  obtain a where a: "a \<in> S"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2465
    using assms by auto
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2466
  then have h0: "independent  ((\<lambda>x. -a + x) ` (S-{a}))"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2467
    using affine_dependent_iff_dependent2 assms by auto
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2468
  then obtain B where B:
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2469
    "(\<lambda>x. -a+x) ` (S - {a}) \<subseteq> B \<and> B \<subseteq> (\<lambda>x. -a+x) ` V \<and> independent B \<and> (\<lambda>x. -a+x) ` V \<subseteq> span B"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2470
     using maximal_independent_subset_extend[of "(\<lambda>x. -a+x) ` (S-{a})" "(\<lambda>x. -a + x) ` V"] assms
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2471
     by blast
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  2472
  define T where "T = (\<lambda>x. a+x) ` insert 0 B"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2473
  then have "T = insert a ((\<lambda>x. a+x) ` B)"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2474
    by auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2475
  then have "affine hull T = (\<lambda>x. a+x) ` span B"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2476
    using affine_hull_insert_span_gen[of a "((\<lambda>x. a+x) ` B)"] translation_assoc[of "-a" a B]
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2477
    by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2478
  then have "V \<subseteq> affine hull T"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2479
    using B assms translation_inverse_subset[of a V "span B"]
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2480
    by auto
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2481
  moreover have "T \<subseteq> V"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2482
    using T_def B a assms by auto
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2483
  ultimately have "affine hull T = affine hull V"
44457
d366fa5551ef declare euclidean_simps [simp] at the point they are proved;
huffman
parents: 44365
diff changeset
  2484
    by (metis Int_absorb1 Int_absorb2 hull_hull hull_mono)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2485
  moreover have "S \<subseteq> T"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2486
    using T_def B translation_inverse_subset[of a "S-{a}" B]
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2487
    by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2488
  moreover have "\<not> affine_dependent T"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2489
    using T_def affine_dependent_translation_eq[of "insert 0 B"]
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2490
      affine_dependent_imp_dependent2 B
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2491
    by auto
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2492
  ultimately show ?thesis using \<open>T \<subseteq> V\<close> by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2493
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2494
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  2495
lemma affine_basis_exists:
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2496
  fixes V :: "'n::euclidean_space set"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2497
  shows "\<exists>B. B \<subseteq> V \<and> \<not> affine_dependent B \<and> affine hull V = affine hull B"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2498
proof (cases "V = {}")
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2499
  case True
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2500
  then show ?thesis
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2501
    using affine_independent_0 by auto
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2502
next
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2503
  case False
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2504
  then obtain x where "x \<in> V" by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2505
  then show ?thesis
63016
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  2506
    using affine_dependent_def[of "{x}"] extend_to_affine_basis_nonempty[of "{x}" V]
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  2507
    by auto
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  2508
qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  2509
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  2510
proposition extend_to_affine_basis:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  2511
  fixes S V :: "'n::euclidean_space set"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  2512
  assumes "\<not> affine_dependent S" "S \<subseteq> V"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  2513
  obtains T where "\<not> affine_dependent T" "S \<subseteq> T" "T \<subseteq> V" "affine hull T = affine hull V"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  2514
proof (cases "S = {}")
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  2515
  case True then show ?thesis
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  2516
    using affine_basis_exists by (metis empty_subsetI that)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  2517
next
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  2518
  case False
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  2519
  then show ?thesis by (metis assms extend_to_affine_basis_nonempty that)
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2520
qed
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2521
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2522
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2523
subsection \<open>Affine Dimension of a Set\<close>
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2524
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  2525
definition aff_dim :: "('a::euclidean_space) set \<Rightarrow> int"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  2526
  where "aff_dim V =
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2527
  (SOME d :: int.
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2528
    \<exists>B. affine hull B = affine hull V \<and> \<not> affine_dependent B \<and> of_nat (card B) = d + 1)"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2529
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2530
lemma aff_dim_basis_exists:
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  2531
  fixes V :: "('n::euclidean_space) set"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2532
  shows "\<exists>B. affine hull B = affine hull V \<and> \<not> affine_dependent B \<and> of_nat (card B) = aff_dim V + 1"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2533
proof -
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2534
  obtain B where "\<not> affine_dependent B \<and> affine hull B = affine hull V"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2535
    using affine_basis_exists[of V] by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2536
  then show ?thesis
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2537
    unfolding aff_dim_def
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2538
      some_eq_ex[of "\<lambda>d. \<exists>B. affine hull B = affine hull V \<and> \<not> affine_dependent B \<and> of_nat (card B) = d + 1"]
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2539
    apply auto
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2540
    apply (rule exI[of _ "int (card B) - (1 :: int)"])
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2541
    apply (rule exI[of _ "B"])
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2542
    apply auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2543
    done
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2544
qed
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2545
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2546
lemma affine_hull_nonempty: "S \<noteq> {} \<longleftrightarrow> affine hull S \<noteq> {}"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2547
proof -
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2548
  have "S = {} \<Longrightarrow> affine hull S = {}"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2549
    using affine_hull_empty by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2550
  moreover have "affine hull S = {} \<Longrightarrow> S = {}"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2551
    unfolding hull_def by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2552
  ultimately show ?thesis by blast
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2553
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2554
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2555
lemma aff_dim_parallel_subspace_aux:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2556
  fixes B :: "'n::euclidean_space set"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2557
  assumes "\<not> affine_dependent B" "a \<in> B"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2558
  shows "finite B \<and> ((card B) - 1 = dim (span ((\<lambda>x. -a+x) ` (B-{a}))))"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2559
proof -
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2560
  have "independent ((\<lambda>x. -a + x) ` (B-{a}))"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2561
    using affine_dependent_iff_dependent2 assms by auto
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2562
  then have fin: "dim (span ((\<lambda>x. -a+x) ` (B-{a}))) = card ((\<lambda>x. -a + x) ` (B-{a}))"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2563
    "finite ((\<lambda>x. -a + x) ` (B - {a}))"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2564
    using indep_card_eq_dim_span[of "(\<lambda>x. -a+x) ` (B-{a})"] by auto
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2565
  show ?thesis
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2566
  proof (cases "(\<lambda>x. -a + x) ` (B - {a}) = {}")
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2567
    case True
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2568
    have "B = insert a ((\<lambda>x. a + x) ` (\<lambda>x. -a + x) ` (B - {a}))"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2569
      using translation_assoc[of "a" "-a" "(B - {a})"] assms by auto
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2570
    then have "B = {a}" using True by auto
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2571
    then show ?thesis using assms fin by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2572
  next
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2573
    case False
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2574
    then have "card ((\<lambda>x. -a + x) ` (B - {a})) > 0"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2575
      using fin by auto
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2576
    moreover have h1: "card ((\<lambda>x. -a + x) ` (B-{a})) = card (B-{a})"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2577
       apply (rule card_image)
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2578
       using translate_inj_on
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53676
diff changeset
  2579
       apply (auto simp del: uminus_add_conv_diff)
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2580
       done
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2581
    ultimately have "card (B-{a}) > 0" by auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2582
    then have *: "finite (B - {a})"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2583
      using card_gt_0_iff[of "(B - {a})"] by auto
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2584
    then have "card (B - {a}) = card B - 1"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2585
      using card_Diff_singleton assms by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2586
    with * show ?thesis using fin h1 by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2587
  qed
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2588
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2589
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2590
lemma aff_dim_parallel_subspace:
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2591
  fixes V L :: "'n::euclidean_space set"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2592
  assumes "V \<noteq> {}"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2593
    and "subspace L"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2594
    and "affine_parallel (affine hull V) L"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2595
  shows "aff_dim V = int (dim L)"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2596
proof -
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2597
  obtain B where
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2598
    B: "affine hull B = affine hull V \<and> \<not> affine_dependent B \<and> int (card B) = aff_dim V + 1"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2599
    using aff_dim_basis_exists by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2600
  then have "B \<noteq> {}"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2601
    using assms B affine_hull_nonempty[of V] affine_hull_nonempty[of B]
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2602
    by auto
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2603
  then obtain a where a: "a \<in> B" by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  2604
  define Lb where "Lb = span ((\<lambda>x. -a+x) ` (B-{a}))"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2605
  moreover have "affine_parallel (affine hull B) Lb"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2606
    using Lb_def B assms affine_hull_span2[of a B] a
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2607
      affine_parallel_commut[of "Lb" "(affine hull B)"]
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2608
    unfolding affine_parallel_def
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2609
    by auto
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2610
  moreover have "subspace Lb"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2611
    using Lb_def subspace_span by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2612
  moreover have "affine hull B \<noteq> {}"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2613
    using assms B affine_hull_nonempty[of V] by auto
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2614
  ultimately have "L = Lb"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2615
    using assms affine_parallel_subspace[of "affine hull B"] affine_affine_hull[of B] B
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2616
    by auto
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2617
  then have "dim L = dim Lb"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2618
    by auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2619
  moreover have "card B - 1 = dim Lb" and "finite B"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2620
    using Lb_def aff_dim_parallel_subspace_aux a B by auto
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2621
  ultimately show ?thesis
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2622
    using B \<open>B \<noteq> {}\<close> card_gt_0_iff[of B] by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2623
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2624
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2625
lemma aff_independent_finite:
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2626
  fixes B :: "'n::euclidean_space set"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2627
  assumes "\<not> affine_dependent B"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2628
  shows "finite B"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2629
proof -
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2630
  {
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2631
    assume "B \<noteq> {}"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2632
    then obtain a where "a \<in> B" by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2633
    then have ?thesis
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2634
      using aff_dim_parallel_subspace_aux assms by auto
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2635
  }
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2636
  then show ?thesis by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2637
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2638
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2639
lemma independent_finite:
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2640
  fixes B :: "'n::euclidean_space set"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2641
  assumes "independent B"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2642
  shows "finite B"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2643
  using affine_dependent_imp_dependent[of B] aff_independent_finite[of B] assms
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2644
  by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2645
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2646
lemma subspace_dim_equal:
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2647
  assumes "subspace (S :: ('n::euclidean_space) set)"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2648
    and "subspace T"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2649
    and "S \<subseteq> T"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2650
    and "dim S \<ge> dim T"
53302
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2651
  shows "S = T"
98fdf6c34142 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2652
proof -
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2653
  obtain B where B: "B \<le> S" "independent B \<and> S \<subseteq> span B" "card B = dim S"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2654
    using basis_exists[of S] by auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2655
  then have "span B \<subseteq> S"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2656
    using span_mono[of B S] span_eq[of S] assms by metis
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2657
  then have "span B = S"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2658
    using B by auto
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2659
  have "dim S = dim T"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2660
    using assms dim_subset[of S T] by auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2661
  then have "T \<subseteq> span B"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2662
    using card_eq_dim[of B T] B independent_finite assms by auto
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2663
  then show ?thesis
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2664
    using assms \<open>span B = S\<close> by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2665
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2666
63016
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  2667
corollary dim_eq_span:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  2668
  fixes S :: "'a::euclidean_space set"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  2669
  shows "\<lbrakk>S \<subseteq> T; dim T \<le> dim S\<rbrakk> \<Longrightarrow> span S = span T"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  2670
by (simp add: span_mono subspace_dim_equal subspace_span)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  2671
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
  2672
lemma span_substd_basis:
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50104
diff changeset
  2673
  assumes d: "d \<subseteq> Basis"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2674
  shows "span d = {x. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2675
  (is "_ = ?B")
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2676
proof -
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2677
  have "d \<subseteq> ?B"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2678
    using d by (auto simp: inner_Basis)
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2679
  moreover have s: "subspace ?B"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2680
    using subspace_substandard[of "\<lambda>i. i \<notin> d"] .
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2681
  ultimately have "span d \<subseteq> ?B"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2682
    using span_mono[of d "?B"] span_eq[of "?B"] by blast
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 53348
diff changeset
  2683
  moreover have *: "card d \<le> dim (span d)"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2684
    using independent_card_le_dim[of d "span d"] independent_substdbasis[OF assms] span_inc[of d]
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2685
    by auto
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 53348
diff changeset
  2686
  moreover from * have "dim ?B \<le> dim (span d)"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2687
    using dim_substandard[OF assms] by auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2688
  ultimately show ?thesis
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2689
    using s subspace_dim_equal[of "span d" "?B"] subspace_span[of d] by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2690
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2691
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2692
lemma basis_to_substdbasis_subspace_isomorphism:
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2693
  fixes B :: "'a::euclidean_space set"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2694
  assumes "independent B"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2695
  shows "\<exists>f d::'a set. card d = card B \<and> linear f \<and> f ` B = d \<and>
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2696
    f ` span B = {x. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0} \<and> inj_on f (span B) \<and> d \<subseteq> Basis"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2697
proof -
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2698
  have B: "card B = dim B"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2699
    using dim_unique[of B B "card B"] assms span_inc[of B] by auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2700
  have "dim B \<le> card (Basis :: 'a set)"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2701
    using dim_subset_UNIV[of B] by simp
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2702
  from ex_card[OF this] obtain d :: "'a set" where d: "d \<subseteq> Basis" and t: "card d = dim B"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2703
    by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2704
  let ?t = "{x::'a::euclidean_space. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0}"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2705
  have "\<exists>f. linear f \<and> f ` B = d \<and> f ` span B = ?t \<and> inj_on f (span B)"
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
  2706
    apply (rule basis_to_basis_subspace_isomorphism[of "span B" ?t B "d"])
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2707
    apply (rule subspace_span)
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2708
    apply (rule subspace_substandard)
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2709
    defer
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2710
    apply (rule span_inc)
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2711
    apply (rule assms)
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2712
    defer
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2713
    unfolding dim_span[of B]
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2714
    apply(rule B)
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2715
    unfolding span_substd_basis[OF d, symmetric]
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2716
    apply (rule span_inc)
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2717
    apply (rule independent_substdbasis[OF d])
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2718
    apply rule
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2719
    apply assumption
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2720
    unfolding t[symmetric] span_substd_basis[OF d] dim_substandard[OF d]
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2721
    apply auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2722
    done
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2723
  with t \<open>card B = dim B\<close> d show ?thesis by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2724
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2725
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2726
lemma aff_dim_empty:
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2727
  fixes S :: "'n::euclidean_space set"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2728
  shows "S = {} \<longleftrightarrow> aff_dim S = -1"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2729
proof -
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2730
  obtain B where *: "affine hull B = affine hull S"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2731
    and "\<not> affine_dependent B"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2732
    and "int (card B) = aff_dim S + 1"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2733
    using aff_dim_basis_exists by auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2734
  moreover
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2735
  from * have "S = {} \<longleftrightarrow> B = {}"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2736
    using affine_hull_nonempty[of B] affine_hull_nonempty[of S] by auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2737
  ultimately show ?thesis
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2738
    using aff_independent_finite[of B] card_gt_0_iff[of B] by auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2739
qed
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2740
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  2741
lemma aff_dim_empty_eq [simp]: "aff_dim ({}::'a::euclidean_space set) = -1"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  2742
  by (simp add: aff_dim_empty [symmetric])
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  2743
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2744
lemma aff_dim_affine_hull: "aff_dim (affine hull S) = aff_dim S"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2745
  unfolding aff_dim_def using hull_hull[of _ S] by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2746
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2747
lemma aff_dim_affine_hull2:
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2748
  assumes "affine hull S = affine hull T"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2749
  shows "aff_dim S = aff_dim T"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2750
  unfolding aff_dim_def using assms by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2751
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  2752
lemma aff_dim_unique:
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2753
  fixes B V :: "'n::euclidean_space set"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2754
  assumes "affine hull B = affine hull V \<and> \<not> affine_dependent B"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2755
  shows "of_nat (card B) = aff_dim V + 1"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2756
proof (cases "B = {}")
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2757
  case True
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2758
  then have "V = {}"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2759
    using affine_hull_nonempty[of V] affine_hull_nonempty[of B] assms
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2760
    by auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2761
  then have "aff_dim V = (-1::int)"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2762
    using aff_dim_empty by auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2763
  then show ?thesis
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2764
    using \<open>B = {}\<close> by auto
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2765
next
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2766
  case False
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2767
  then obtain a where a: "a \<in> B" by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  2768
  define Lb where "Lb = span ((\<lambda>x. -a+x) ` (B-{a}))"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2769
  have "affine_parallel (affine hull B) Lb"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2770
    using Lb_def affine_hull_span2[of a B] a
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2771
      affine_parallel_commut[of "Lb" "(affine hull B)"]
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2772
    unfolding affine_parallel_def by auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2773
  moreover have "subspace Lb"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2774
    using Lb_def subspace_span by auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2775
  ultimately have "aff_dim B = int(dim Lb)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2776
    using aff_dim_parallel_subspace[of B Lb] \<open>B \<noteq> {}\<close> by auto
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2777
  moreover have "(card B) - 1 = dim Lb" "finite B"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  2778
    using Lb_def aff_dim_parallel_subspace_aux a assms by auto
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2779
  ultimately have "of_nat (card B) = aff_dim B + 1"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2780
    using \<open>B \<noteq> {}\<close> card_gt_0_iff[of B] by auto
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2781
  then show ?thesis
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2782
    using aff_dim_affine_hull2 assms by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2783
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2784
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  2785
lemma aff_dim_affine_independent:
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2786
  fixes B :: "'n::euclidean_space set"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2787
  assumes "\<not> affine_dependent B"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2788
  shows "of_nat (card B) = aff_dim B + 1"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2789
  using aff_dim_unique[of B B] assms by auto
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2790
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  2791
lemma affine_independent_iff_card:
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  2792
    fixes s :: "'a::euclidean_space set"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  2793
    shows "~ affine_dependent s \<longleftrightarrow> finite s \<and> aff_dim s = int(card s) - 1"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  2794
  apply (rule iffI)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  2795
  apply (simp add: aff_dim_affine_independent aff_independent_finite)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  2796
  by (metis affine_basis_exists [of s] aff_dim_unique card_subset_eq diff_add_cancel of_nat_eq_iff)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  2797
62948
7700f467892b lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 62843
diff changeset
  2798
lemma aff_dim_sing [simp]:
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2799
  fixes a :: "'n::euclidean_space"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2800
  shows "aff_dim {a} = 0"
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2801
  using aff_dim_affine_independent[of "{a}"] affine_independent_1 by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2802
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2803
lemma aff_dim_inner_basis_exists:
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  2804
  fixes V :: "('n::euclidean_space) set"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2805
  shows "\<exists>B. B \<subseteq> V \<and> affine hull B = affine hull V \<and>
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2806
    \<not> affine_dependent B \<and> of_nat (card B) = aff_dim V + 1"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2807
proof -
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2808
  obtain B where B: "\<not> affine_dependent B" "B \<subseteq> V" "affine hull B = affine hull V"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2809
    using affine_basis_exists[of V] by auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2810
  then have "of_nat(card B) = aff_dim V+1" using aff_dim_unique by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2811
  with B show ?thesis by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2812
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2813
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2814
lemma aff_dim_le_card:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2815
  fixes V :: "'n::euclidean_space set"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2816
  assumes "finite V"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2817
  shows "aff_dim V \<le> of_nat (card V) - 1"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2818
proof -
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2819
  obtain B where B: "B \<subseteq> V" "of_nat (card B) = aff_dim V + 1"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2820
    using aff_dim_inner_basis_exists[of V] by auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2821
  then have "card B \<le> card V"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2822
    using assms card_mono by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2823
  with B show ?thesis by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2824
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2825
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2826
lemma aff_dim_parallel_eq:
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2827
  fixes S T :: "'n::euclidean_space set"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2828
  assumes "affine_parallel (affine hull S) (affine hull T)"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2829
  shows "aff_dim S = aff_dim T"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2830
proof -
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2831
  {
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2832
    assume "T \<noteq> {}" "S \<noteq> {}"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2833
    then obtain L where L: "subspace L \<and> affine_parallel (affine hull T) L"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2834
      using affine_parallel_subspace[of "affine hull T"]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2835
        affine_affine_hull[of T] affine_hull_nonempty
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2836
      by auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2837
    then have "aff_dim T = int (dim L)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2838
      using aff_dim_parallel_subspace \<open>T \<noteq> {}\<close> by auto
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2839
    moreover have *: "subspace L \<and> affine_parallel (affine hull S) L"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2840
       using L affine_parallel_assoc[of "affine hull S" "affine hull T" L] assms by auto
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2841
    moreover from * have "aff_dim S = int (dim L)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  2842
      using aff_dim_parallel_subspace \<open>S \<noteq> {}\<close> by auto
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2843
    ultimately have ?thesis by auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2844
  }
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2845
  moreover
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2846
  {
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2847
    assume "S = {}"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2848
    then have "S = {}" and "T = {}"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2849
      using assms affine_hull_nonempty
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2850
      unfolding affine_parallel_def
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2851
      by auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2852
    then have ?thesis using aff_dim_empty by auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2853
  }
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2854
  moreover
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2855
  {
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2856
    assume "T = {}"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2857
    then have "S = {}" and "T = {}"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2858
      using assms affine_hull_nonempty
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2859
      unfolding affine_parallel_def
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2860
      by auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2861
    then have ?thesis
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2862
      using aff_dim_empty by auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2863
  }
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2864
  ultimately show ?thesis by blast
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2865
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2866
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2867
lemma aff_dim_translation_eq:
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2868
  fixes a :: "'n::euclidean_space"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2869
  shows "aff_dim ((\<lambda>x. a + x) ` S) = aff_dim S"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2870
proof -
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2871
  have "affine_parallel (affine hull S) (affine hull ((\<lambda>x. a + x) ` S))"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2872
    unfolding affine_parallel_def
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2873
    apply (rule exI[of _ "a"])
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2874
    using affine_hull_translation[of a S]
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2875
    apply auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2876
    done
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2877
  then show ?thesis
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2878
    using aff_dim_parallel_eq[of S "(\<lambda>x. a + x) ` S"] by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2879
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2880
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2881
lemma aff_dim_affine:
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2882
  fixes S L :: "'n::euclidean_space set"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2883
  assumes "S \<noteq> {}"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2884
    and "affine S"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2885
    and "subspace L"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2886
    and "affine_parallel S L"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2887
  shows "aff_dim S = int (dim L)"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2888
proof -
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2889
  have *: "affine hull S = S"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2890
    using assms affine_hull_eq[of S] by auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2891
  then have "affine_parallel (affine hull S) L"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2892
    using assms by (simp add: *)
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2893
  then show ?thesis
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2894
    using assms aff_dim_parallel_subspace[of S L] by blast
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2895
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2896
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2897
lemma dim_affine_hull:
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2898
  fixes S :: "'n::euclidean_space set"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2899
  shows "dim (affine hull S) = dim S"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2900
proof -
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2901
  have "dim (affine hull S) \<ge> dim S"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2902
    using dim_subset by auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2903
  moreover have "dim (span S) \<ge> dim (affine hull S)"
60303
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60176
diff changeset
  2904
    using dim_subset affine_hull_subset_span by blast
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2905
  moreover have "dim (span S) = dim S"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2906
    using dim_span by auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2907
  ultimately show ?thesis by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2908
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2909
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2910
lemma aff_dim_subspace:
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2911
  fixes S :: "'n::euclidean_space set"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2912
  assumes "S \<noteq> {}"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2913
    and "subspace S"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2914
  shows "aff_dim S = int (dim S)"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2915
  using aff_dim_affine[of S S] assms subspace_imp_affine[of S] affine_parallel_reflex[of S]
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2916
  by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2917
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2918
lemma aff_dim_zero:
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2919
  fixes S :: "'n::euclidean_space set"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2920
  assumes "0 \<in> affine hull S"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2921
  shows "aff_dim S = int (dim S)"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2922
proof -
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2923
  have "subspace (affine hull S)"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2924
    using subspace_affine[of "affine hull S"] affine_affine_hull assms
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2925
    by auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2926
  then have "aff_dim (affine hull S) = int (dim (affine hull S))"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2927
    using assms aff_dim_subspace[of "affine hull S"] by auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2928
  then show ?thesis
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2929
    using aff_dim_affine_hull[of S] dim_affine_hull[of S]
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2930
    by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2931
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2932
63016
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  2933
lemma aff_dim_eq_dim:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  2934
  fixes S :: "'n::euclidean_space set"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  2935
  assumes "a \<in> affine hull S"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  2936
  shows "aff_dim S = int (dim ((\<lambda>x. -a+x) ` S))"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  2937
proof -
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  2938
  have "0 \<in> affine hull ((\<lambda>x. -a+x) ` S)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  2939
    unfolding Convex_Euclidean_Space.affine_hull_translation
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  2940
    using assms by (simp add: ab_group_add_class.ab_left_minus image_iff)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  2941
  with aff_dim_zero show ?thesis
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  2942
    by (metis aff_dim_translation_eq)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  2943
qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  2944
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2945
lemma aff_dim_univ: "aff_dim (UNIV :: 'n::euclidean_space set) = int(DIM('n))"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2946
  using aff_dim_subspace[of "(UNIV :: 'n::euclidean_space set)"]
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2947
    dim_UNIV[where 'a="'n::euclidean_space"]
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2948
  by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2949
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2950
lemma aff_dim_geq:
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2951
  fixes V :: "'n::euclidean_space set"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2952
  shows "aff_dim V \<ge> -1"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2953
proof -
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2954
  obtain B where "affine hull B = affine hull V"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2955
    and "\<not> affine_dependent B"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2956
    and "int (card B) = aff_dim V + 1"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2957
    using aff_dim_basis_exists by auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2958
  then show ?thesis by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2959
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2960
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  2961
lemma independent_card_le_aff_dim:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2962
  fixes B :: "'n::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2963
  assumes "B \<subseteq> V"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2964
  assumes "\<not> affine_dependent B"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2965
  shows "int (card B) \<le> aff_dim V + 1"
63016
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  2966
proof -
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  2967
  obtain T where T: "\<not> affine_dependent T \<and> B \<subseteq> T \<and> T \<subseteq> V \<and> affine hull T = affine hull V"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  2968
    by (metis assms extend_to_affine_basis[of B V])
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2969
  then have "of_nat (card T) = aff_dim V + 1"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2970
    using aff_dim_unique by auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2971
  then show ?thesis
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2972
    using T card_mono[of T B] aff_independent_finite[of T] by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2973
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2974
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2975
lemma aff_dim_subset:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2976
  fixes S T :: "'n::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2977
  assumes "S \<subseteq> T"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2978
  shows "aff_dim S \<le> aff_dim T"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2979
proof -
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2980
  obtain B where B: "\<not> affine_dependent B" "B \<subseteq> S" "affine hull B = affine hull S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2981
    "of_nat (card B) = aff_dim S + 1"
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2982
    using aff_dim_inner_basis_exists[of S] by auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2983
  then have "int (card B) \<le> aff_dim T + 1"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2984
    using assms independent_card_le_aff_dim[of B T] by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2985
  with B show ?thesis by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2986
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2987
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  2988
lemma aff_dim_le_DIM:
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2989
  fixes S :: "'n::euclidean_space set"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2990
  shows "aff_dim S \<le> int (DIM('n))"
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  2991
proof -
53339
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2992
  have "aff_dim (UNIV :: 'n::euclidean_space set) = int(DIM('n))"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2993
    using aff_dim_univ by auto
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2994
  then show "aff_dim (S:: 'n::euclidean_space set) \<le> int(DIM('n))"
0dc28fd72c7d tuned proofs;
wenzelm
parents: 53333
diff changeset
  2995
    using assms aff_dim_subset[of S "(UNIV :: ('n::euclidean_space) set)"] subset_UNIV by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2996
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2997
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  2998
lemma affine_dim_equal:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  2999
  fixes S :: "'n::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3000
  assumes "affine S" "affine T" "S \<noteq> {}" "S \<subseteq> T" "aff_dim S = aff_dim T"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3001
  shows "S = T"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3002
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3003
  obtain a where "a \<in> S" using assms by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3004
  then have "a \<in> T" using assms by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  3005
  define LS where "LS = {y. \<exists>x \<in> S. (-a) + x = y}"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3006
  then have ls: "subspace LS" "affine_parallel S LS"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  3007
    using assms parallel_subspace_explicit[of S a LS] \<open>a \<in> S\<close> by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3008
  then have h1: "int(dim LS) = aff_dim S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3009
    using assms aff_dim_affine[of S LS] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3010
  have "T \<noteq> {}" using assms by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  3011
  define LT where "LT = {y. \<exists>x \<in> T. (-a) + x = y}"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3012
  then have lt: "subspace LT \<and> affine_parallel T LT"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  3013
    using assms parallel_subspace_explicit[of T a LT] \<open>a \<in> T\<close> by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3014
  then have "int(dim LT) = aff_dim T"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  3015
    using assms aff_dim_affine[of T LT] \<open>T \<noteq> {}\<close> by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3016
  then have "dim LS = dim LT"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3017
    using h1 assms by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3018
  moreover have "LS \<le> LT"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3019
    using LS_def LT_def assms by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3020
  ultimately have "LS = LT"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3021
    using subspace_dim_equal[of LS LT] ls lt by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3022
  moreover have "S = {x. \<exists>y \<in> LS. a+y=x}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3023
    using LS_def by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3024
  moreover have "T = {x. \<exists>y \<in> LT. a+y=x}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3025
    using LT_def by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3026
  ultimately show ?thesis by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3027
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3028
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  3029
lemma affine_hull_UNIV:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3030
  fixes S :: "'n::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3031
  assumes "aff_dim S = int(DIM('n))"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3032
  shows "affine hull S = (UNIV :: ('n::euclidean_space) set)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3033
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3034
  have "S \<noteq> {}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3035
    using assms aff_dim_empty[of S] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3036
  have h0: "S \<subseteq> affine hull S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3037
    using hull_subset[of S _] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3038
  have h1: "aff_dim (UNIV :: ('n::euclidean_space) set) = aff_dim S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3039
    using aff_dim_univ assms by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3040
  then have h2: "aff_dim (affine hull S) \<le> aff_dim (UNIV :: ('n::euclidean_space) set)"
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  3041
    using aff_dim_le_DIM[of "affine hull S"] assms h0 by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3042
  have h3: "aff_dim S \<le> aff_dim (affine hull S)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3043
    using h0 aff_dim_subset[of S "affine hull S"] assms by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3044
  then have h4: "aff_dim (affine hull S) = aff_dim (UNIV :: ('n::euclidean_space) set)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3045
    using h0 h1 h2 by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3046
  then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3047
    using affine_dim_equal[of "affine hull S" "(UNIV :: ('n::euclidean_space) set)"]
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  3048
      affine_affine_hull[of S] affine_UNIV assms h4 h0 \<open>S \<noteq> {}\<close>
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3049
    by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3050
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3051
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  3052
lemma disjoint_affine_hull:
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  3053
  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
  3054
  assumes "~ affine_dependent s" "t \<subseteq> s" "u \<subseteq> s" "t \<inter> u = {}"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  3055
    shows "(affine hull t) \<inter> (affine hull u) = {}"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  3056
proof -
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  3057
  have "finite s" using assms by (simp add: aff_independent_finite)
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  3058
  then have "finite t" "finite u" using assms finite_subset by blast+
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  3059
  { fix y
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  3060
    assume yt: "y \<in> affine hull t" and yu: "y \<in> affine hull u"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  3061
    then obtain a b
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  3062
           where a1 [simp]: "setsum a t = 1" and [simp]: "setsum (\<lambda>v. a v *\<^sub>R v) t = y"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  3063
             and [simp]: "setsum b u = 1" "setsum (\<lambda>v. b v *\<^sub>R v) u = y"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  3064
      by (auto simp: affine_hull_finite \<open>finite t\<close> \<open>finite u\<close>)
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  3065
    define c where "c x = (if x \<in> t then a x else if x \<in> u then -(b x) else 0)" for x
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  3066
    have [simp]: "s \<inter> t = t" "s \<inter> - t \<inter> u = u" using assms by auto
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  3067
    have "setsum c s = 0"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  3068
      by (simp add: c_def comm_monoid_add_class.setsum.If_cases \<open>finite s\<close> setsum_negf)
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  3069
    moreover have "~ (\<forall>v\<in>s. c v = 0)"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  3070
      by (metis (no_types) IntD1 \<open>s \<inter> t = t\<close> a1 c_def setsum_not_0 zero_neq_one)
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  3071
    moreover have "(\<Sum>v\<in>s. c v *\<^sub>R v) = 0"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  3072
      by (simp add: c_def if_smult setsum_negf
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  3073
             comm_monoid_add_class.setsum.If_cases \<open>finite s\<close>)
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  3074
    ultimately have False
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  3075
      using assms \<open>finite s\<close> by (auto simp: affine_dependent_explicit)
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  3076
  }
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  3077
  then show ?thesis by blast
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  3078
qed
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  3079
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3080
lemma aff_dim_convex_hull:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3081
  fixes S :: "'n::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3082
  shows "aff_dim (convex hull S) = aff_dim S"
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  3083
  using aff_dim_affine_hull[of S] convex_hull_subset_affine_hull[of S]
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3084
    hull_subset[of S "convex"] aff_dim_subset[of S "convex hull S"]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3085
    aff_dim_subset[of "convex hull S" "affine hull S"]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3086
  by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3087
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3088
lemma aff_dim_cball:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3089
  fixes a :: "'n::euclidean_space"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3090
  assumes "e > 0"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3091
  shows "aff_dim (cball a e) = int (DIM('n))"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3092
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3093
  have "(\<lambda>x. a + x) ` (cball 0 e) \<subseteq> cball a e"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3094
    unfolding cball_def dist_norm by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3095
  then have "aff_dim (cball (0 :: 'n::euclidean_space) e) \<le> aff_dim (cball a e)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3096
    using aff_dim_translation_eq[of a "cball 0 e"]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3097
          aff_dim_subset[of "op + a ` cball 0 e" "cball a e"]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3098
    by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3099
  moreover have "aff_dim (cball (0 :: 'n::euclidean_space) e) = int (DIM('n))"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3100
    using hull_inc[of "(0 :: 'n::euclidean_space)" "cball 0 e"]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3101
      centre_in_cball[of "(0 :: 'n::euclidean_space)"] assms
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3102
    by (simp add: dim_cball[of e] aff_dim_zero[of "cball 0 e"])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3103
  ultimately show ?thesis
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  3104
    using aff_dim_le_DIM[of "cball a e"] by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3105
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3106
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3107
lemma aff_dim_open:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3108
  fixes S :: "'n::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3109
  assumes "open S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3110
    and "S \<noteq> {}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3111
  shows "aff_dim S = int (DIM('n))"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3112
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3113
  obtain x where "x \<in> S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3114
    using assms by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3115
  then obtain e where e: "e > 0" "cball x e \<subseteq> S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3116
    using open_contains_cball[of S] assms by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3117
  then have "aff_dim (cball x e) \<le> aff_dim S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3118
    using aff_dim_subset by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3119
  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
  3120
    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
  3121
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3122
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3123
lemma low_dim_interior:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3124
  fixes S :: "'n::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3125
  assumes "\<not> aff_dim S = int (DIM('n))"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3126
  shows "interior S = {}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3127
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3128
  have "aff_dim(interior S) \<le> aff_dim S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3129
    using interior_subset aff_dim_subset[of "interior S" S] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3130
  then show ?thesis
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  3131
    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
  3132
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3133
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  3134
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
  3135
  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
  3136
  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
  3137
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
  3138
63016
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  3139
corollary aff_dim_nonempty_interior:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  3140
  fixes S :: "'a::euclidean_space set"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  3141
  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
  3142
by (metis low_dim_interior)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
  3143
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3144
subsection \<open>Caratheodory's theorem.\<close>
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3145
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3146
lemma convex_hull_caratheodory_aff_dim:
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3147
  fixes p :: "('a::euclidean_space) set"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3148
  shows "convex hull p =
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3149
    {y. \<exists>s u. finite s \<and> s \<subseteq> p \<and> card s \<le> aff_dim p + 1 \<and>
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3150
      (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> setsum (\<lambda>v. u v *\<^sub>R v) s = y}"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3151
  unfolding convex_hull_explicit set_eq_iff mem_Collect_eq
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3152
proof (intro allI iffI)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3153
  fix y
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3154
  let ?P = "\<lambda>n. \<exists>s u. finite s \<and> card s = n \<and> s \<subseteq> p \<and> (\<forall>x\<in>s. 0 \<le> u x) \<and>
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3155
    setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3156
  assume "\<exists>s u. finite s \<and> s \<subseteq> p \<and> (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3157
  then obtain N where "?P N" by auto
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3158
  then have "\<exists>n\<le>N. (\<forall>k<n. \<not> ?P k) \<and> ?P n"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3159
    apply (rule_tac ex_least_nat_le)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3160
    apply auto
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3161
    done
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3162
  then obtain n where "?P n" and smallest: "\<forall>k<n. \<not> ?P k"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3163
    by blast
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3164
  then obtain s u where obt: "finite s" "card s = n" "s\<subseteq>p" "\<forall>x\<in>s. 0 \<le> u x"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3165
    "setsum u s = 1"  "(\<Sum>v\<in>s. u v *\<^sub>R v) = y" by auto
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3166
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3167
  have "card s \<le> aff_dim p + 1"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3168
  proof (rule ccontr, simp only: not_le)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3169
    assume "aff_dim p + 1 < card s"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3170
    then have "affine_dependent s"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3171
      using affine_dependent_biggerset[OF obt(1)] independent_card_le_aff_dim not_less obt(3)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3172
      by blast
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3173
    then obtain w v where wv: "setsum w s = 0" "v\<in>s" "w v \<noteq> 0" "(\<Sum>v\<in>s. w v *\<^sub>R v) = 0"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3174
      using affine_dependent_explicit_finite[OF obt(1)] by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  3175
    define i where "i = (\<lambda>v. (u v) / (- w v)) ` {v\<in>s. w v < 0}"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  3176
    define t where "t = Min i"
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3177
    have "\<exists>x\<in>s. w x < 0"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3178
    proof (rule ccontr, simp add: not_less)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3179
      assume as:"\<forall>x\<in>s. 0 \<le> w x"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3180
      then have "setsum w (s - {v}) \<ge> 0"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3181
        apply (rule_tac setsum_nonneg)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3182
        apply auto
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3183
        done
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3184
      then have "setsum w s > 0"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3185
        unfolding setsum.remove[OF obt(1) \<open>v\<in>s\<close>]
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3186
        using as[THEN bspec[where x=v]]  \<open>v\<in>s\<close>  \<open>w v \<noteq> 0\<close> by auto
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3187
      then show False using wv(1) by auto
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3188
    qed
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3189
    then have "i \<noteq> {}" unfolding i_def by auto
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3190
    then have "t \<ge> 0"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3191
      using Min_ge_iff[of i 0 ] and obt(1)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3192
      unfolding t_def i_def
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3193
      using obt(4)[unfolded le_less]
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3194
      by (auto simp: divide_le_0_iff)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3195
    have t: "\<forall>v\<in>s. u v + t * w v \<ge> 0"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3196
    proof
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3197
      fix v
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3198
      assume "v \<in> s"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3199
      then have v: "0 \<le> u v"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3200
        using obt(4)[THEN bspec[where x=v]] by auto
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3201
      show "0 \<le> u v + t * w v"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3202
      proof (cases "w v < 0")
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3203
        case False
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3204
        thus ?thesis using v \<open>t\<ge>0\<close> by auto
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3205
      next
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3206
        case True
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3207
        then have "t \<le> u v / (- w v)"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3208
          using \<open>v\<in>s\<close> unfolding t_def i_def
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3209
          apply (rule_tac Min_le)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3210
          using obt(1) apply auto
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3211
          done
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3212
        then show ?thesis
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3213
          unfolding real_0_le_add_iff
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3214
          using pos_le_divide_eq[OF True[unfolded neg_0_less_iff_less[symmetric]]]
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3215
          by auto
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3216
      qed
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3217
    qed
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3218
    obtain a where "a \<in> s" and "t = (\<lambda>v. (u v) / (- w v)) a" and "w a < 0"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3219
      using Min_in[OF _ \<open>i\<noteq>{}\<close>] and obt(1) unfolding i_def t_def by auto
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3220
    then have a: "a \<in> s" "u a + t * w a = 0" by auto
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3221
    have *: "\<And>f. setsum f (s - {a}) = setsum f s - ((f a)::'b::ab_group_add)"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3222
      unfolding setsum.remove[OF obt(1) \<open>a\<in>s\<close>] by auto
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3223
    have "(\<Sum>v\<in>s. u v + t * w v) = 1"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3224
      unfolding setsum.distrib wv(1) setsum_right_distrib[symmetric] obt(5) by auto
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3225
    moreover have "(\<Sum>v\<in>s. u v *\<^sub>R v + (t * w v) *\<^sub>R v) - (u a *\<^sub>R a + (t * w a) *\<^sub>R a) = y"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3226
      unfolding setsum.distrib obt(6) scaleR_scaleR[symmetric] scaleR_right.setsum [symmetric] wv(4)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3227
      using a(2) [THEN eq_neg_iff_add_eq_0 [THEN iffD2]] by simp
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3228
    ultimately have "?P (n - 1)"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3229
      apply (rule_tac x="(s - {a})" in exI)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3230
      apply (rule_tac x="\<lambda>v. u v + t * w v" in exI)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3231
      using obt(1-3) and t and a
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3232
      apply (auto simp add: * scaleR_left_distrib)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3233
      done
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3234
    then show False
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3235
      using smallest[THEN spec[where x="n - 1"]] by auto
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3236
  qed
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3237
  then show "\<exists>s u. finite s \<and> s \<subseteq> p \<and> card s \<le> aff_dim p + 1 \<and>
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3238
      (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3239
    using obt by auto
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3240
qed auto
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3241
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3242
lemma caratheodory_aff_dim:
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3243
  fixes p :: "('a::euclidean_space) set"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3244
  shows "convex hull p = {x. \<exists>s. finite s \<and> s \<subseteq> p \<and> card s \<le> aff_dim p + 1 \<and> x \<in> convex hull s}"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3245
        (is "?lhs = ?rhs")
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3246
proof
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3247
  show "?lhs \<subseteq> ?rhs"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3248
    apply (subst convex_hull_caratheodory_aff_dim)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3249
    apply clarify
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3250
    apply (rule_tac x="s" in exI)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3251
    apply (simp add: hull_subset convex_explicit [THEN iffD1, OF convex_convex_hull])
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3252
    done
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3253
next
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3254
  show "?rhs \<subseteq> ?lhs"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3255
    using hull_mono by blast
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3256
qed
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3257
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3258
lemma convex_hull_caratheodory:
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3259
  fixes p :: "('a::euclidean_space) set"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3260
  shows "convex hull p =
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3261
            {y. \<exists>s u. finite s \<and> s \<subseteq> p \<and> card s \<le> DIM('a) + 1 \<and>
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3262
              (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> setsum (\<lambda>v. u v *\<^sub>R v) s = y}"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3263
        (is "?lhs = ?rhs")
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3264
proof (intro set_eqI iffI)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3265
  fix x
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3266
  assume "x \<in> ?lhs" then show "x \<in> ?rhs"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3267
    apply (simp only: convex_hull_caratheodory_aff_dim Set.mem_Collect_eq)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3268
    apply (erule ex_forward)+
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  3269
    using aff_dim_le_DIM [of p]
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3270
    apply simp
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3271
    done
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3272
next
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3273
  fix x
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3274
  assume "x \<in> ?rhs" then show "x \<in> ?lhs"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3275
    by (auto simp add: convex_hull_explicit)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3276
qed
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3277
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3278
theorem caratheodory:
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3279
  "convex hull p =
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3280
    {x::'a::euclidean_space. \<exists>s. finite s \<and> s \<subseteq> p \<and>
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3281
      card s \<le> DIM('a) + 1 \<and> x \<in> convex hull s}"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3282
proof safe
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3283
  fix x
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3284
  assume "x \<in> convex hull p"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3285
  then obtain s u where "finite s" "s \<subseteq> p" "card s \<le> DIM('a) + 1"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3286
    "\<forall>x\<in>s. 0 \<le> u x" "setsum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3287
    unfolding convex_hull_caratheodory by auto
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3288
  then show "\<exists>s. finite s \<and> s \<subseteq> p \<and> card s \<le> DIM('a) + 1 \<and> x \<in> convex hull s"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3289
    apply (rule_tac x=s in exI)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3290
    using hull_subset[of s convex]
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3291
    using convex_convex_hull[unfolded convex_explicit, of s,
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3292
      THEN spec[where x=s], THEN spec[where x=u]]
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3293
    apply auto
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3294
    done
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3295
next
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3296
  fix x s
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3297
  assume  "finite s" "s \<subseteq> p" "card s \<le> DIM('a) + 1" "x \<in> convex hull s"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3298
  then show "x \<in> convex hull p"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3299
    using hull_mono[OF \<open>s\<subseteq>p\<close>] by auto
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3300
qed
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3301
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3302
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  3303
subsection \<open>Relative interior of a set\<close>
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3304
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3305
definition "rel_interior S =
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3306
  {x. \<exists>T. openin (subtopology euclidean (affine hull S)) T \<and> x \<in> T \<and> T \<subseteq> S}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3307
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3308
lemma rel_interior:
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3309
  "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
  3310
  unfolding rel_interior_def[of S] openin_open[of "affine hull S"]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3311
  apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3312
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3313
  fix x T
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3314
  assume *: "x \<in> S" "open T" "x \<in> T" "T \<inter> affine hull S \<subseteq> S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3315
  then have **: "x \<in> T \<inter> affine hull S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3316
    using hull_inc by auto
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  3317
  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
  3318
    apply (rule_tac x = "T \<inter> (affine hull S)" in exI)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3319
    using * **
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3320
    apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3321
    done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3322
qed
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3323
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3324
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)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3325
  by (auto simp add: rel_interior)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3326
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3327
lemma mem_rel_interior_ball:
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3328
  "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
  3329
  apply (simp add: rel_interior, safe)
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3330
  apply (force simp add: open_contains_ball)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3331
  apply (rule_tac x = "ball x e" in exI)
44457
d366fa5551ef declare euclidean_simps [simp] at the point they are proved;
huffman
parents: 44365
diff changeset
  3332
  apply simp
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3333
  done
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3334
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  3335
lemma rel_interior_ball:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3336
  "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
  3337
  using mem_rel_interior_ball [of _ S] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3338
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3339
lemma mem_rel_interior_cball:
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3340
  "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
  3341
  apply (simp add: rel_interior, safe)
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3342
  apply (force simp add: open_contains_cball)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3343
  apply (rule_tac x = "ball x e" in exI)
44457
d366fa5551ef declare euclidean_simps [simp] at the point they are proved;
huffman
parents: 44365
diff changeset
  3344
  apply (simp add: subset_trans [OF ball_subset_cball])
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3345
  apply auto
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3346
  done
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3347
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3348
lemma rel_interior_cball:
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3349
  "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
  3350
  using mem_rel_interior_cball [of _ S] by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3351
60303
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60176
diff changeset
  3352
lemma rel_interior_empty [simp]: "rel_interior {} = {}"
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  3353
   by (auto simp add: rel_interior_def)
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3354
60303
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60176
diff changeset
  3355
lemma affine_hull_sing [simp]: "affine hull {a :: 'n::euclidean_space} = {a}"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3356
  by (metis affine_hull_eq affine_sing)
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3357
60303
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60176
diff changeset
  3358
lemma rel_interior_sing [simp]: "rel_interior {a :: 'n::euclidean_space} = {a}"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3359
  unfolding rel_interior_ball affine_hull_sing
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3360
  apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3361
  apply (rule_tac x = "1 :: real" in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3362
  apply simp
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3363
  done
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3364
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3365
lemma subset_rel_interior:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3366
  fixes S T :: "'n::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3367
  assumes "S \<subseteq> T"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3368
    and "affine hull S = affine hull T"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3369
  shows "rel_interior S \<subseteq> rel_interior T"
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  3370
  using assms by (auto simp add: rel_interior_def)
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  3371
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3372
lemma rel_interior_subset: "rel_interior S \<subseteq> S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3373
  by (auto simp add: rel_interior_def)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3374
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3375
lemma rel_interior_subset_closure: "rel_interior S \<subseteq> closure S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3376
  using rel_interior_subset by (auto simp add: closure_def)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3377
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3378
lemma interior_subset_rel_interior: "interior S \<subseteq> rel_interior S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3379
  by (auto simp add: rel_interior interior_def)
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3380
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3381
lemma interior_rel_interior:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3382
  fixes S :: "'n::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3383
  assumes "aff_dim S = int(DIM('n))"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3384
  shows "rel_interior S = interior S"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3385
proof -
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3386
  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
  3387
    using assms affine_hull_UNIV[of S] by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3388
  then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3389
    unfolding rel_interior interior_def by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3390
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3391
60303
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60176
diff changeset
  3392
lemma rel_interior_interior:
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60176
diff changeset
  3393
  fixes S :: "'n::euclidean_space set"
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60176
diff changeset
  3394
  assumes "affine hull S = UNIV"
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60176
diff changeset
  3395
  shows "rel_interior S = interior S"
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60176
diff changeset
  3396
  using assms unfolding rel_interior interior_def by auto
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60176
diff changeset
  3397
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3398
lemma rel_interior_open:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3399
  fixes S :: "'n::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3400
  assumes "open S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3401
  shows "rel_interior S = S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3402
  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
  3403
60800
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
  3404
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
  3405
  by (simp add: interior_open)
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
  3406
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3407
lemma interior_rel_interior_gen:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3408
  fixes S :: "'n::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3409
  shows "interior S = (if aff_dim S = int(DIM('n)) then rel_interior S else {})"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3410
  by (metis interior_rel_interior low_dim_interior)
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3411
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  3412
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
  3413
  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
  3414
  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
  3415
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
  3416
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  3417
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
  3418
  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
  3419
  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
  3420
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
  3421
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  3422
lemma rel_interior_affine_hull [simp]:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3423
  fixes S :: "'n::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3424
  shows "rel_interior (affine hull S) = affine hull S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3425
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3426
  have *: "rel_interior (affine hull S) \<subseteq> affine hull S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3427
    using rel_interior_subset by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3428
  {
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3429
    fix x
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3430
    assume x: "x \<in> affine hull S"
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  3431
    define e :: real where "e = 1"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3432
    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
  3433
      using hull_hull[of _ S] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3434
    then have "x \<in> rel_interior (affine hull S)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3435
      using x rel_interior_ball[of "affine hull S"] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3436
  }
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3437
  then show ?thesis using * by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3438
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3439
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  3440
lemma rel_interior_UNIV [simp]: "rel_interior (UNIV :: ('n::euclidean_space) set) = UNIV"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3441
  by (metis open_UNIV rel_interior_open)
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3442
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3443
lemma rel_interior_convex_shrink:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3444
  fixes S :: "'a::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3445
  assumes "convex S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3446
    and "c \<in> rel_interior S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3447
    and "x \<in> S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3448
    and "0 < e"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3449
    and "e \<le> 1"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3450
  shows "x - e *\<^sub>R (x - c) \<in> rel_interior S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3451
proof -
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  3452
  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
  3453
    using assms(2) unfolding  mem_rel_interior_ball by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3454
  {
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3455
    fix y
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3456
    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
  3457
    have *: "y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  3458
      using \<open>e > 0\<close> by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3459
    have "x \<in> affine hull S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3460
      using assms hull_subset[of S] by auto
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  3461
    moreover have "1 / e + - ((1 - e) / e) = 1"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  3462
      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
  3463
    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
  3464
      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
  3465
      by (simp add: algebra_simps)
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61880
diff changeset
  3466
    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
  3467
      unfolding dist_norm norm_scaleR[symmetric]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3468
      apply (rule arg_cong[where f=norm])
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  3469
      using \<open>e > 0\<close>
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3470
      apply (auto simp add: euclidean_eq_iff[where 'a='a] field_simps inner_simps)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3471
      done
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61880
diff changeset
  3472
    also have "\<dots> = \<bar>1/e\<bar> * norm (x - e *\<^sub>R (x - c) - y)"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3473
      by (auto intro!:arg_cong[where f=norm] simp add: algebra_simps)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3474
    also have "\<dots> < d"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  3475
      using as[unfolded dist_norm] and \<open>e > 0\<close>
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  3476
      by (auto simp add:pos_divide_less_eq[OF \<open>e > 0\<close>] mult.commute)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3477
    finally have "y \<in> S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3478
      apply (subst *)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3479
      apply (rule assms(1)[unfolded convex_alt,rule_format])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3480
      apply (rule d[unfolded subset_eq,rule_format])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3481
      unfolding mem_ball
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3482
      using assms(3-5) **
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3483
      apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3484
      done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3485
  }
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3486
  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
  3487
    by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3488
  moreover have "e * d > 0"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  3489
    using \<open>e > 0\<close> \<open>d > 0\<close> by simp
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3490
  moreover have c: "c \<in> S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3491
    using assms rel_interior_subset by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3492
  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
  3493
    using convexD_alt[of S x c e]
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3494
    apply (simp add: algebra_simps)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3495
    using assms
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3496
    apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3497
    done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3498
  ultimately show ?thesis
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  3499
    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
  3500
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3501
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3502
lemma interior_real_semiline:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3503
  fixes a :: real
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3504
  shows "interior {a..} = {a<..}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3505
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3506
  {
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3507
    fix y
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3508
    assume "a < y"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3509
    then have "y \<in> interior {a..}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3510
      apply (simp add: mem_interior)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3511
      apply (rule_tac x="(y-a)" in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3512
      apply (auto simp add: dist_norm)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3513
      done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3514
  }
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3515
  moreover
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3516
  {
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3517
    fix y
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3518
    assume "y \<in> interior {a..}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3519
    then obtain e where e: "e > 0" "cball y e \<subseteq> {a..}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3520
      using mem_interior_cball[of y "{a..}"] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3521
    moreover from e have "y - e \<in> cball y e"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3522
      by (auto simp add: 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
  3523
    ultimately have "a \<le> y - e" by blast
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3524
    then have "a < y" using e by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3525
  }
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3526
  ultimately show ?thesis by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3527
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3528
61880
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3529
lemma continuous_ge_on_Ioo:
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3530
  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
  3531
  shows "g (x::real) \<ge> (a::real)"
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3532
proof-
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3533
  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
  3534
  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
  3535
  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
  3536
  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
  3537
    by (auto simp: continuous_on_closed_vimage)
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3538
  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
  3539
  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
  3540
qed
61880
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3541
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3542
lemma interior_real_semiline':
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3543
  fixes a :: real
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3544
  shows "interior {..a} = {..<a}"
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3545
proof -
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3546
  {
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3547
    fix y
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3548
    assume "a > y"
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3549
    then have "y \<in> interior {..a}"
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3550
      apply (simp add: mem_interior)
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3551
      apply (rule_tac x="(a-y)" in exI)
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3552
      apply (auto simp add: dist_norm)
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3553
      done
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3554
  }
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3555
  moreover
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3556
  {
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3557
    fix y
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3558
    assume "y \<in> interior {..a}"
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3559
    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
  3560
      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
  3561
    moreover from e have "y + e \<in> cball y e"
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3562
      by (auto simp add: cball_def dist_norm)
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3563
    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
  3564
    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
  3565
  }
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3566
  ultimately show ?thesis by auto
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3567
qed
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3568
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3569
lemma interior_atLeastAtMost_real: "interior {a..b} = {a<..<b :: real}"
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3570
proof-
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3571
  have "{a..b} = {a..} \<inter> {..b}" by auto
62087
44841d07ef1d revisions to limits and derivatives, plus new lemmas
paulson
parents: 61952
diff changeset
  3572
  also have "interior ... = {a<..} \<inter> {..<b}"
61880
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3573
    by (simp add: interior_real_semiline interior_real_semiline')
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3574
  also have "... = {a<..<b}" by auto
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3575
  finally show ?thesis .
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3576
qed
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3577
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3578
lemma frontier_real_Iic:
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3579
  fixes a :: real
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3580
  shows "frontier {..a} = {a}"
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3581
  unfolding frontier_def by (auto simp add: interior_real_semiline')
ff4d33058566 moved some theorems from the CLT proof; reordered some theorems / notation
hoelzl
parents: 61848
diff changeset
  3582
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  3583
lemma rel_interior_real_box:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3584
  fixes a b :: real
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3585
  assumes "a < b"
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  3586
  shows "rel_interior {a .. b} = {a <..< b}"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3587
proof -
54775
2d3df8633dad prefer box over greaterThanLessThan on euclidean_space
immler
parents: 54465
diff changeset
  3588
  have "box a b \<noteq> {}"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3589
    using assms
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3590
    unfolding set_eq_iff
56189
c4daa97ac57a removed dependencies on theory Ordered_Euclidean_Space
immler
parents: 56188
diff changeset
  3591
    by (auto intro!: exI[of _ "(a + b) / 2"] simp: box_def)
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3592
  then show ?thesis
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  3593
    using interior_rel_interior_gen[of "cbox a b", symmetric]
62390
842917225d56 more canonical names
nipkow
parents: 62131
diff changeset
  3594
    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
  3595
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3596
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3597
lemma rel_interior_real_semiline:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3598
  fixes a :: real
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3599
  shows "rel_interior {a..} = {a<..}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3600
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3601
  have *: "{a<..} \<noteq> {}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3602
    unfolding set_eq_iff by (auto intro!: exI[of _ "a + 1"])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3603
  then show ?thesis using interior_real_semiline interior_rel_interior_gen[of "{a..}"]
62390
842917225d56 more canonical names
nipkow
parents: 62131
diff changeset
  3604
    by (auto split: if_split_asm)
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3605
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3606
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  3607
subsubsection \<open>Relative open sets\<close>
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3608
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3609
definition "rel_open S \<longleftrightarrow> rel_interior S = S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3610
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3611
lemma rel_open: "rel_open S \<longleftrightarrow> openin (subtopology euclidean (affine hull S)) S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3612
  unfolding rel_open_def rel_interior_def
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3613
  apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3614
  using openin_subopen[of "subtopology euclidean (affine hull S)" S]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3615
  apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3616
  done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3617
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3618
lemma opein_rel_interior: "openin (subtopology euclidean (affine hull S)) (rel_interior S)"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3619
  apply (simp add: rel_interior_def)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3620
  apply (subst openin_subopen)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3621
  apply blast
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3622
  done
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3623
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  3624
lemma affine_rel_open:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3625
  fixes S :: "'n::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3626
  assumes "affine S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3627
  shows "rel_open S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3628
  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
  3629
  using assms rel_interior_affine_hull[of S] affine_hull_eq[of S]
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3630
  by metis
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3631
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  3632
lemma affine_closed:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3633
  fixes S :: "'n::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3634
  assumes "affine S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3635
  shows "closed S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3636
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3637
  {
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3638
    assume "S \<noteq> {}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3639
    then obtain L where L: "subspace L" "affine_parallel S L"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3640
      using assms affine_parallel_subspace[of S] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3641
    then obtain a where a: "S = (op + a ` L)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3642
      using affine_parallel_def[of L S] affine_parallel_commut by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3643
    from L have "closed L" using closed_subspace by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3644
    then have "closed S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3645
      using closed_translation a by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3646
  }
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3647
  then show ?thesis by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3648
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3649
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3650
lemma closure_affine_hull:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3651
  fixes S :: "'n::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3652
  shows "closure S \<subseteq> affine hull S"
44524
04ad69081646 generalize some lemmas
huffman
parents: 44523
diff changeset
  3653
  by (intro closure_minimal hull_subset affine_closed affine_affine_hull)
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3654
62948
7700f467892b lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 62843
diff changeset
  3655
lemma closure_same_affine_hull [simp]:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3656
  fixes S :: "'n::euclidean_space set"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3657
  shows "affine hull (closure S) = affine hull S"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3658
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3659
  have "affine hull (closure S) \<subseteq> affine hull S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3660
    using hull_mono[of "closure S" "affine hull S" "affine"]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3661
      closure_affine_hull[of S] hull_hull[of "affine" S]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3662
    by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3663
  moreover have "affine hull (closure S) \<supseteq> affine hull S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3664
    using hull_mono[of "S" "closure S" "affine"] closure_subset by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3665
  ultimately show ?thesis by auto
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  3666
qed
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  3667
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  3668
lemma closure_aff_dim:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3669
  fixes S :: "'n::euclidean_space set"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3670
  shows "aff_dim (closure S) = aff_dim S"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3671
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3672
  have "aff_dim S \<le> aff_dim (closure S)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3673
    using aff_dim_subset closure_subset by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3674
  moreover have "aff_dim (closure S) \<le> aff_dim (affine hull S)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3675
    using aff_dim_subset closure_affine_hull by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3676
  moreover have "aff_dim (affine hull S) = aff_dim S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3677
    using aff_dim_affine_hull by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3678
  ultimately show ?thesis by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3679
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3680
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3681
lemma rel_interior_closure_convex_shrink:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3682
  fixes S :: "_::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3683
  assumes "convex S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3684
    and "c \<in> rel_interior S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3685
    and "x \<in> closure S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3686
    and "e > 0"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3687
    and "e \<le> 1"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3688
  shows "x - e *\<^sub>R (x - c) \<in> rel_interior S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3689
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3690
  obtain d where "d > 0" and d: "ball c d \<inter> affine hull S \<subseteq> S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3691
    using assms(2) unfolding mem_rel_interior_ball by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3692
  have "\<exists>y \<in> S. norm (y - x) * (1 - e) < e * d"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3693
  proof (cases "x \<in> S")
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3694
    case True
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  3695
    then show ?thesis using \<open>e > 0\<close> \<open>d > 0\<close>
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3696
      apply (rule_tac bexI[where x=x])
56544
b60d5d119489 made mult_pos_pos a simp rule
nipkow
parents: 56541
diff changeset
  3697
      apply (auto)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3698
      done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3699
  next
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3700
    case False
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3701
    then have x: "x islimpt S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3702
      using assms(3)[unfolded closure_def] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3703
    show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3704
    proof (cases "e = 1")
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3705
      case True
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3706
      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
  3707
        using x[unfolded islimpt_approachable,THEN spec[where x=1]] by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3708
      then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3709
        apply (rule_tac x=y in bexI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3710
        unfolding True
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  3711
        using \<open>d > 0\<close>
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3712
        apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3713
        done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3714
    next
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3715
      case False
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3716
      then have "0 < e * d / (1 - e)" and *: "1 - e > 0"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  3717
        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
  3718
      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
  3719
        using x[unfolded islimpt_approachable,THEN spec[where x="e*d / (1 - e)"]] by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3720
      then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3721
        apply (rule_tac x=y in bexI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3722
        unfolding dist_norm
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3723
        using pos_less_divide_eq[OF *]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3724
        apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3725
        done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3726
    qed
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3727
  qed
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3728
  then obtain y where "y \<in> S" and y: "norm (y - x) * (1 - e) < e * d"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3729
    by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  3730
  define z where "z = c + ((1 - e) / e) *\<^sub>R (x - y)"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3731
  have *: "x - e *\<^sub>R (x - c) = y - e *\<^sub>R (y - z)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  3732
    unfolding z_def using \<open>e > 0\<close>
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3733
    by (auto simp add: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3734
  have zball: "z \<in> ball c d"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3735
    using mem_ball z_def dist_norm[of c]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3736
    using y and assms(4,5)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3737
    by (auto simp add:field_simps norm_minus_commute)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3738
  have "x \<in> affine hull S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3739
    using closure_affine_hull assms by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3740
  moreover have "y \<in> affine hull S"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  3741
    using \<open>y \<in> S\<close> hull_subset[of S] by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3742
  moreover have "c \<in> affine hull S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3743
    using assms rel_interior_subset hull_subset[of S] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3744
  ultimately have "z \<in> affine hull S"
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  3745
    using z_def affine_affine_hull[of S]
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3746
      mem_affine_3_minus [of "affine hull S" c x y "(1 - e) / e"]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3747
      assms
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3748
    by (auto simp add: field_simps)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3749
  then have "z \<in> S" using d zball by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3750
  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
  3751
    using zball open_ball[of c d] openE[of "ball c d" z] by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3752
  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
  3753
    by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3754
  then have "ball z d1 \<inter> affine hull S \<subseteq> S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3755
    using d by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3756
  then have "z \<in> rel_interior S"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  3757
    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
  3758
  then have "y - e *\<^sub>R (y - z) \<in> rel_interior S"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  3759
    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
  3760
  then show ?thesis using * by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3761
qed
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3762
62620
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  3763
lemma rel_interior_eq:
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  3764
   "rel_interior s = s \<longleftrightarrow> openin(subtopology euclidean (affine hull s)) s"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  3765
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
  3766
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  3767
lemma rel_interior_openin:
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  3768
   "openin(subtopology euclidean (affine hull s)) s \<Longrightarrow> rel_interior s = s"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  3769
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
  3770
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3771
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  3772
subsubsection\<open>Relative interior preserves under linear transformations\<close>
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3773
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3774
lemma rel_interior_translation_aux:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3775
  fixes a :: "'n::euclidean_space"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3776
  shows "((\<lambda>x. a + x) ` rel_interior S) \<subseteq> rel_interior ((\<lambda>x. a + x) ` S)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3777
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3778
  {
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3779
    fix x
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3780
    assume x: "x \<in> rel_interior S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3781
    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
  3782
      using mem_rel_interior[of x S] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3783
    then have "open ((\<lambda>x. a + x) ` T)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3784
      and "a + x \<in> ((\<lambda>x. a + x) ` T) \<inter> ((\<lambda>x. a + x) ` S)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3785
      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
  3786
      using affine_hull_translation[of a S] open_translation[of T a] x by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3787
    then have "a + x \<in> rel_interior ((\<lambda>x. a + x) ` S)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3788
      using mem_rel_interior[of "a+x" "((\<lambda>x. a + x) ` S)"] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3789
  }
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3790
  then show ?thesis by auto
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
  3791
qed
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3792
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3793
lemma rel_interior_translation:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3794
  fixes a :: "'n::euclidean_space"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3795
  shows "rel_interior ((\<lambda>x. a + x) ` S) = (\<lambda>x. a + x) ` rel_interior S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3796
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3797
  have "(\<lambda>x. (-a) + x) ` rel_interior ((\<lambda>x. a + x) ` S) \<subseteq> rel_interior S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3798
    using rel_interior_translation_aux[of "-a" "(\<lambda>x. a + x) ` S"]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3799
      translation_assoc[of "-a" "a"]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3800
    by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3801
  then have "((\<lambda>x. a + x) ` rel_interior S) \<supseteq> rel_interior ((\<lambda>x. a + x) ` S)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3802
    using translation_inverse_subset[of a "rel_interior (op + a ` S)" "rel_interior S"]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3803
    by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3804
  then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3805
    using rel_interior_translation_aux[of a S] by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3806
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3807
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3808
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3809
lemma affine_hull_linear_image:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3810
  assumes "bounded_linear f"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3811
  shows "f ` (affine hull s) = affine hull f ` s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3812
  apply rule
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3813
  unfolding subset_eq ball_simps
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3814
  apply (rule_tac[!] hull_induct, rule hull_inc)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3815
  prefer 3
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3816
  apply (erule imageE)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3817
  apply (rule_tac x=xa in image_eqI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3818
  apply assumption
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3819
  apply (rule hull_subset[unfolded subset_eq, rule_format])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3820
  apply assumption
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3821
proof -
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3822
  interpret f: bounded_linear f by fact
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3823
  show "affine {x. f x \<in> affine hull f ` s}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3824
    unfolding affine_def
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3825
    by (auto simp add: f.scaleR f.add affine_affine_hull[unfolded affine_def, rule_format])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3826
  show "affine {x. x \<in> f ` (affine hull s)}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3827
    using affine_affine_hull[unfolded affine_def, of s]
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3828
    unfolding affine_def by (auto simp add: f.scaleR [symmetric] f.add [symmetric])
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3829
qed auto
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3830
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3831
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3832
lemma rel_interior_injective_on_span_linear_image:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3833
  fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3834
    and S :: "'m::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3835
  assumes "bounded_linear f"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3836
    and "inj_on f (span S)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3837
  shows "rel_interior (f ` S) = f ` (rel_interior S)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3838
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3839
  {
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3840
    fix z
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3841
    assume z: "z \<in> rel_interior (f ` S)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3842
    then have "z \<in> f ` S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3843
      using rel_interior_subset[of "f ` S"] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3844
    then obtain x where x: "x \<in> S" "f x = z" by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3845
    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
  3846
      using z rel_interior_cball[of "f ` S"] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3847
    obtain K where K: "K > 0" "\<And>x. norm (f x) \<le> norm x * K"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3848
     using assms Real_Vector_Spaces.bounded_linear.pos_bounded[of f] by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  3849
    define e1 where "e1 = 1 / K"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3850
    then have e1: "e1 > 0" "\<And>x. e1 * norm (f x) \<le> norm x"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3851
      using K pos_le_divide_eq[of e1] by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  3852
    define e where "e = e1 * e2"
56544
b60d5d119489 made mult_pos_pos a simp rule
nipkow
parents: 56541
diff changeset
  3853
    then have "e > 0" using e1 e2 by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3854
    {
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3855
      fix y
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3856
      assume y: "y \<in> cball x e \<inter> affine hull S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3857
      then have h1: "f y \<in> affine hull (f ` S)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3858
        using affine_hull_linear_image[of f S] assms by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3859
      from y have "norm (x-y) \<le> e1 * e2"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3860
        using cball_def[of x e] dist_norm[of x y] e_def by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3861
      moreover have "f x - f y = f (x - y)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3862
        using assms linear_sub[of f x y] linear_conv_bounded_linear[of f] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3863
      moreover have "e1 * norm (f (x-y)) \<le> norm (x - y)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3864
        using e1 by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3865
      ultimately have "e1 * norm ((f x)-(f y)) \<le> e1 * e2"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3866
        by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3867
      then have "f y \<in> cball z e2"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3868
        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
  3869
      then have "f y \<in> f ` S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3870
        using y e2 h1 by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3871
      then have "y \<in> S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3872
        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
  3873
          inj_on_image_mem_iff [OF \<open>inj_on f (span S)\<close>]
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3874
        by (metis Int_iff span_inc subsetCE)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3875
    }
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3876
    then have "z \<in> f ` (rel_interior S)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  3877
      using mem_rel_interior_cball[of x S] \<open>e > 0\<close> x by auto
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  3878
  }
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3879
  moreover
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3880
  {
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3881
    fix x
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3882
    assume x: "x \<in> rel_interior S"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  3883
    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
  3884
      using rel_interior_cball[of S] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3885
    have "x \<in> S" using x rel_interior_subset by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3886
    then have *: "f x \<in> f ` S" by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3887
    have "\<forall>x\<in>span S. f x = 0 \<longrightarrow> x = 0"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3888
      using assms subspace_span linear_conv_bounded_linear[of f]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3889
        linear_injective_on_subspace_0[of f "span S"]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3890
      by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3891
    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
  3892
      using assms injective_imp_isometric[of "span S" f]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3893
        subspace_span[of S] closed_subspace[of "span S"]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3894
      by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  3895
    define e where "e = e1 * e2"
56544
b60d5d119489 made mult_pos_pos a simp rule
nipkow
parents: 56541
diff changeset
  3896
    hence "e > 0" using e1 e2 by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3897
    {
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3898
      fix y
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3899
      assume y: "y \<in> cball (f x) e \<inter> affine hull (f ` S)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3900
      then have "y \<in> f ` (affine hull S)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3901
        using affine_hull_linear_image[of f S] assms by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3902
      then obtain xy where xy: "xy \<in> affine hull S" "f xy = y" by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3903
      with y have "norm (f x - f xy) \<le> e1 * e2"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3904
        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
  3905
      moreover have "f x - f xy = f (x - xy)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3906
        using assms linear_sub[of f x xy] linear_conv_bounded_linear[of f] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3907
      moreover have *: "x - xy \<in> span S"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  3908
        using subspace_sub[of "span S" x xy] subspace_span \<open>x \<in> S\<close> xy
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3909
          affine_hull_subset_span[of S] span_inc
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3910
        by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3911
      moreover from * have "e1 * norm (x - xy) \<le> norm (f (x - xy))"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3912
        using e1 by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3913
      ultimately have "e1 * norm (x - xy) \<le> e1 * e2"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3914
        by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3915
      then have "xy \<in> cball x e2"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3916
        using cball_def[of x e2] dist_norm[of x xy] e1 by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3917
      then have "y \<in> f ` S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3918
        using xy e2 by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3919
    }
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3920
    then have "f x \<in> rel_interior (f ` S)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  3921
      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
  3922
  }
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3923
  ultimately show ?thesis by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3924
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3925
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3926
lemma rel_interior_injective_linear_image:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3927
  fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3928
  assumes "bounded_linear f"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3929
    and "inj f"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3930
  shows "rel_interior (f ` S) = f ` (rel_interior S)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3931
  using assms rel_interior_injective_on_span_linear_image[of f S]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3932
    subset_inj_on[of f "UNIV" "span S"]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3933
  by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3934
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3935
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  3936
subsection\<open>Some Properties of subset of standard basis\<close>
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3937
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3938
lemma affine_hull_substd_basis:
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3939
  assumes "d \<subseteq> Basis"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3940
  shows "affine hull (insert 0 d) = {x::'a::euclidean_space. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3941
  (is "affine hull (insert 0 ?A) = ?B")
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3942
proof -
61076
bdc1e2f0a86a eliminated \<Colon>;
wenzelm
parents: 60974
diff changeset
  3943
  have *: "\<And>A. op + (0::'a) ` A = A" "\<And>A. op + (- (0::'a)) ` A = A"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3944
    by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3945
  show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3946
    unfolding affine_hull_insert_span_gen span_substd_basis[OF assms,symmetric] * ..
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3947
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3948
60303
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60176
diff changeset
  3949
lemma affine_hull_convex_hull [simp]: "affine hull (convex hull S) = affine hull S"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3950
  by (metis Int_absorb1 Int_absorb2 convex_hull_subset_affine_hull hull_hull hull_mono hull_subset)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3951
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  3952
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  3953
subsection \<open>Openness and compactness are preserved by convex hull operation.\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3954
34964
4e8be3c04d37 Replaced vec1 and dest_vec1 by abbreviation.
hoelzl
parents: 34915
diff changeset
  3955
lemma open_convex_hull[intro]:
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3956
  fixes s :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3957
  assumes "open s"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3958
  shows "open (convex hull s)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3959
  unfolding open_contains_cball convex_hull_explicit
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3960
  unfolding mem_Collect_eq ball_simps(8)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3961
proof (rule, rule)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3962
  fix a
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  3963
  assume "\<exists>sa u. finite sa \<and> sa \<subseteq> s \<and> (\<forall>x\<in>sa. 0 \<le> u x) \<and> setsum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = a"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3964
  then obtain t u where obt: "finite t" "t\<subseteq>s" "\<forall>x\<in>t. 0 \<le> u x" "setsum u t = 1" "(\<Sum>v\<in>t. u v *\<^sub>R v) = a"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3965
    by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3966
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3967
  from assms[unfolded open_contains_cball] obtain b
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3968
    where b: "\<forall>x\<in>s. 0 < b x \<and> cball x (b x) \<subseteq> s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3969
    using bchoice[of s "\<lambda>x e. e > 0 \<and> cball x e \<subseteq> s"] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3970
  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
  3971
    using obt by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  3972
  define i where "i = b ` t"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3973
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3974
  show "\<exists>e > 0.
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3975
    cball a e \<subseteq> {y. \<exists>sa u. finite sa \<and> sa \<subseteq> s \<and> (\<forall>x\<in>sa. 0 \<le> u x) \<and> setsum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = y}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3976
    apply (rule_tac x = "Min i" in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3977
    unfolding subset_eq
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3978
    apply rule
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3979
    defer
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3980
    apply rule
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3981
    unfolding mem_Collect_eq
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3982
  proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3983
    show "0 < Min i"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  3984
      unfolding i_def and Min_gr_iff[OF finite_imageI[OF obt(1)] \<open>b ` t\<noteq>{}\<close>]
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3985
      using b
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3986
      apply simp
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3987
      apply rule
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3988
      apply (erule_tac x=x in ballE)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  3989
      using \<open>t\<subseteq>s\<close>
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3990
      apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3991
      done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3992
  next
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3993
    fix y
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3994
    assume "y \<in> cball a (Min i)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3995
    then have y: "norm (a - y) \<le> Min i"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3996
      unfolding dist_norm[symmetric] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3997
    {
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3998
      fix x
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  3999
      assume "x \<in> t"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4000
      then have "Min i \<le> b x"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4001
        unfolding i_def
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4002
        apply (rule_tac Min_le)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4003
        using obt(1)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4004
        apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4005
        done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4006
      then have "x + (y - a) \<in> cball x (b x)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4007
        using y unfolding mem_cball dist_norm by auto
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4008
      moreover from \<open>x\<in>t\<close> have "x \<in> s"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4009
        using obt(2) by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4010
      ultimately have "x + (y - a) \<in> s"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  4011
        using y and b[THEN bspec[where x=x]] unfolding subset_eq by fast
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4012
    }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4013
    moreover
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4014
    have *: "inj_on (\<lambda>v. v + (y - a)) t"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4015
      unfolding inj_on_def by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4016
    have "(\<Sum>v\<in>(\<lambda>v. v + (y - a)) ` t. u (v - (y - a))) = 1"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  4017
      unfolding setsum.reindex[OF *] o_def using obt(4) by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4018
    moreover have "(\<Sum>v\<in>(\<lambda>v. v + (y - a)) ` t. u (v - (y - a)) *\<^sub>R v) = y"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  4019
      unfolding setsum.reindex[OF *] o_def using obt(4,5)
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  4020
      by (simp add: setsum.distrib setsum_subtractf scaleR_left.setsum[symmetric] scaleR_right_distrib)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4021
    ultimately
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4022
    show "\<exists>sa u. finite sa \<and> (\<forall>x\<in>sa. x \<in> s) \<and> (\<forall>x\<in>sa. 0 \<le> u x) \<and> setsum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = y"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4023
      apply (rule_tac x="(\<lambda>v. v + (y - a)) ` t" in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4024
      apply (rule_tac x="\<lambda>v. u (v - (y - a))" in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4025
      using obt(1, 3)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4026
      apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4027
      done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4028
  qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4029
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4030
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4031
lemma compact_convex_combinations:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4032
  fixes s t :: "'a::real_normed_vector set"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4033
  assumes "compact s" "compact t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4034
  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
  4035
proof -
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4036
  let ?X = "{0..1} \<times> s \<times> t"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4037
  let ?h = "(\<lambda>z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4038
  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"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4039
    apply (rule set_eqI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4040
    unfolding image_iff mem_Collect_eq
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4041
    apply rule
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4042
    apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4043
    apply (rule_tac x=u in rev_bexI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4044
    apply simp
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4045
    apply (erule rev_bexI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4046
    apply (erule rev_bexI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4047
    apply simp
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4048
    apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4049
    done
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  4050
  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
  4051
    unfolding continuous_on by (rule ballI) (intro tendsto_intros)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4052
  then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4053
    unfolding *
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4054
    apply (rule compact_continuous_image)
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  4055
    apply (intro compact_Times compact_Icc assms)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4056
    done
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4057
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4058
44525
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  4059
lemma finite_imp_compact_convex_hull:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4060
  fixes s :: "'a::real_normed_vector set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4061
  assumes "finite s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4062
  shows "compact (convex hull s)"
44525
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  4063
proof (cases "s = {}")
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4064
  case True
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4065
  then show ?thesis by simp
44525
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  4066
next
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4067
  case False
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4068
  with assms show ?thesis
44525
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  4069
  proof (induct rule: finite_ne_induct)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4070
    case (singleton x)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4071
    show ?case by simp
44525
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  4072
  next
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  4073
    case (insert x A)
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  4074
    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
  4075
    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
  4076
    have "continuous_on ?T ?f"
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  4077
      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
  4078
    moreover have "compact ?T"
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  4079
      by (intro compact_Times compact_Icc insert)
44525
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  4080
    ultimately have "compact (?f ` ?T)"
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  4081
      by (rule compact_continuous_image)
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  4082
    also have "?f ` ?T = convex hull (insert x A)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4083
      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
  4084
      apply safe
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  4085
      apply (rule_tac x=a in exI, simp)
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  4086
      apply (rule_tac x="1 - a" in exI, simp)
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  4087
      apply fast
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  4088
      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
  4089
      done
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  4090
    finally show "compact (convex hull (insert x A))" .
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  4091
  qed
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  4092
qed
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  4093
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4094
lemma compact_convex_hull:
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4095
  fixes s :: "'a::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4096
  assumes "compact s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4097
  shows "compact (convex hull s)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4098
proof (cases "s = {}")
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4099
  case True
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4100
  then show ?thesis using compact_empty by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4101
next
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4102
  case False
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4103
  then obtain w where "w \<in> s" by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4104
  show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4105
    unfolding caratheodory[of s]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4106
  proof (induct ("DIM('a) + 1"))
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4107
    case 0
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4108
    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
  4109
      using compact_empty by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4110
    from 0 show ?case unfolding * by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4111
  next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4112
    case (Suc n)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4113
    show ?case
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4114
    proof (cases "n = 0")
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4115
      case True
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4116
      have "{x. \<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t} = s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4117
        unfolding set_eq_iff and mem_Collect_eq
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4118
      proof (rule, rule)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4119
        fix x
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4120
        assume "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4121
        then obtain t where t: "finite t" "t \<subseteq> s" "card t \<le> Suc n" "x \<in> convex hull t"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4122
          by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4123
        show "x \<in> s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4124
        proof (cases "card t = 0")
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4125
          case True
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4126
          then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4127
            using t(4) unfolding card_0_eq[OF t(1)] by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4128
        next
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4129
          case False
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4130
          then have "card t = Suc 0" using t(3) \<open>n=0\<close> by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4131
          then obtain a where "t = {a}" unfolding card_Suc_eq by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4132
          then show ?thesis using t(2,4) by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4133
        qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4134
      next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4135
        fix x assume "x\<in>s"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4136
        then show "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4137
          apply (rule_tac x="{x}" in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4138
          unfolding convex_hull_singleton
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4139
          apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4140
          done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4141
      qed
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4142
      then show ?thesis using assms by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4143
    next
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4144
      case False
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4145
      have "{x. \<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t} =
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4146
        {(1 - u) *\<^sub>R x + u *\<^sub>R y | x y u.
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4147
          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}}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4148
        unfolding set_eq_iff and mem_Collect_eq
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4149
      proof (rule, rule)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4150
        fix x
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4151
        assume "\<exists>u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \<and>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4152
          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)"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4153
        then obtain u v c t where obt: "x = (1 - c) *\<^sub>R u + c *\<^sub>R v"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4154
          "0 \<le> c \<and> c \<le> 1" "u \<in> s" "finite t" "t \<subseteq> s" "card t \<le> n"  "v \<in> convex hull t"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4155
          by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4156
        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
  4157
          apply (rule convexD_alt)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4158
          using obt(2) and convex_convex_hull and hull_subset[of "insert u t" convex]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4159
          using obt(7) and hull_mono[of t "insert u t"]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4160
          apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4161
          done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4162
        ultimately 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
  4163
          apply (rule_tac x="insert u t" in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4164
          apply (auto simp add: card_insert_if)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4165
          done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4166
      next
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4167
        fix x
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4168
        assume "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4169
        then obtain t where t: "finite t" "t \<subseteq> s" "card t \<le> Suc n" "x \<in> convex hull t"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4170
          by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4171
        show "\<exists>u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \<and>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4172
          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)"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4173
        proof (cases "card t = Suc n")
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4174
          case False
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4175
          then have "card t \<le> n" using t(3) by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4176
          then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4177
            apply (rule_tac x=w in exI, rule_tac x=x in exI, rule_tac x=1 in exI)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4178
            using \<open>w\<in>s\<close> and t
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4179
            apply (auto intro!: exI[where x=t])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4180
            done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4181
        next
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4182
          case True
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4183
          then obtain a u where au: "t = insert a u" "a\<notin>u"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4184
            apply (drule_tac card_eq_SucD)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4185
            apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4186
            done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4187
          show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4188
          proof (cases "u = {}")
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4189
            case True
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4190
            then have "x = a" using t(4)[unfolded au] by auto
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4191
            show ?thesis unfolding \<open>x = a\<close>
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4192
              apply (rule_tac x=a in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4193
              apply (rule_tac x=a in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4194
              apply (rule_tac x=1 in exI)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4195
              using t and \<open>n \<noteq> 0\<close>
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4196
              unfolding au
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4197
              apply (auto intro!: exI[where x="{a}"])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4198
              done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4199
          next
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4200
            case False
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4201
            obtain ux vx b where obt: "ux\<ge>0" "vx\<ge>0" "ux + vx = 1"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4202
              "b \<in> convex hull u" "x = ux *\<^sub>R a + vx *\<^sub>R b"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4203
              using t(4)[unfolded au convex_hull_insert[OF False]]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4204
              by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4205
            have *: "1 - vx = ux" using obt(3) by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4206
            show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4207
              apply (rule_tac x=a in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4208
              apply (rule_tac x=b in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4209
              apply (rule_tac x=vx in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4210
              using obt and t(1-3)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4211
              unfolding au and * using card_insert_disjoint[OF _ au(2)]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4212
              apply (auto intro!: exI[where x=u])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4213
              done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4214
          qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4215
        qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4216
      qed
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4217
      then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4218
        using compact_convex_combinations[OF assms Suc] by simp
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4219
    qed
36362
06475a1547cb fix lots of looping simp calls and other warnings
huffman
parents: 36341
diff changeset
  4220
  qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4221
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4222
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4223
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4224
subsection \<open>Extremal points of a simplex are some vertices.\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4225
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4226
lemma dist_increases_online:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4227
  fixes a b d :: "'a::real_inner"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4228
  assumes "d \<noteq> 0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4229
  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
  4230
proof (cases "inner a d - inner b d > 0")
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4231
  case True
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4232
  then have "0 < inner d d + (inner a d * 2 - inner b d * 2)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4233
    apply (rule_tac add_pos_pos)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4234
    using assms
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4235
    apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4236
    done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4237
  then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4238
    apply (rule_tac disjI2)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4239
    unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4240
    apply  (simp add: algebra_simps inner_commute)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4241
    done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4242
next
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4243
  case False
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4244
  then have "0 < inner d d + (inner b d * 2 - inner a d * 2)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4245
    apply (rule_tac add_pos_nonneg)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4246
    using assms
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4247
    apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4248
    done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4249
  then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4250
    apply (rule_tac disjI1)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4251
    unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4252
    apply (simp add: algebra_simps inner_commute)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4253
    done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4254
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4255
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4256
lemma norm_increases_online:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4257
  fixes d :: "'a::real_inner"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4258
  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
  4259
  using dist_increases_online[of d a 0] unfolding dist_norm by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4260
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4261
lemma simplex_furthest_lt:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4262
  fixes s :: "'a::real_inner set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4263
  assumes "finite s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4264
  shows "\<forall>x \<in> convex hull s.  x \<notin> s \<longrightarrow> (\<exists>y \<in> convex hull s. norm (x - a) < norm(y - a))"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4265
  using assms
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4266
proof induct
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4267
  fix x s
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4268
  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))"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4269
  show "\<forall>xa\<in>convex hull insert x s. xa \<notin> insert x s \<longrightarrow>
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4270
    (\<exists>y\<in>convex hull insert x s. norm (xa - a) < norm (y - a))"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4271
  proof (rule, rule, cases "s = {}")
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4272
    case False
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4273
    fix y
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4274
    assume y: "y \<in> convex hull insert x s" "y \<notin> insert x s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4275
    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
  4276
      using y(1)[unfolded convex_hull_insert[OF False]] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4277
    show "\<exists>z\<in>convex hull insert x s. norm (y - a) < norm (z - a)"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4278
    proof (cases "y \<in> convex hull s")
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4279
      case True
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4280
      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
  4281
        using as(3)[THEN bspec[where x=y]] and y(2) by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4282
      then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4283
        apply (rule_tac x=z in bexI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4284
        unfolding convex_hull_insert[OF False]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4285
        apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4286
        done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4287
    next
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4288
      case False
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4289
      show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4290
        using obt(3)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4291
      proof (cases "u = 0", case_tac[!] "v = 0")
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4292
        assume "u = 0" "v \<noteq> 0"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4293
        then have "y = b" using obt by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4294
        then show ?thesis using False and obt(4) by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4295
      next
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4296
        assume "u \<noteq> 0" "v = 0"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4297
        then have "y = x" using obt by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4298
        then show ?thesis using y(2) by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4299
      next
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4300
        assume "u \<noteq> 0" "v \<noteq> 0"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4301
        then obtain w where w: "w>0" "w<u" "w<v"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4302
          using real_lbound_gt_zero[of u v] and obt(1,2) by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4303
        have "x \<noteq> b"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4304
        proof
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4305
          assume "x = b"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4306
          then have "y = b" unfolding obt(5)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4307
            using obt(3) by (auto simp add: scaleR_left_distrib[symmetric])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4308
          then show False using obt(4) and False by simp
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4309
        qed
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4310
        then have *: "w *\<^sub>R (x - b) \<noteq> 0" using w(1) by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4311
        show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4312
          using dist_increases_online[OF *, of a y]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4313
        proof (elim disjE)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4314
          assume "dist a y < dist a (y + w *\<^sub>R (x - b))"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4315
          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
  4316
            unfolding dist_commute[of a]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4317
            unfolding dist_norm obt(5)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4318
            by (simp add: algebra_simps)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4319
          moreover have "(u + w) *\<^sub>R x + (v - w) *\<^sub>R b \<in> convex hull insert x s"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4320
            unfolding convex_hull_insert[OF \<open>s\<noteq>{}\<close>] and mem_Collect_eq
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4321
            apply (rule_tac x="u + w" in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4322
            apply rule
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4323
            defer
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4324
            apply (rule_tac x="v - w" in exI)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4325
            using \<open>u \<ge> 0\<close> and w and obt(3,4)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4326
            apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4327
            done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4328
          ultimately show ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4329
        next
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4330
          assume "dist a y < dist a (y - w *\<^sub>R (x - b))"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4331
          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
  4332
            unfolding dist_commute[of a]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4333
            unfolding dist_norm obt(5)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4334
            by (simp add: algebra_simps)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4335
          moreover have "(u - w) *\<^sub>R x + (v + w) *\<^sub>R b \<in> convex hull insert x s"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4336
            unfolding convex_hull_insert[OF \<open>s\<noteq>{}\<close>] and mem_Collect_eq
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4337
            apply (rule_tac x="u - w" in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4338
            apply rule
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4339
            defer
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4340
            apply (rule_tac x="v + w" in exI)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4341
            using \<open>u \<ge> 0\<close> and w and obt(3,4)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4342
            apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4343
            done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4344
          ultimately show ?thesis by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4345
        qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4346
      qed auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4347
    qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4348
  qed auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4349
qed (auto simp add: assms)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4350
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4351
lemma simplex_furthest_le:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4352
  fixes s :: "'a::real_inner set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4353
  assumes "finite s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4354
    and "s \<noteq> {}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4355
  shows "\<exists>y\<in>s. \<forall>x\<in> convex hull s. norm (x - a) \<le> norm (y - a)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4356
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4357
  have "convex hull s \<noteq> {}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4358
    using hull_subset[of s convex] and assms(2) by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4359
  then obtain x where x: "x \<in> convex hull s" "\<forall>y\<in>convex hull s. norm (y - a) \<le> norm (x - a)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4360
    using distance_attains_sup[OF finite_imp_compact_convex_hull[OF assms(1)], of a]
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4361
    unfolding dist_commute[of a]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4362
    unfolding dist_norm
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4363
    by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4364
  show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4365
  proof (cases "x \<in> s")
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4366
    case False
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4367
    then obtain y where "y \<in> convex hull s" "norm (x - a) < norm (y - a)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4368
      using simplex_furthest_lt[OF assms(1), THEN bspec[where x=x]] and x(1)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4369
      by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4370
    then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4371
      using x(2)[THEN bspec[where x=y]] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4372
  next
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4373
    case True
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4374
    with x show ?thesis by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4375
  qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4376
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4377
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4378
lemma simplex_furthest_le_exists:
44525
fbb777aec0d4 generalize lemma finite_imp_compact_convex_hull and related lemmas
huffman
parents: 44524
diff changeset
  4379
  fixes s :: "('a::real_inner) set"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4380
  shows "finite s \<Longrightarrow> \<forall>x\<in>(convex hull s). \<exists>y\<in>s. norm (x - a) \<le> norm (y - a)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4381
  using simplex_furthest_le[of s] by (cases "s = {}") auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4382
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4383
lemma simplex_extremal_le:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4384
  fixes s :: "'a::real_inner set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4385
  assumes "finite s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4386
    and "s \<noteq> {}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4387
  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)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4388
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4389
  have "convex hull s \<noteq> {}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4390
    using hull_subset[of s convex] and assms(2) by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4391
  then obtain u v where obt: "u \<in> convex hull s" "v \<in> convex hull s"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4392
    "\<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
  4393
    using compact_sup_maxdistance[OF finite_imp_compact_convex_hull[OF assms(1)]]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4394
    by (auto simp: dist_norm)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4395
  then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4396
  proof (cases "u\<notin>s \<or> v\<notin>s", elim disjE)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4397
    assume "u \<notin> s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4398
    then obtain y where "y \<in> convex hull s" "norm (u - v) < norm (y - v)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4399
      using simplex_furthest_lt[OF assms(1), THEN bspec[where x=u]] and obt(1)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4400
      by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4401
    then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4402
      using obt(3)[THEN bspec[where x=y], THEN bspec[where x=v]] and obt(2)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4403
      by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4404
  next
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4405
    assume "v \<notin> s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4406
    then obtain y where "y \<in> convex hull s" "norm (v - u) < norm (y - u)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4407
      using simplex_furthest_lt[OF assms(1), THEN bspec[where x=v]] and obt(2)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4408
      by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4409
    then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4410
      using obt(3)[THEN bspec[where x=u], THEN bspec[where x=y]] and obt(1)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4411
      by (auto simp add: norm_minus_commute)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4412
  qed auto
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  4413
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4414
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4415
lemma simplex_extremal_le_exists:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4416
  fixes s :: "'a::real_inner set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4417
  shows "finite s \<Longrightarrow> x \<in> convex hull s \<Longrightarrow> y \<in> convex hull s \<Longrightarrow>
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4418
    \<exists>u\<in>s. \<exists>v\<in>s. norm (x - y) \<le> norm (u - v)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4419
  using convex_hull_empty simplex_extremal_le[of s]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4420
  by(cases "s = {}") auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4421
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4422
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4423
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
  4424
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4425
definition closest_point :: "'a::{real_inner,heine_borel} set \<Rightarrow> 'a \<Rightarrow> 'a"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4426
  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
  4427
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4428
lemma closest_point_exists:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4429
  assumes "closed s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4430
    and "s \<noteq> {}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4431
  shows "closest_point s a \<in> s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4432
    and "\<forall>y\<in>s. dist a (closest_point s a) \<le> dist a y"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4433
  unfolding closest_point_def
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4434
  apply(rule_tac[!] someI2_ex)
62381
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
  4435
  apply (auto intro: distance_attains_inf[OF assms(1,2), of a])
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4436
  done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4437
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4438
lemma closest_point_in_set: "closed s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> closest_point s a \<in> s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4439
  by (meson closest_point_exists)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4440
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4441
lemma closest_point_le: "closed s \<Longrightarrow> x \<in> s \<Longrightarrow> dist a (closest_point s a) \<le> dist a x"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4442
  using closest_point_exists[of s] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4443
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4444
lemma closest_point_self:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4445
  assumes "x \<in> s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4446
  shows "closest_point s x = x"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4447
  unfolding closest_point_def
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4448
  apply (rule some1_equality, rule ex1I[of _ x])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4449
  using assms
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4450
  apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4451
  done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4452
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4453
lemma closest_point_refl: "closed s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> closest_point s x = x \<longleftrightarrow> x \<in> s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4454
  using closest_point_in_set[of s x] closest_point_self[of x s]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4455
  by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4456
36337
87b6c83d7ed7 generalize constant closest_point
huffman
parents: 36071
diff changeset
  4457
lemma closer_points_lemma:
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4458
  assumes "inner y z > 0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4459
  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
  4460
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4461
  have z: "inner z z > 0"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4462
    unfolding inner_gt_zero_iff using assms by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4463
  then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4464
    using assms
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4465
    apply (rule_tac x = "inner y z / inner z z" in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4466
    apply rule
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4467
    defer
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4468
  proof rule+
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4469
    fix v
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4470
    assume "0 < v" and "v \<le> inner y z / inner z z"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4471
    then show "norm (v *\<^sub>R z - y) < norm y"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4472
      unfolding norm_lt using z and assms
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4473
      by (simp add: field_simps inner_diff inner_commute mult_strict_left_mono[OF _ \<open>0<v\<close>])
56541
0e3abadbef39 made divide_pos_pos a simp rule
nipkow
parents: 56536
diff changeset
  4474
  qed auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4475
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4476
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4477
lemma closer_point_lemma:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4478
  assumes "inner (y - x) (z - x) > 0"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4479
  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
  4480
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4481
  obtain u where "u > 0"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4482
    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
  4483
    using closer_points_lemma[OF assms] by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4484
  show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4485
    apply (rule_tac x="min u 1" in exI)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4486
    using u[THEN spec[where x="min u 1"]] and \<open>u > 0\<close>
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4487
    unfolding dist_norm by (auto simp add: norm_minus_commute field_simps)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4488
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4489
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4490
lemma any_closest_point_dot:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4491
  assumes "convex s" "closed s" "x \<in> s" "y \<in> s" "\<forall>z\<in>s. dist a x \<le> dist a z"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4492
  shows "inner (a - x) (y - x) \<le> 0"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4493
proof (rule ccontr)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4494
  assume "\<not> ?thesis"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4495
  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
  4496
    using closer_point_lemma[of a x y] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4497
  let ?z = "(1 - u) *\<^sub>R x + u *\<^sub>R y"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4498
  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
  4499
    using convexD_alt[OF assms(1,3,4), of u] using u by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4500
  then show False
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4501
    using assms(5)[THEN bspec[where x="?z"]] and u(3)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4502
    by (auto simp add: dist_commute algebra_simps)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4503
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4504
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4505
lemma any_closest_point_unique:
36337
87b6c83d7ed7 generalize constant closest_point
huffman
parents: 36071
diff changeset
  4506
  fixes x :: "'a::real_inner"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4507
  assumes "convex s" "closed s" "x \<in> s" "y \<in> s"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4508
    "\<forall>z\<in>s. dist a x \<le> dist a z" "\<forall>z\<in>s. dist a y \<le> dist a z"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4509
  shows "x = y"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4510
  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
  4511
  unfolding norm_pths(1) and norm_le_square
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4512
  by (auto simp add: algebra_simps)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4513
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4514
lemma closest_point_unique:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4515
  assumes "convex s" "closed s" "x \<in> s" "\<forall>z\<in>s. dist a x \<le> dist a z"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4516
  shows "x = closest_point s a"
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  4517
  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
  4518
  using closest_point_exists[OF assms(2)] and assms(3) by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4519
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4520
lemma closest_point_dot:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4521
  assumes "convex s" "closed s" "x \<in> s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4522
  shows "inner (a - closest_point s a) (x - closest_point s a) \<le> 0"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4523
  apply (rule any_closest_point_dot[OF assms(1,2) _ assms(3)])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4524
  using closest_point_exists[OF assms(2)] and assms(3)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4525
  apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4526
  done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4527
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4528
lemma closest_point_lt:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4529
  assumes "convex s" "closed s" "x \<in> s" "x \<noteq> closest_point s a"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4530
  shows "dist a (closest_point s a) < dist a x"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4531
  apply (rule ccontr)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4532
  apply (rule_tac notE[OF assms(4)])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4533
  apply (rule closest_point_unique[OF assms(1-3), of a])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4534
  using closest_point_le[OF assms(2), of _ a]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4535
  apply fastforce
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4536
  done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4537
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4538
lemma closest_point_lipschitz:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4539
  assumes "convex s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4540
    and "closed s" "s \<noteq> {}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4541
  shows "dist (closest_point s x) (closest_point s y) \<le> dist x y"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4542
proof -
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4543
  have "inner (x - closest_point s x) (closest_point s y - closest_point s x) \<le> 0"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4544
    and "inner (y - closest_point s y) (closest_point s x - closest_point s y) \<le> 0"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4545
    apply (rule_tac[!] any_closest_point_dot[OF assms(1-2)])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4546
    using closest_point_exists[OF assms(2-3)]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4547
    apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4548
    done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4549
  then show ?thesis unfolding dist_norm and norm_le
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4550
    using inner_ge_zero[of "(x - closest_point s x) - (y - closest_point s y)"]
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4551
    by (simp add: inner_add inner_diff inner_commute)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4552
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4553
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4554
lemma continuous_at_closest_point:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4555
  assumes "convex s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4556
    and "closed s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4557
    and "s \<noteq> {}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4558
  shows "continuous (at x) (closest_point s)"
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  4559
  unfolding continuous_at_eps_delta
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4560
  using le_less_trans[OF closest_point_lipschitz[OF assms]] by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4561
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4562
lemma continuous_on_closest_point:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4563
  assumes "convex s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4564
    and "closed s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4565
    and "s \<noteq> {}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4566
  shows "continuous_on t (closest_point s)"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4567
  by (metis continuous_at_imp_continuous_on continuous_at_closest_point[OF assms])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4568
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4569
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4570
subsubsection \<open>Various point-to-set separating/supporting hyperplane theorems.\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4571
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4572
lemma supporting_hyperplane_closed_point:
36337
87b6c83d7ed7 generalize constant closest_point
huffman
parents: 36071
diff changeset
  4573
  fixes z :: "'a::{real_inner,heine_borel}"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4574
  assumes "convex s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4575
    and "closed s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4576
    and "s \<noteq> {}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4577
    and "z \<notin> s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4578
  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)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4579
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4580
  obtain y where "y \<in> s" and y: "\<forall>x\<in>s. 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
  4581
    by (metis distance_attains_inf[OF assms(2-3)]) 
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4582
  show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4583
    apply (rule_tac x="y - z" in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4584
    apply (rule_tac x="inner (y - z) y" in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4585
    apply (rule_tac x=y in bexI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4586
    apply rule
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4587
    defer
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4588
    apply rule
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4589
    defer
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4590
    apply rule
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4591
    apply (rule ccontr)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4592
    using \<open>y \<in> s\<close>
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4593
  proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4594
    show "inner (y - z) z < inner (y - z) y"
61762
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61738
diff changeset
  4595
      apply (subst diff_gt_0_iff_gt [symmetric])
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4596
      unfolding inner_diff_right[symmetric] and inner_gt_zero_iff
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4597
      using \<open>y\<in>s\<close> \<open>z\<notin>s\<close>
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4598
      apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4599
      done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4600
  next
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4601
    fix x
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4602
    assume "x \<in> s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4603
    have *: "\<forall>u. 0 \<le> u \<and> u \<le> 1 \<longrightarrow> dist z y \<le> dist z ((1 - u) *\<^sub>R y + u *\<^sub>R x)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4604
      using assms(1)[unfolded convex_alt] and y and \<open>x\<in>s\<close> and \<open>y\<in>s\<close> by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4605
    assume "\<not> inner (y - z) y \<le> inner (y - z) x"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4606
    then obtain v where "v > 0" "v \<le> 1" "dist (y + v *\<^sub>R (x - y)) z < dist y z"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4607
      using closer_point_lemma[of z y x] by (auto simp add: inner_diff)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4608
    then show False
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4609
      using *[THEN spec[where x=v]] by (auto simp add: dist_commute algebra_simps)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4610
  qed auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4611
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4612
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4613
lemma separating_hyperplane_closed_point:
36337
87b6c83d7ed7 generalize constant closest_point
huffman
parents: 36071
diff changeset
  4614
  fixes z :: "'a::{real_inner,heine_borel}"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4615
  assumes "convex s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4616
    and "closed s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4617
    and "z \<notin> s"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4618
  shows "\<exists>a b. inner a z < b \<and> (\<forall>x\<in>s. inner a x > b)"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4619
proof (cases "s = {}")
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4620
  case True
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4621
  then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4622
    apply (rule_tac x="-z" in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4623
    apply (rule_tac x=1 in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4624
    using less_le_trans[OF _ inner_ge_zero[of z]]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4625
    apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4626
    done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4627
next
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4628
  case False
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4629
  obtain y where "y \<in> s" and y: "\<forall>x\<in>s. 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
  4630
    by (metis distance_attains_inf[OF assms(2) False])
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4631
  show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4632
    apply (rule_tac x="y - z" in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4633
    apply (rule_tac x="inner (y - z) z + (norm (y - z))\<^sup>2 / 2" in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4634
    apply rule
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4635
    defer
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4636
    apply rule
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4637
  proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4638
    fix x
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4639
    assume "x \<in> s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4640
    have "\<not> 0 < inner (z - y) (x - y)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4641
      apply (rule notI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4642
      apply (drule closer_point_lemma)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4643
    proof -
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4644
      assume "\<exists>u>0. u \<le> 1 \<and> dist (y + u *\<^sub>R (x - y)) z < dist y z"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4645
      then obtain u where "u > 0" "u \<le> 1" "dist (y + u *\<^sub>R (x - y)) z < dist y z"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4646
        by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4647
      then show False using y[THEN bspec[where x="y + u *\<^sub>R (x - y)"]]
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4648
        using assms(1)[unfolded convex_alt, THEN bspec[where x=y]]
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4649
        using \<open>x\<in>s\<close> \<open>y\<in>s\<close> by (auto simp add: dist_commute algebra_simps)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4650
    qed
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4651
    moreover have "0 < (norm (y - z))\<^sup>2"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4652
      using \<open>y\<in>s\<close> \<open>z\<notin>s\<close> by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4653
    then have "0 < inner (y - z) (y - z)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4654
      unfolding power2_norm_eq_inner by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51524
diff changeset
  4655
    ultimately show "inner (y - z) z + (norm (y - z))\<^sup>2 / 2 < inner (y - z) x"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4656
      unfolding power2_norm_eq_inner and not_less
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4657
      by (auto simp add: field_simps inner_commute inner_diff)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4658
  qed (insert \<open>y\<in>s\<close> \<open>z\<notin>s\<close>, auto)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4659
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4660
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4661
lemma separating_hyperplane_closed_0:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4662
  assumes "convex (s::('a::euclidean_space) set)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4663
    and "closed s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4664
    and "0 \<notin> s"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4665
  shows "\<exists>a b. a \<noteq> 0 \<and> 0 < b \<and> (\<forall>x\<in>s. inner a x > b)"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4666
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
  4667
  case True
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4668
  have "norm ((SOME i. i\<in>Basis)::'a) = 1" "(SOME i. i\<in>Basis) \<noteq> (0::'a)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4669
    defer
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4670
    apply (subst norm_le_zero_iff[symmetric])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4671
    apply (auto simp: SOME_Basis)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4672
    done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4673
  then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4674
    apply (rule_tac x="SOME i. i\<in>Basis" in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4675
    apply (rule_tac x=1 in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4676
    using True using DIM_positive[where 'a='a]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4677
    apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4678
    done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4679
next
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4680
  case False
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4681
  then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4682
    using False using separating_hyperplane_closed_point[OF assms]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4683
    apply (elim exE)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4684
    unfolding inner_zero_right
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4685
    apply (rule_tac x=a in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4686
    apply (rule_tac x=b in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4687
    apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4688
    done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4689
qed
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4690
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4691
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4692
subsubsection \<open>Now set-to-set for closed/compact sets\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4693
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4694
lemma separating_hyperplane_closed_compact:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4695
  fixes s :: "'a::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4696
  assumes "convex s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4697
    and "closed s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4698
    and "convex t"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4699
    and "compact t"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4700
    and "t \<noteq> {}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4701
    and "s \<inter> t = {}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4702
  shows "\<exists>a b. (\<forall>x\<in>s. inner a x < b) \<and> (\<forall>x\<in>t. inner a x > b)"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4703
proof (cases "s = {}")
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4704
  case True
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4705
  obtain b where b: "b > 0" "\<forall>x\<in>t. norm x \<le> b"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4706
    using compact_imp_bounded[OF assms(4)] unfolding bounded_pos by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4707
  obtain z :: 'a where z: "norm z = b + 1"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4708
    using vector_choose_size[of "b + 1"] and b(1) by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4709
  then have "z \<notin> t" using b(2)[THEN bspec[where x=z]] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4710
  then obtain a b where ab: "inner a z < b" "\<forall>x\<in>t. b < inner a x"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4711
    using separating_hyperplane_closed_point[OF assms(3) compact_imp_closed[OF assms(4)], of z]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4712
    by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4713
  then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4714
    using True by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4715
next
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4716
  case False
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4717
  then obtain y where "y \<in> s" by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4718
  obtain a b where "0 < b" "\<forall>x\<in>{x - y |x y. x \<in> s \<and> y \<in> t}. b < inner a x"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4719
    using separating_hyperplane_closed_point[OF convex_differences[OF assms(1,3)], of 0]
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4720
    using closed_compact_differences[OF assms(2,4)]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4721
    using assms(6) by auto blast
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4722
  then have ab: "\<forall>x\<in>s. \<forall>y\<in>t. b + inner a y < inner a x"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4723
    apply -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4724
    apply rule
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4725
    apply rule
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4726
    apply (erule_tac x="x - y" in ballE)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4727
    apply (auto simp add: inner_diff)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4728
    done
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  4729
  define k where "k = (SUP x:t. a \<bullet> x)"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4730
  show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4731
    apply (rule_tac x="-a" in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4732
    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
  4733
    apply (intro conjI ballI)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4734
    unfolding inner_minus_left and neg_less_iff_less
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4735
  proof -
54263
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54258
diff changeset
  4736
    fix x assume "x \<in> t"
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54258
diff changeset
  4737
    then have "inner a x - b / 2 < k"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4738
      unfolding k_def
54263
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54258
diff changeset
  4739
    proof (subst less_cSUP_iff)
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54258
diff changeset
  4740
      show "t \<noteq> {}" by fact
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54258
diff changeset
  4741
      show "bdd_above (op \<bullet> a ` t)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4742
        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
  4743
        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
  4744
    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
  4745
    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
  4746
      by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4747
  next
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4748
    fix x
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4749
    assume "x \<in> s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4750
    then have "k \<le> inner a x - b"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4751
      unfolding k_def
54263
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54258
diff changeset
  4752
      apply (rule_tac cSUP_least)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4753
      using assms(5)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4754
      using ab[THEN bspec[where x=x]]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4755
      apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4756
      done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4757
    then show "k + b / 2 < inner a x"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4758
      using \<open>0 < b\<close> by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4759
  qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4760
qed
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4761
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4762
lemma separating_hyperplane_compact_closed:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4763
  fixes s :: "'a::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4764
  assumes "convex s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4765
    and "compact s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4766
    and "s \<noteq> {}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4767
    and "convex t"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4768
    and "closed t"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4769
    and "s \<inter> t = {}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4770
  shows "\<exists>a b. (\<forall>x\<in>s. inner a x < b) \<and> (\<forall>x\<in>t. inner a x > b)"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4771
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4772
  obtain a b where "(\<forall>x\<in>t. inner a x < b) \<and> (\<forall>x\<in>s. b < inner a x)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4773
    using separating_hyperplane_closed_compact[OF assms(4-5,1-2,3)] and assms(6)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4774
    by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4775
  then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4776
    apply (rule_tac x="-a" in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4777
    apply (rule_tac x="-b" in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4778
    apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4779
    done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4780
qed
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4781
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4782
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4783
subsubsection \<open>General case without assuming closure and getting non-strict separation\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4784
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4785
lemma separating_hyperplane_set_0:
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 36844
diff changeset
  4786
  assumes "convex s" "(0::'a::euclidean_space) \<notin> s"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4787
  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
  4788
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4789
  let ?k = "\<lambda>c. {x::'a. 0 \<le> inner c x}"
62381
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
  4790
  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
  4791
  proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4792
    obtain c where c: "f = ?k ` c" "c \<subseteq> s" "finite c"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4793
      using finite_subset_image[OF as(2,1)] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4794
    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
  4795
      using separating_hyperplane_closed_0[OF convex_convex_hull, of c]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4796
      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
  4797
      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
  4798
      by force
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4799
    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
  4800
      apply (rule_tac x = "inverse(norm a) *\<^sub>R a" in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4801
      using hull_subset[of c convex]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4802
      unfolding subset_eq and inner_scaleR
56536
aefb4a8da31f made mult_nonneg_nonneg a simp rule
nipkow
parents: 56480
diff changeset
  4803
      by (auto simp add: inner_commute del: ballE elim!: ballE)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4804
    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
  4805
      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
  4806
  qed
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
  4807
  have "frontier (cball 0 1) \<inter> (\<Inter>(?k ` s)) \<noteq> {}"
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
  4808
    apply (rule compact_imp_fip)
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
  4809
    apply (rule compact_frontier[OF compact_cball])
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
  4810
    using * closed_halfspace_ge
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
  4811
    by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4812
  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
  4813
    unfolding frontier_cball dist_norm sphere_def by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4814
  then show ?thesis
62381
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
  4815
    by (metis inner_commute mem_Collect_eq norm_eq_zero zero_neq_one)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4816
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4817
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4818
lemma separating_hyperplane_sets:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4819
  fixes s t :: "'a::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4820
  assumes "convex s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4821
    and "convex t"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4822
    and "s \<noteq> {}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4823
    and "t \<noteq> {}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4824
    and "s \<inter> t = {}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4825
  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
  4826
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4827
  from separating_hyperplane_set_0[OF convex_differences[OF assms(2,1)]]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4828
  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"
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
  4829
    using assms(3-5) by fastforce
54263
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54258
diff changeset
  4830
  then have *: "\<And>x y. x \<in> t \<Longrightarrow> y \<in> s \<Longrightarrow> inner a y \<le> inner a x"
33270
paulson
parents: 33175
diff changeset
  4831
    by (force simp add: inner_diff)
54263
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54258
diff changeset
  4832
  then have bdd: "bdd_above ((op \<bullet> a)`s)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4833
    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
  4834
  show ?thesis
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4835
    using \<open>a\<noteq>0\<close>
54263
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54258
diff changeset
  4836
    by (intro exI[of _ a] exI[of _ "SUP x:s. a \<bullet> x"])
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4837
       (auto intro!: cSUP_upper bdd cSUP_least \<open>a \<noteq> 0\<close> \<open>s \<noteq> {}\<close> *)
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4838
qed
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4839
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4840
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4841
subsection \<open>More convexity generalities\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4842
62948
7700f467892b lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 62843
diff changeset
  4843
lemma convex_closure [intro,simp]:
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4844
  fixes s :: "'a::real_normed_vector set"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4845
  assumes "convex s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4846
  shows "convex (closure s)"
53676
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4847
  apply (rule convexI)
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4848
  apply (unfold closure_sequential, elim exE)
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4849
  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
  4850
  apply (rule,rule)
53676
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4851
  apply (rule convexD [OF assms])
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4852
  apply (auto del: tendsto_const intro!: tendsto_intros)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4853
  done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4854
62948
7700f467892b lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 62843
diff changeset
  4855
lemma convex_interior [intro,simp]:
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4856
  fixes s :: "'a::real_normed_vector set"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4857
  assumes "convex s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4858
  shows "convex (interior s)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4859
  unfolding convex_alt Ball_def mem_interior
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4860
  apply (rule,rule,rule,rule,rule,rule)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4861
  apply (elim exE conjE)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4862
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4863
  fix x y u
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4864
  assume u: "0 \<le> u" "u \<le> (1::real)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4865
  fix e d
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4866
  assume ed: "ball x e \<subseteq> s" "ball y d \<subseteq> s" "0<d" "0<e"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4867
  show "\<exists>e>0. ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) e \<subseteq> s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4868
    apply (rule_tac x="min d e" in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4869
    apply rule
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4870
    unfolding subset_eq
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4871
    defer
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4872
    apply rule
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4873
  proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4874
    fix z
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4875
    assume "z \<in> ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) (min d e)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4876
    then have "(1- u) *\<^sub>R (z - u *\<^sub>R (y - x)) + u *\<^sub>R (z + (1 - u) *\<^sub>R (y - x)) \<in> s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4877
      apply (rule_tac assms[unfolded convex_alt, rule_format])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4878
      using ed(1,2) and u
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4879
      unfolding subset_eq mem_ball Ball_def dist_norm
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4880
      apply (auto simp add: algebra_simps)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4881
      done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4882
    then show "z \<in> s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4883
      using u by (auto simp add: algebra_simps)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4884
  qed(insert u ed(3-4), auto)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4885
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4886
34964
4e8be3c04d37 Replaced vec1 and dest_vec1 by abbreviation.
hoelzl
parents: 34915
diff changeset
  4887
lemma convex_hull_eq_empty[simp]: "convex hull s = {} \<longleftrightarrow> s = {}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4888
  using hull_subset[of s convex] convex_hull_empty by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4889
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4890
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4891
subsection \<open>Moving and scaling convex hulls.\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4892
53676
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4893
lemma convex_hull_set_plus:
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4894
  "convex hull (s + t) = convex hull s + convex hull t"
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4895
  unfolding set_plus_image
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4896
  apply (subst convex_hull_linear_image [symmetric])
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4897
  apply (simp add: linear_iff scaleR_right_distrib)
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4898
  apply (simp add: convex_hull_Times)
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4899
  done
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4900
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4901
lemma translation_eq_singleton_plus: "(\<lambda>x. a + x) ` t = {a} + t"
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4902
  unfolding set_plus_def by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4903
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4904
lemma convex_hull_translation:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4905
  "convex hull ((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (convex hull s)"
53676
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4906
  unfolding translation_eq_singleton_plus
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4907
  by (simp only: convex_hull_set_plus convex_hull_singleton)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4908
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4909
lemma convex_hull_scaling:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4910
  "convex hull ((\<lambda>x. c *\<^sub>R x) ` s) = (\<lambda>x. c *\<^sub>R x) ` (convex hull s)"
53676
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4911
  using linear_scaleR by (rule convex_hull_linear_image [symmetric])
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4912
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4913
lemma convex_hull_affinity:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4914
  "convex hull ((\<lambda>x. a + c *\<^sub>R x) ` s) = (\<lambda>x. a + c *\<^sub>R x) ` (convex hull s)"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4915
  by(simp only: image_image[symmetric] convex_hull_scaling convex_hull_translation)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4916
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4917
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4918
subsection \<open>Convexity of cone hulls\<close>
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  4919
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  4920
lemma convex_cone_hull:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4921
  assumes "convex S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4922
  shows "convex (cone hull S)"
53676
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4923
proof (rule convexI)
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4924
  fix x y
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4925
  assume xy: "x \<in> cone hull S" "y \<in> cone hull S"
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4926
  then have "S \<noteq> {}"
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4927
    using cone_hull_empty_iff[of S] by auto
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4928
  fix u v :: real
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4929
  assume uv: "u \<ge> 0" "v \<ge> 0" "u + v = 1"
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4930
  then have *: "u *\<^sub>R x \<in> cone hull S" "v *\<^sub>R y \<in> cone hull S"
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4931
    using cone_cone_hull[of S] xy cone_def[of "cone hull S"] by auto
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4932
  from * obtain cx :: real and xx where x: "u *\<^sub>R x = cx *\<^sub>R xx" "cx \<ge> 0" "xx \<in> S"
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4933
    using cone_hull_expl[of S] by auto
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4934
  from * obtain cy :: real and yy where y: "v *\<^sub>R y = cy *\<^sub>R yy" "cy \<ge> 0" "yy \<in> S"
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4935
    using cone_hull_expl[of S] by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4936
  {
53676
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4937
    assume "cx + cy \<le> 0"
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4938
    then have "u *\<^sub>R x = 0" and "v *\<^sub>R y = 0"
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4939
      using x y by auto
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4940
    then have "u *\<^sub>R x + v *\<^sub>R y = 0"
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4941
      by auto
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4942
    then have "u *\<^sub>R x + v *\<^sub>R y \<in> cone hull S"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4943
      using cone_hull_contains_0[of S] \<open>S \<noteq> {}\<close> by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  4944
  }
53676
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4945
  moreover
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4946
  {
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4947
    assume "cx + cy > 0"
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4948
    then have "(cx / (cx + cy)) *\<^sub>R xx + (cy / (cx + cy)) *\<^sub>R yy \<in> S"
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4949
      using assms mem_convex_alt[of S xx yy cx cy] x y by auto
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4950
    then have "cx *\<^sub>R xx + cy *\<^sub>R yy \<in> cone hull S"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4951
      using mem_cone_hull[of "(cx/(cx+cy)) *\<^sub>R xx + (cy/(cx+cy)) *\<^sub>R yy" S "cx+cy"] \<open>cx+cy>0\<close>
53676
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4952
      by (auto simp add: scaleR_right_distrib)
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4953
    then have "u *\<^sub>R x + v *\<^sub>R y \<in> cone hull S"
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4954
      using x y by auto
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4955
  }
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4956
  moreover have "cx + cy \<le> 0 \<or> cx + cy > 0" by auto
476ef9b468d2 tuned proofs about 'convex'
huffman
parents: 53620
diff changeset
  4957
  ultimately show "u *\<^sub>R x + v *\<^sub>R y \<in> cone hull S" by blast
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  4958
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  4959
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  4960
lemma cone_convex_hull:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4961
  assumes "cone S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4962
  shows "cone (convex hull S)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4963
proof (cases "S = {}")
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4964
  case True
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4965
  then show ?thesis by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4966
next
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4967
  case False
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  4968
  then have *: "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> op *\<^sub>R c ` S = S)"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  4969
    using cone_iff[of S] assms by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4970
  {
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4971
    fix c :: real
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4972
    assume "c > 0"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4973
    then have "op *\<^sub>R c ` (convex hull S) = convex hull (op *\<^sub>R c ` S)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4974
      using convex_hull_scaling[of _ S] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4975
    also have "\<dots> = convex hull S"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4976
      using * \<open>c > 0\<close> by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4977
    finally have "op *\<^sub>R c ` (convex hull S) = convex hull S"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4978
      by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  4979
  }
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4980
  then have "0 \<in> convex hull S" "\<And>c. c > 0 \<Longrightarrow> (op *\<^sub>R c ` (convex hull S)) = (convex hull S)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4981
    using * hull_subset[of S convex] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4982
  then show ?thesis
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4983
    using \<open>S \<noteq> {}\<close> cone_iff[of "convex hull S"] by auto
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4984
qed
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4985
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  4986
subsection \<open>Convex set as intersection of halfspaces\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4987
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4988
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
  4989
  fixes s :: "('a::euclidean_space) set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  4990
  assumes "closed s" "convex s"
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60420
diff changeset
  4991
  shows "s = \<Inter>{h. s \<subseteq> h \<and> (\<exists>a b. h = {x. inner a x \<le> b})}"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4992
  apply (rule set_eqI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4993
  apply rule
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4994
  unfolding Inter_iff Ball_def mem_Collect_eq
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4995
  apply (rule,rule,erule conjE)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4996
proof -
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  4997
  fix x
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  4998
  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
  4999
  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
  5000
    by blast
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5001
  then show "x \<in> s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5002
    apply (rule_tac ccontr)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5003
    apply (drule separating_hyperplane_closed_point[OF assms(2,1)])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5004
    apply (erule exE)+
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5005
    apply (erule_tac x="-a" in allE)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5006
    apply (erule_tac x="-b" in allE)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5007
    apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5008
    done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5009
qed auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5010
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5011
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5012
subsection \<open>Radon's theorem (from Lars Schewe)\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5013
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5014
lemma radon_ex_lemma:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5015
  assumes "finite c" "affine_dependent c"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5016
  shows "\<exists>u. setsum u c = 0 \<and> (\<exists>v\<in>c. u v \<noteq> 0) \<and> setsum (\<lambda>v. u v *\<^sub>R v) c = 0"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5017
proof -
55697
abec82f4e3e9 tuned proofs;
wenzelm
parents: 54780
diff changeset
  5018
  from assms(2)[unfolded affine_dependent_explicit]
abec82f4e3e9 tuned proofs;
wenzelm
parents: 54780
diff changeset
  5019
  obtain s u where
abec82f4e3e9 tuned proofs;
wenzelm
parents: 54780
diff changeset
  5020
      "finite s" "s \<subseteq> c" "setsum u s = 0" "\<exists>v\<in>s. u v \<noteq> 0" "(\<Sum>v\<in>s. u v *\<^sub>R v) = 0"
abec82f4e3e9 tuned proofs;
wenzelm
parents: 54780
diff changeset
  5021
    by blast
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5022
  then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5023
    apply (rule_tac x="\<lambda>v. if v\<in>s then u v else 0" in exI)
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  5024
    unfolding if_smult scaleR_zero_left and setsum.inter_restrict[OF assms(1), symmetric]
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5025
    apply (auto simp add: Int_absorb1)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5026
    done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5027
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5028
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5029
lemma radon_s_lemma:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5030
  assumes "finite s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5031
    and "setsum f s = (0::real)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5032
  shows "setsum f {x\<in>s. 0 < f x} = - setsum f {x\<in>s. f x < 0}"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5033
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5034
  have *: "\<And>x. (if f x < 0 then f x else 0) + (if 0 < f x then f x else 0) = f x"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5035
    by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5036
  show ?thesis
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
  5037
    unfolding add_eq_0_iff[symmetric] and setsum.inter_filter[OF assms(1)]
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  5038
      and setsum.distrib[symmetric] and *
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5039
    using assms(2)
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
  5040
    by assumption
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5041
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5042
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5043
lemma radon_v_lemma:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5044
  assumes "finite s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5045
    and "setsum f s = 0"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5046
    and "\<forall>x. g x = (0::real) \<longrightarrow> f x = (0::'a::euclidean_space)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5047
  shows "(setsum f {x\<in>s. 0 < g x}) = - setsum f {x\<in>s. g x < 0}"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5048
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5049
  have *: "\<And>x. (if 0 < g x then f x else 0) + (if g x < 0 then f x else 0) = f x"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5050
    using assms(3) by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5051
  show ?thesis
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  5052
    unfolding eq_neg_iff_add_eq_0 and setsum.inter_filter[OF assms(1)]
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  5053
      and setsum.distrib[symmetric] and *
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5054
    using assms(2)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5055
    apply assumption
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5056
    done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5057
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5058
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5059
lemma radon_partition:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5060
  assumes "finite c" "affine_dependent c"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5061
  shows "\<exists>m p. m \<inter> p = {} \<and> m \<union> p = c \<and> (convex hull m) \<inter> (convex hull p) \<noteq> {}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5062
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5063
  obtain u v where uv: "setsum u c = 0" "v\<in>c" "u v \<noteq> 0"  "(\<Sum>v\<in>c. u v *\<^sub>R v) = 0"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5064
    using radon_ex_lemma[OF assms] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5065
  have fin: "finite {x \<in> c. 0 < u x}" "finite {x \<in> c. 0 > u x}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5066
    using assms(1) by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  5067
  define z  where "z = inverse (setsum u {x\<in>c. u x > 0}) *\<^sub>R setsum (\<lambda>x. u x *\<^sub>R x) {x\<in>c. u x > 0}"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5068
  have "setsum u {x \<in> c. 0 < u x} \<noteq> 0"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5069
  proof (cases "u v \<ge> 0")
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5070
    case False
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5071
    then have "u v < 0" by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5072
    then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5073
    proof (cases "\<exists>w\<in>{x \<in> c. 0 < u x}. u w > 0")
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5074
      case True
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5075
      then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5076
        using setsum_nonneg_eq_0_iff[of _ u, OF fin(1)] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5077
    next
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5078
      case False
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5079
      then have "setsum u c \<le> setsum (\<lambda>x. if x=v then u v else 0) c"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5080
        apply (rule_tac setsum_mono)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5081
        apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5082
        done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5083
      then show ?thesis
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5084
        unfolding setsum.delta[OF assms(1)] using uv(2) and \<open>u v < 0\<close> and uv(1) by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5085
    qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5086
  qed (insert setsum_nonneg_eq_0_iff[of _ u, OF fin(1)] uv(2-3), auto)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5087
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5088
  then have *: "setsum u {x\<in>c. u x > 0} > 0"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5089
    unfolding less_le
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5090
    apply (rule_tac conjI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5091
    apply (rule_tac setsum_nonneg)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5092
    apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5093
    done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5094
  moreover have "setsum u ({x \<in> c. 0 < u x} \<union> {x \<in> c. u x < 0}) = setsum u c"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5095
    "(\<Sum>x\<in>{x \<in> c. 0 < u x} \<union> {x \<in> c. u x < 0}. u x *\<^sub>R x) = (\<Sum>x\<in>c. u x *\<^sub>R x)"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5096
    using assms(1)
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  5097
    apply (rule_tac[!] setsum.mono_neutral_left)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5098
    apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5099
    done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5100
  then have "setsum u {x \<in> c. 0 < u x} = - setsum u {x \<in> c. 0 > u x}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5101
    "(\<Sum>x\<in>{x \<in> c. 0 < u x}. u x *\<^sub>R x) = - (\<Sum>x\<in>{x \<in> c. 0 > u x}. u x *\<^sub>R x)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5102
    unfolding eq_neg_iff_add_eq_0
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5103
    using uv(1,4)
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  5104
    by (auto simp add: setsum.union_inter_neutral[OF fin, symmetric])
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  5105
  moreover have "\<forall>x\<in>{v \<in> c. u v < 0}. 0 \<le> inverse (setsum u {x \<in> c. 0 < u x}) * - u x"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5106
    apply rule
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5107
    apply (rule mult_nonneg_nonneg)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5108
    using *
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5109
    apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5110
    done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5111
  ultimately have "z \<in> convex hull {v \<in> c. u v \<le> 0}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5112
    unfolding convex_hull_explicit mem_Collect_eq
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5113
    apply (rule_tac x="{v \<in> c. u v < 0}" in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5114
    apply (rule_tac x="\<lambda>y. inverse (setsum u {x\<in>c. u x > 0}) * - u y" in exI)
49530
wenzelm
parents: 49529
diff changeset
  5115
    using assms(1) unfolding scaleR_scaleR[symmetric] scaleR_right.setsum [symmetric] and z_def
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5116
    apply (auto simp add: setsum_negf setsum_right_distrib[symmetric])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5117
    done
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  5118
  moreover have "\<forall>x\<in>{v \<in> c. 0 < u v}. 0 \<le> inverse (setsum u {x \<in> c. 0 < u x}) * u x"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5119
    apply rule
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5120
    apply (rule mult_nonneg_nonneg)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5121
    using *
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5122
    apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5123
    done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5124
  then have "z \<in> convex hull {v \<in> c. u v > 0}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5125
    unfolding convex_hull_explicit mem_Collect_eq
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5126
    apply (rule_tac x="{v \<in> c. 0 < u v}" in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5127
    apply (rule_tac x="\<lambda>y. inverse (setsum u {x\<in>c. u x > 0}) * u y" in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5128
    using assms(1)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5129
    unfolding scaleR_scaleR[symmetric] scaleR_right.setsum [symmetric] and z_def
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5130
    using *
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5131
    apply (auto simp add: setsum_negf setsum_right_distrib[symmetric])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5132
    done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5133
  ultimately show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5134
    apply (rule_tac x="{v\<in>c. u v \<le> 0}" in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5135
    apply (rule_tac x="{v\<in>c. u v > 0}" in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5136
    apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5137
    done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5138
qed
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5139
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5140
lemma radon:
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5141
  assumes "affine_dependent c"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5142
  obtains m p where "m \<subseteq> c" "p \<subseteq> c" "m \<inter> p = {}" "(convex hull m) \<inter> (convex hull p) \<noteq> {}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5143
proof -
55697
abec82f4e3e9 tuned proofs;
wenzelm
parents: 54780
diff changeset
  5144
  from assms[unfolded affine_dependent_explicit]
abec82f4e3e9 tuned proofs;
wenzelm
parents: 54780
diff changeset
  5145
  obtain s u where
abec82f4e3e9 tuned proofs;
wenzelm
parents: 54780
diff changeset
  5146
      "finite s" "s \<subseteq> c" "setsum u s = 0" "\<exists>v\<in>s. u v \<noteq> 0" "(\<Sum>v\<in>s. u v *\<^sub>R v) = 0"
abec82f4e3e9 tuned proofs;
wenzelm
parents: 54780
diff changeset
  5147
    by blast
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5148
  then have *: "finite s" "affine_dependent s" and s: "s \<subseteq> c"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5149
    unfolding affine_dependent_explicit by auto
55697
abec82f4e3e9 tuned proofs;
wenzelm
parents: 54780
diff changeset
  5150
  from radon_partition[OF *]
abec82f4e3e9 tuned proofs;
wenzelm
parents: 54780
diff changeset
  5151
  obtain m p where "m \<inter> p = {}" "m \<union> p = s" "convex hull m \<inter> convex hull p \<noteq> {}"
abec82f4e3e9 tuned proofs;
wenzelm
parents: 54780
diff changeset
  5152
    by blast
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5153
  then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5154
    apply (rule_tac that[of p m])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5155
    using s
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5156
    apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5157
    done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5158
qed
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5159
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5160
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5161
subsection \<open>Helly's theorem\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5162
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5163
lemma helly_induct:
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5164
  fixes f :: "'a::euclidean_space set set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5165
  assumes "card f = n"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5166
    and "n \<ge> DIM('a) + 1"
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60420
diff changeset
  5167
    and "\<forall>s\<in>f. convex s" "\<forall>t\<subseteq>f. card t = DIM('a) + 1 \<longrightarrow> \<Inter>t \<noteq> {}"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5168
  shows "\<Inter>f \<noteq> {}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5169
  using assms
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5170
proof (induct n arbitrary: f)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5171
  case 0
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5172
  then show ?case by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5173
next
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5174
  case (Suc n)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5175
  have "finite f"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5176
    using \<open>card f = Suc n\<close> by (auto intro: card_ge_0_finite)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5177
  show "\<Inter>f \<noteq> {}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5178
    apply (cases "n = DIM('a)")
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5179
    apply (rule Suc(5)[rule_format])
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5180
    unfolding \<open>card f = Suc n\<close>
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5181
  proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5182
    assume ng: "n \<noteq> DIM('a)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5183
    then have "\<exists>X. \<forall>s\<in>f. X s \<in> \<Inter>(f - {s})"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5184
      apply (rule_tac bchoice)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5185
      unfolding ex_in_conv
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5186
      apply (rule, rule Suc(1)[rule_format])
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5187
      unfolding card_Diff_singleton_if[OF \<open>finite f\<close>] \<open>card f = Suc n\<close>
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5188
      defer
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5189
      defer
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5190
      apply (rule Suc(4)[rule_format])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5191
      defer
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5192
      apply (rule Suc(5)[rule_format])
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5193
      using Suc(3) \<open>finite f\<close>
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5194
      apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5195
      done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5196
    then obtain X where X: "\<forall>s\<in>f. X s \<in> \<Inter>(f - {s})" by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5197
    show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5198
    proof (cases "inj_on X f")
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5199
      case False
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5200
      then obtain s t where st: "s\<noteq>t" "s\<in>f" "t\<in>f" "X s = X t"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5201
        unfolding inj_on_def by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5202
      then have *: "\<Inter>f = \<Inter>(f - {s}) \<inter> \<Inter>(f - {t})" by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5203
      show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5204
        unfolding *
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5205
        unfolding ex_in_conv[symmetric]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5206
        apply (rule_tac x="X s" in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5207
        apply rule
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5208
        apply (rule X[rule_format])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5209
        using X st
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5210
        apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5211
        done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5212
    next
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5213
      case True
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5214
      then obtain m p where mp: "m \<inter> p = {}" "m \<union> p = X ` f" "convex hull m \<inter> convex hull p \<noteq> {}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5215
        using radon_partition[of "X ` f"] and affine_dependent_biggerset[of "X ` f"]
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5216
        unfolding card_image[OF True] and \<open>card f = Suc n\<close>
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5217
        using Suc(3) \<open>finite f\<close> and ng
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5218
        by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5219
      have "m \<subseteq> X ` f" "p \<subseteq> X ` f"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5220
        using mp(2) by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5221
      then obtain g h where gh:"m = X ` g" "p = X ` h" "g \<subseteq> f" "h \<subseteq> f"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5222
        unfolding subset_image_iff by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5223
      then have "f \<union> (g \<union> h) = f" by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5224
      then have f: "f = g \<union> h"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5225
        using inj_on_Un_image_eq_iff[of X f "g \<union> h"] and True
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5226
        unfolding mp(2)[unfolded image_Un[symmetric] gh]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5227
        by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5228
      have *: "g \<inter> h = {}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5229
        using mp(1)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5230
        unfolding gh
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5231
        using inj_on_image_Int[OF True gh(3,4)]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5232
        by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5233
      have "convex hull (X ` h) \<subseteq> \<Inter>g" "convex hull (X ` g) \<subseteq> \<Inter>h"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5234
        apply (rule_tac [!] hull_minimal)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5235
        using Suc gh(3-4)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5236
        unfolding subset_eq
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5237
        apply (rule_tac [2] convex_Inter, rule_tac [4] convex_Inter)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5238
        apply rule
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5239
        prefer 3
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5240
        apply rule
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5241
      proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5242
        fix x
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5243
        assume "x \<in> X ` g"
55697
abec82f4e3e9 tuned proofs;
wenzelm
parents: 54780
diff changeset
  5244
        then obtain y where "y \<in> g" "x = X y"
abec82f4e3e9 tuned proofs;
wenzelm
parents: 54780
diff changeset
  5245
          unfolding image_iff ..
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5246
        then show "x \<in> \<Inter>h"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5247
          using X[THEN bspec[where x=y]] using * f by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5248
      next
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5249
        fix x
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5250
        assume "x \<in> X ` h"
55697
abec82f4e3e9 tuned proofs;
wenzelm
parents: 54780
diff changeset
  5251
        then obtain y where "y \<in> h" "x = X y"
abec82f4e3e9 tuned proofs;
wenzelm
parents: 54780
diff changeset
  5252
          unfolding image_iff ..
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5253
        then show "x \<in> \<Inter>g"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5254
          using X[THEN bspec[where x=y]] using * f by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5255
      qed auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5256
      then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5257
        unfolding f using mp(3)[unfolded gh] by blast
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5258
    qed
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5259
  qed auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5260
qed
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5261
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5262
lemma helly:
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5263
  fixes f :: "'a::euclidean_space set set"
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 36844
diff changeset
  5264
  assumes "card f \<ge> DIM('a) + 1" "\<forall>s\<in>f. convex s"
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60420
diff changeset
  5265
    and "\<forall>t\<subseteq>f. card t = DIM('a) + 1 \<longrightarrow> \<Inter>t \<noteq> {}"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5266
  shows "\<Inter>f \<noteq> {}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5267
  apply (rule helly_induct)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5268
  using assms
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5269
  apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5270
  done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5271
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5272
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5273
subsection \<open>Homeomorphism of all convex compact sets with nonempty interior\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5274
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5275
lemma compact_frontier_line_lemma:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5276
  fixes s :: "'a::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5277
  assumes "compact s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5278
    and "0 \<in> s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5279
    and "x \<noteq> 0"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5280
  obtains u where "0 \<le> u" and "(u *\<^sub>R x) \<in> frontier s" "\<forall>v>u. (v *\<^sub>R x) \<notin> s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5281
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5282
  obtain b where b: "b > 0" "\<forall>x\<in>s. norm x \<le> b"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5283
    using compact_imp_bounded[OF assms(1), unfolded bounded_pos] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5284
  let ?A = "{y. \<exists>u. 0 \<le> u \<and> u \<le> b / norm(x) \<and> (y = u *\<^sub>R x)}"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5285
  have A: "?A = (\<lambda>u. u *\<^sub>R x) ` {0 .. b / norm x}"
36431
340755027840 move definitions and theorems for type real^1 to separate theory file
huffman
parents: 36365
diff changeset
  5286
    by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5287
  have *: "\<And>x A B. x\<in>A \<Longrightarrow> x\<in>B \<Longrightarrow> A\<inter>B \<noteq> {}" by blast
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5288
  have "compact ?A"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5289
    unfolding A
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5290
    apply (rule compact_continuous_image)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5291
    apply (rule continuous_at_imp_continuous_on)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5292
    apply rule
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5293
    apply (intro continuous_intros)
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  5294
    apply (rule compact_Icc)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5295
    done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5296
  moreover have "{y. \<exists>u\<ge>0. u \<le> b / norm x \<and> y = u *\<^sub>R x} \<inter> s \<noteq> {}"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5297
    apply(rule *[OF _ assms(2)])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5298
    unfolding mem_Collect_eq
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5299
    using \<open>b > 0\<close> assms(3)
56571
f4635657d66f added divide_nonneg_nonneg and co; made it a simp rule
hoelzl
parents: 56544
diff changeset
  5300
    apply auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5301
    done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5302
  ultimately obtain u y where obt: "u\<ge>0" "u \<le> b / norm x" "y = u *\<^sub>R x"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5303
    "y \<in> ?A" "y \<in> s" "\<forall>z\<in>?A \<inter> s. dist 0 z \<le> dist 0 y"
62843
313d3b697c9a Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents: 62626
diff changeset
  5304
    using distance_attains_sup[OF compact_Int[OF _ assms(1), of ?A], of 0] by blast
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5305
  have "norm x > 0"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5306
    using assms(3)[unfolded zero_less_norm_iff[symmetric]] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5307
  {
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5308
    fix v
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5309
    assume as: "v > u" "v *\<^sub>R x \<in> s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5310
    then have "v \<le> b / norm x"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5311
      using b(2)[rule_format, OF as(2)]
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5312
      using \<open>u\<ge>0\<close>
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5313
      unfolding pos_le_divide_eq[OF \<open>norm x > 0\<close>]
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5314
      by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5315
    then have "norm (v *\<^sub>R x) \<le> norm y"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5316
      apply (rule_tac obt(6)[rule_format, unfolded dist_0_norm])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5317
      apply (rule IntI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5318
      defer
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5319
      apply (rule as(2))
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5320
      unfolding mem_Collect_eq
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5321
      apply (rule_tac x=v in exI)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5322
      using as(1) \<open>u\<ge>0\<close>
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5323
      apply (auto simp add: field_simps)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5324
      done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5325
    then have False
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5326
      unfolding obt(3) using \<open>u\<ge>0\<close> \<open>norm x > 0\<close> \<open>v > u\<close>
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5327
      by (auto simp add:field_simps)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5328
  } note u_max = this
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5329
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5330
  have "u *\<^sub>R x \<in> frontier s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5331
    unfolding frontier_straddle
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5332
    apply (rule,rule,rule)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5333
    apply (rule_tac x="u *\<^sub>R x" in bexI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5334
    unfolding obt(3)[symmetric]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5335
    prefer 3
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5336
    apply (rule_tac x="(u + (e / 2) / norm x) *\<^sub>R x" in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5337
    apply (rule, rule)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5338
  proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5339
    fix e
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5340
    assume "e > 0" and as: "(u + e / 2 / norm x) *\<^sub>R x \<in> s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5341
    then have "u + e / 2 / norm x > u"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5342
      using \<open>norm x > 0\<close> by (auto simp del:zero_less_norm_iff)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5343
    then show False using u_max[OF _ as] by auto
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5344
  qed (insert \<open>y\<in>s\<close>, auto simp add: dist_norm scaleR_left_distrib obt(3))
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5345
  then show ?thesis by(metis that[of u] u_max obt(1))
36071
c8ae8e56d42e tuned many proofs a bit
nipkow
parents: 35577
diff changeset
  5346
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5347
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5348
lemma starlike_compact_projective:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5349
  assumes "compact s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5350
    and "cball (0::'a::euclidean_space) 1 \<subseteq> s "
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5351
    and "\<forall>x\<in>s. \<forall>u. 0 \<le> u \<and> u < 1 \<longrightarrow> u *\<^sub>R x \<in> s - frontier s"
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 36844
diff changeset
  5352
  shows "s homeomorphic (cball (0::'a::euclidean_space) 1)"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5353
proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5354
  have fs: "frontier s \<subseteq> s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5355
    apply (rule frontier_subset_closed)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5356
    using compact_imp_closed[OF assms(1)]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5357
    apply simp
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5358
    done
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  5359
  define pi where [abs_def]: "pi x = inverse (norm x) *\<^sub>R x" for x :: 'a
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5360
  have "0 \<notin> frontier s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5361
    unfolding frontier_straddle
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5362
    apply (rule notI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5363
    apply (erule_tac x=1 in allE)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5364
    using assms(2)[unfolded subset_eq Ball_def mem_cball]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5365
    apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5366
    done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5367
  have injpi: "\<And>x y. pi x = pi y \<and> norm x = norm y \<longleftrightarrow> x = y"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5368
    unfolding pi_def by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5369
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5370
  have contpi: "continuous_on (UNIV - {0}) pi"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5371
    apply (rule continuous_at_imp_continuous_on)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5372
    apply rule unfolding pi_def
44647
e4de7750cdeb modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents: 44629
diff changeset
  5373
    apply (intro continuous_intros)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5374
    apply simp
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5375
    done
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  5376
  define sphere :: "'a set" where "sphere = {x. norm x = 1}"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5377
  have pi: "\<And>x. x \<noteq> 0 \<Longrightarrow> pi x \<in> sphere" "\<And>x u. u>0 \<Longrightarrow> pi (u *\<^sub>R x) = pi x"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5378
    unfolding pi_def sphere_def by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5379
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5380
  have "0 \<in> s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5381
    using assms(2) and centre_in_cball[of 0 1] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5382
  have front_smul: "\<forall>x\<in>frontier s. \<forall>u\<ge>0. u *\<^sub>R x \<in> s \<longleftrightarrow> u \<le> 1"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5383
  proof (rule,rule,rule)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5384
    fix x and u :: real
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5385
    assume x: "x \<in> frontier s" and "0 \<le> u"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5386
    then have "x \<noteq> 0"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5387
      using \<open>0 \<notin> frontier s\<close> by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5388
    obtain v where v: "0 \<le> v" "v *\<^sub>R x \<in> frontier s" "\<forall>w>v. w *\<^sub>R x \<notin> s"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5389
      using compact_frontier_line_lemma[OF assms(1) \<open>0\<in>s\<close> \<open>x\<noteq>0\<close>] by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5390
    have "v = 1"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5391
      apply (rule ccontr)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5392
      unfolding neq_iff
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5393
      apply (erule disjE)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5394
    proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5395
      assume "v < 1"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5396
      then show False
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62398
diff changeset
  5397
        using v(3)[THEN spec[where x=1]] using x fs by (simp add: pth_1 subset_iff)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5398
    next
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5399
      assume "v > 1"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5400
      then show False
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5401
        using assms(3)[THEN bspec[where x="v *\<^sub>R x"], THEN spec[where x="inverse v"]]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5402
        using v and x and fs
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5403
        unfolding inverse_less_1_iff by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5404
    qed
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5405
    show "u *\<^sub>R x \<in> s \<longleftrightarrow> u \<le> 1"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5406
      apply rule
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5407
      using v(3)[unfolded \<open>v=1\<close>, THEN spec[where x=u]]
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5408
    proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5409
      assume "u \<le> 1"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5410
      then show "u *\<^sub>R x \<in> s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5411
      apply (cases "u = 1")
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5412
        using assms(3)[THEN bspec[where x=x], THEN spec[where x=u]]
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5413
        using \<open>0\<le>u\<close> and x and fs
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62398
diff changeset
  5414
        by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5415
    qed auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5416
  qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5417
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5418
  have "\<exists>surf. homeomorphism (frontier s) sphere pi surf"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5419
    apply (rule homeomorphism_compact)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5420
    apply (rule compact_frontier[OF assms(1)])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5421
    apply (rule continuous_on_subset[OF contpi])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5422
    defer
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5423
    apply (rule set_eqI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5424
    apply rule
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5425
    unfolding inj_on_def
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5426
    prefer 3
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5427
    apply(rule,rule,rule)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5428
  proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5429
    fix x
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5430
    assume "x \<in> pi ` frontier s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5431
    then obtain y where "y \<in> frontier s" "x = pi y" by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5432
    then show "x \<in> sphere"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5433
      using pi(1)[of y] and \<open>0 \<notin> frontier s\<close> by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5434
  next
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5435
    fix x
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5436
    assume "x \<in> sphere"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5437
    then have "norm x = 1" "x \<noteq> 0"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5438
      unfolding sphere_def by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5439
    then obtain u where "0 \<le> u" "u *\<^sub>R x \<in> frontier s" "\<forall>v>u. v *\<^sub>R x \<notin> s"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5440
      using compact_frontier_line_lemma[OF assms(1) \<open>0\<in>s\<close>, of x] by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5441
    then show "x \<in> pi ` frontier s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5442
      unfolding image_iff le_less pi_def
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5443
      apply (rule_tac x="u *\<^sub>R x" in bexI)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5444
      using \<open>norm x = 1\<close> \<open>0 \<notin> frontier s\<close>
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5445
      apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5446
      done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5447
  next
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5448
    fix x y
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5449
    assume as: "x \<in> frontier s" "y \<in> frontier s" "pi x = pi y"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5450
    then have xys: "x \<in> s" "y \<in> s"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5451
      using fs by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5452
    from as(1,2) have nor: "norm x \<noteq> 0" "norm y \<noteq> 0"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5453
      using \<open>0\<notin>frontier s\<close> by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5454
    from nor have x: "x = norm x *\<^sub>R ((inverse (norm y)) *\<^sub>R y)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5455
      unfolding as(3)[unfolded pi_def, symmetric] by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5456
    from nor have y: "y = norm y *\<^sub>R ((inverse (norm x)) *\<^sub>R x)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5457
      unfolding as(3)[unfolded pi_def] by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5458
    have "0 \<le> norm y * inverse (norm x)" and "0 \<le> norm x * inverse (norm y)"
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5459
      using nor
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5460
      apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5461
      done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5462
    then have "norm x = norm y"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5463
      apply -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5464
      apply (rule ccontr)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5465
      unfolding neq_iff
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5466
      using x y and front_smul[THEN bspec, OF as(1), THEN spec[where x="norm y * (inverse (norm x))"]]
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5467
      using front_smul[THEN bspec, OF as(2), THEN spec[where x="norm x * (inverse (norm y))"]]
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5468
      using xys nor
56571
f4635657d66f added divide_nonneg_nonneg and co; made it a simp rule
hoelzl
parents: 56544
diff changeset
  5469
      apply (auto simp add: field_simps)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5470
      done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5471
    then show "x = y"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5472
      apply (subst injpi[symmetric])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5473
      using as(3)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5474
      apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5475
      done
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5476
  qed (insert \<open>0 \<notin> frontier s\<close>, auto)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5477
  then obtain surf where
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5478
    surf: "\<forall>x\<in>frontier s. surf (pi x) = x"  "pi ` frontier s = sphere" "continuous_on (frontier s) pi"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5479
    "\<forall>y\<in>sphere. pi (surf y) = y" "surf ` sphere = frontier s" "continuous_on sphere surf"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5480
    unfolding homeomorphism_def by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5481
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5482
  have cont_surfpi: "continuous_on (UNIV -  {0}) (surf \<circ> pi)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5483
    apply (rule continuous_on_compose)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5484
    apply (rule contpi)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5485
    apply (rule continuous_on_subset[of sphere])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5486
    apply (rule surf(6))
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5487
    using pi(1)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5488
    apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5489
    done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5490
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5491
  {
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5492
    fix x
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5493
    assume as: "x \<in> cball (0::'a) 1"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5494
    have "norm x *\<^sub>R surf (pi x) \<in> s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5495
    proof (cases "x=0 \<or> norm x = 1")
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5496
      case False
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5497
      then have "pi x \<in> sphere" "norm x < 1"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5498
        using pi(1)[of x] as by(auto simp add: dist_norm)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5499
      then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5500
        apply (rule_tac assms(3)[rule_format, THEN DiffD1])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5501
        apply (rule_tac fs[unfolded subset_eq, rule_format])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5502
        unfolding surf(5)[symmetric]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5503
        apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5504
        done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5505
    next
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5506
      case True
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5507
      then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5508
        apply rule
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5509
        defer
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5510
        unfolding pi_def
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5511
        apply (rule fs[unfolded subset_eq, rule_format])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5512
        unfolding surf(5)[unfolded sphere_def, symmetric]
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5513
        using \<open>0\<in>s\<close>
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5514
        apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5515
        done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5516
    qed
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5517
  } note hom = this
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5518
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5519
  {
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5520
    fix x
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5521
    assume "x \<in> s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5522
    then have "x \<in> (\<lambda>x. norm x *\<^sub>R surf (pi x)) ` cball 0 1"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5523
    proof (cases "x = 0")
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5524
      case True
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5525
      show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5526
        unfolding image_iff True
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5527
        apply (rule_tac x=0 in bexI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5528
        apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5529
        done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5530
    next
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5531
      let ?a = "inverse (norm (surf (pi x)))"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5532
      case False
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5533
      then have invn: "inverse (norm x) \<noteq> 0" by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5534
      from False have pix: "pi x\<in>sphere" using pi(1) by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5535
      then have "pi (surf (pi x)) = pi x"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5536
        apply (rule_tac surf(4)[rule_format])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5537
        apply assumption
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5538
        done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5539
      then have **: "norm x *\<^sub>R (?a *\<^sub>R surf (pi x)) = x"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5540
        apply (rule_tac scaleR_left_imp_eq[OF invn])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5541
        unfolding pi_def
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5542
        using invn
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5543
        apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5544
        done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5545
      then have *: "?a * norm x > 0" and "?a > 0" "?a \<noteq> 0"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5546
        using surf(5) \<open>0\<notin>frontier s\<close>
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5547
        apply -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5548
        apply (rule mult_pos_pos)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5549
        using False[unfolded zero_less_norm_iff[symmetric]]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5550
        apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5551
        done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5552
      have "norm (surf (pi x)) \<noteq> 0"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5553
        using ** False by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5554
      then have "norm x = norm ((?a * norm x) *\<^sub>R surf (pi x))"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5555
        unfolding norm_scaleR abs_mult abs_norm_cancel abs_of_pos[OF \<open>?a > 0\<close>] by auto
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  5556
      moreover have "pi x = pi ((inverse (norm (surf (pi x))) * norm x) *\<^sub>R surf (pi x))"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5557
        unfolding pi(2)[OF *] surf(4)[rule_format, OF pix] ..
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5558
      moreover have "surf (pi x) \<in> frontier s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5559
        using surf(5) pix by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5560
      then have "dist 0 (inverse (norm (surf (pi x))) *\<^sub>R x) \<le> 1"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5561
        unfolding dist_norm
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5562
        using ** and *
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5563
        using front_smul[THEN bspec[where x="surf (pi x)"], THEN spec[where x="norm x * ?a"]]
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5564
        using False \<open>x\<in>s\<close>
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5565
        by (auto simp add: field_simps)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5566
      ultimately show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5567
        unfolding image_iff
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5568
        apply (rule_tac x="inverse (norm (surf(pi x))) *\<^sub>R x" in bexI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5569
        apply (subst injpi[symmetric])
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5570
        unfolding abs_mult abs_norm_cancel abs_of_pos[OF \<open>?a > 0\<close>]
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5571
        unfolding pi(2)[OF \<open>?a > 0\<close>]
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5572
        apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5573
        done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5574
    qed
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5575
  } note hom2 = this
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5576
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5577
  show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5578
    apply (subst homeomorphic_sym)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5579
    apply (rule homeomorphic_compact[where f="\<lambda>x. norm x *\<^sub>R surf (pi x)"])
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5580
    apply (rule compact_cball)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5581
    defer
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5582
    apply (rule set_eqI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5583
    apply rule
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5584
    apply (erule imageE)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5585
    apply (drule hom)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5586
    prefer 4
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5587
    apply (rule continuous_at_imp_continuous_on)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5588
    apply rule
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5589
    apply (rule_tac [3] hom2)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5590
  proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5591
    fix x :: 'a
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5592
    assume as: "x \<in> cball 0 1"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5593
    then show "continuous (at x) (\<lambda>x. norm x *\<^sub>R surf (pi x))"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5594
    proof (cases "x = 0")
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5595
      case False
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5596
      then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5597
        apply (intro continuous_intros)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5598
        using cont_surfpi
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5599
        unfolding continuous_on_eq_continuous_at[OF open_delete[OF open_UNIV]] o_def
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5600
        apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5601
        done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5602
    next
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5603
      case True
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5604
      obtain B where B: "\<forall>x\<in>s. norm x \<le> B"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5605
        using compact_imp_bounded[OF assms(1)] unfolding bounded_iff by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5606
      then have "B > 0"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5607
        using assms(2)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5608
        unfolding subset_eq
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5609
        apply (erule_tac x="SOME i. i\<in>Basis" in ballE)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5610
        defer
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5611
        apply (erule_tac x="SOME i. i\<in>Basis" in ballE)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5612
        unfolding Ball_def mem_cball dist_norm
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5613
        using DIM_positive[where 'a='a]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5614
        apply (auto simp: SOME_Basis)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5615
        done
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5616
      show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5617
        unfolding True continuous_at Lim_at
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5618
        apply(rule,rule)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5619
        apply(rule_tac x="e / B" in exI)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5620
        apply rule
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5621
        apply (rule divide_pos_pos)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5622
        prefer 3
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5623
        apply(rule,rule,erule conjE)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5624
        unfolding norm_zero scaleR_zero_left dist_norm diff_0_right norm_scaleR abs_norm_cancel
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5625
      proof -
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5626
        fix e and x :: 'a
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5627
        assume as: "norm x < e / B" "0 < norm x" "e > 0"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5628
        then have "surf (pi x) \<in> frontier s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5629
          using pi(1)[of x] unfolding surf(5)[symmetric] by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5630
        then have "norm (surf (pi x)) \<le> B"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5631
          using B fs by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5632
        then have "norm x * norm (surf (pi x)) \<le> norm x * B"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5633
          using as(2) by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5634
        also have "\<dots> < e / B * B"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5635
          apply (rule mult_strict_right_mono)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5636
          using as(1) \<open>B>0\<close>
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5637
          apply auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5638
          done
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5639
        also have "\<dots> = e" using \<open>B > 0\<close> by auto
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5640
        finally show "norm x * norm (surf (pi x)) < e" .
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5641
      qed (insert \<open>B>0\<close>, auto)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5642
    qed
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5643
  next
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5644
    {
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5645
      fix x
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5646
      assume as: "surf (pi x) = 0"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5647
      have "x = 0"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5648
      proof (rule ccontr)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5649
        assume "x \<noteq> 0"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5650
        then have "pi x \<in> sphere"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5651
          using pi(1) by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5652
        then have "surf (pi x) \<in> frontier s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5653
          using surf(5) by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5654
        then show False
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5655
          using \<open>0\<notin>frontier s\<close> unfolding as by simp
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5656
      qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5657
    } note surf_0 = this
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5658
    show "inj_on (\<lambda>x. norm x *\<^sub>R surf (pi x)) (cball 0 1)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5659
      unfolding inj_on_def
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5660
    proof (rule,rule,rule)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5661
      fix x y
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5662
      assume as: "x \<in> cball 0 1" "y \<in> cball 0 1" "norm x *\<^sub>R surf (pi x) = norm y *\<^sub>R surf (pi y)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5663
      then show "x = y"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5664
      proof (cases "x=0 \<or> y=0")
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5665
        case True
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5666
        then show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5667
          using as by (auto elim: surf_0)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5668
      next
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5669
        case False
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5670
        then have "pi (surf (pi x)) = pi (surf (pi y))"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5671
          using as(3)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5672
          using pi(2)[of "norm x" "surf (pi x)"] pi(2)[of "norm y" "surf (pi y)"]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5673
          by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5674
        moreover have "pi x \<in> sphere" "pi y \<in> sphere"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5675
          using pi(1) False by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5676
        ultimately have *: "pi x = pi y"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5677
          using surf(4)[THEN bspec[where x="pi x"]] surf(4)[THEN bspec[where x="pi y"]]
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5678
          by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5679
        moreover have "norm x = norm y"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5680
          using as(3)[unfolded *] using False
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5681
          by (auto dest:surf_0)
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5682
        ultimately show ?thesis
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5683
          using injpi by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5684
      qed
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5685
    qed
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5686
  qed auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5687
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5688
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44467
diff changeset
  5689
lemma homeomorphic_convex_compact_lemma:
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5690
  fixes s :: "'a::euclidean_space set"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5691
  assumes "convex s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5692
    and "compact s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5693
    and "cball 0 1 \<subseteq> s"
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 36844
diff changeset
  5694
  shows "s homeomorphic (cball (0::'a) 1)"
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44467
diff changeset
  5695
proof (rule starlike_compact_projective[OF assms(2-3)], clarify)
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5696
  fix x u
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5697
  assume "x \<in> s" and "0 \<le> u" and "u < (1::real)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5698
  have "open (ball (u *\<^sub>R x) (1 - u))"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5699
    by (rule open_ball)
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44467
diff changeset
  5700
  moreover have "u *\<^sub>R x \<in> ball (u *\<^sub>R x) (1 - u)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5701
    unfolding centre_in_ball using \<open>u < 1\<close> by simp
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44467
diff changeset
  5702
  moreover have "ball (u *\<^sub>R x) (1 - u) \<subseteq> s"
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44467
diff changeset
  5703
  proof
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5704
    fix y
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5705
    assume "y \<in> ball (u *\<^sub>R x) (1 - u)"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5706
    then have "dist (u *\<^sub>R x) y < 1 - u"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5707
      unfolding mem_ball .
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5708
    with \<open>u < 1\<close> have "inverse (1 - u) *\<^sub>R (y - u *\<^sub>R x) \<in> cball 0 1"
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44467
diff changeset
  5709
      by (simp add: dist_norm inverse_eq_divide norm_minus_commute)
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44467
diff changeset
  5710
    with assms(3) have "inverse (1 - u) *\<^sub>R (y - u *\<^sub>R x) \<in> s" ..
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44467
diff changeset
  5711
    with assms(1) have "(1 - u) *\<^sub>R ((y - u *\<^sub>R x) /\<^sub>R (1 - u)) + u *\<^sub>R x \<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
  5712
      using \<open>x \<in> s\<close> \<open>0 \<le> u\<close> \<open>u < 1\<close> [THEN less_imp_le] by (rule convexD_alt)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5713
    then show "y \<in> s" using \<open>u < 1\<close>
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5714
      by simp
44519
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44467
diff changeset
  5715
  qed
ea85d78a994e simplify definition of 'interior';
huffman
parents: 44467
diff changeset
  5716
  ultimately have "u *\<^sub>R x \<in> interior s" ..
53347
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5717
  then show "u *\<^sub>R x \<in> s - frontier s"
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5718
    using frontier_def and interior_subset by auto
547610c26257 tuned proofs;
wenzelm
parents: 53339
diff changeset
  5719
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5720
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5721
lemma homeomorphic_convex_compact_cball:
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5722
  fixes e :: real
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5723
    and s :: "'a::euclidean_space set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5724
  assumes "convex s"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5725
    and "compact s"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5726
    and "interior s \<noteq> {}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5727
    and "e > 0"
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 36844
diff changeset
  5728
  shows "s homeomorphic (cball (b::'a) e)"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5729
proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5730
  obtain a where "a \<in> interior s"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5731
    using assms(3) by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5732
  then obtain d where "d > 0" and d: "cball a d \<subseteq> s"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5733
    unfolding mem_interior_cball by auto
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 36844
diff changeset
  5734
  let ?d = "inverse d" and ?n = "0::'a"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5735
  have "cball ?n 1 \<subseteq> (\<lambda>x. inverse d *\<^sub>R (x - a)) ` s"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5736
    apply rule
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5737
    apply (rule_tac x="d *\<^sub>R x + a" in image_eqI)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5738
    defer
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5739
    apply (rule d[unfolded subset_eq, rule_format])
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5740
    using \<open>d > 0\<close>
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5741
    unfolding mem_cball dist_norm
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5742
    apply (auto simp add: mult_right_le_one_le)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5743
    done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5744
  then have "(\<lambda>x. inverse d *\<^sub>R (x - a)) ` s homeomorphic cball ?n 1"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5745
    using homeomorphic_convex_compact_lemma[of "(\<lambda>x. ?d *\<^sub>R -a + ?d *\<^sub>R x) ` s",
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5746
      OF convex_affinity compact_affinity]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5747
    using assms(1,2)
57865
dcfb33c26f50 tuned proofs -- fewer warnings;
wenzelm
parents: 57512
diff changeset
  5748
    by (auto simp add: scaleR_right_diff_distrib)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5749
  then show ?thesis
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5750
    apply (rule_tac homeomorphic_trans[OF _ homeomorphic_balls(2)[of 1 _ ?n]])
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5751
    apply (rule homeomorphic_trans[OF homeomorphic_affinity[of "?d" s "?d *\<^sub>R -a"]])
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5752
    using \<open>d>0\<close> \<open>e>0\<close>
57865
dcfb33c26f50 tuned proofs -- fewer warnings;
wenzelm
parents: 57512
diff changeset
  5753
    apply (auto simp add: scaleR_right_diff_distrib)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5754
    done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5755
qed
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5756
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5757
lemma homeomorphic_convex_compact:
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5758
  fixes s :: "'a::euclidean_space set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5759
    and t :: "'a set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5760
  assumes "convex s" "compact s" "interior s \<noteq> {}"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5761
    and "convex t" "compact t" "interior t \<noteq> {}"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5762
  shows "s homeomorphic t"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5763
  using assms
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5764
  by (meson zero_less_one homeomorphic_trans homeomorphic_convex_compact_cball homeomorphic_sym)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5765
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5766
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5767
subsection \<open>Epigraphs of convex functions\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5768
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5769
definition "epigraph s (f :: _ \<Rightarrow> real) = {xy. fst xy \<in> s \<and> f (fst xy) \<le> snd xy}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5770
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5771
lemma mem_epigraph: "(x, y) \<in> epigraph s f \<longleftrightarrow> x \<in> s \<and> f x \<le> y"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5772
  unfolding epigraph_def by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5773
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5774
lemma convex_epigraph: "convex (epigraph s f) \<longleftrightarrow> convex_on s f \<and> convex s"
36338
7808fbc9c3b4 generalize more constants and lemmas
huffman
parents: 36337
diff changeset
  5775
  unfolding convex_def convex_on_def
7808fbc9c3b4 generalize more constants and lemmas
huffman
parents: 36337
diff changeset
  5776
  unfolding Ball_def split_paired_All epigraph_def
7808fbc9c3b4 generalize more constants and lemmas
huffman
parents: 36337
diff changeset
  5777
  unfolding mem_Collect_eq fst_conv snd_conv fst_add snd_add fst_scaleR snd_scaleR Ball_def[symmetric]
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5778
  apply safe
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5779
  defer
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5780
  apply (erule_tac x=x in allE)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5781
  apply (erule_tac x="f x" in allE)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5782
  apply safe
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5783
  apply (erule_tac x=xa in allE)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5784
  apply (erule_tac x="f xa" in allE)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5785
  prefer 3
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5786
  apply (rule_tac y="u * f a + v * f aa" in order_trans)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5787
  defer
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5788
  apply (auto intro!:mult_left_mono add_mono)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5789
  done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5790
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5791
lemma convex_epigraphI: "convex_on s f \<Longrightarrow> convex s \<Longrightarrow> convex (epigraph s f)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5792
  unfolding convex_epigraph by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5793
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5794
lemma convex_epigraph_convex: "convex s \<Longrightarrow> convex_on s f \<longleftrightarrow> convex(epigraph s f)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5795
  by (simp add: convex_epigraph)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5796
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5797
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5798
subsubsection \<open>Use this to derive general bound property of convex function\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5799
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5800
lemma convex_on:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5801
  assumes "convex s"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5802
  shows "convex_on s f \<longleftrightarrow>
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5803
    (\<forall>k u x. (\<forall>i\<in>{1..k::nat}. 0 \<le> u i \<and> x i \<in> s) \<and> setsum u {1..k} = 1 \<longrightarrow>
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5804
      f (setsum (\<lambda>i. u i *\<^sub>R x i) {1..k} ) \<le> setsum (\<lambda>i. u i * f(x i)) {1..k})"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5805
  unfolding convex_epigraph_convex[OF assms] convex epigraph_def Ball_def mem_Collect_eq
36338
7808fbc9c3b4 generalize more constants and lemmas
huffman
parents: 36337
diff changeset
  5806
  unfolding fst_setsum snd_setsum fst_scaleR snd_scaleR
7808fbc9c3b4 generalize more constants and lemmas
huffman
parents: 36337
diff changeset
  5807
  apply safe
7808fbc9c3b4 generalize more constants and lemmas
huffman
parents: 36337
diff changeset
  5808
  apply (drule_tac x=k in spec)
7808fbc9c3b4 generalize more constants and lemmas
huffman
parents: 36337
diff changeset
  5809
  apply (drule_tac x=u in spec)
7808fbc9c3b4 generalize more constants and lemmas
huffman
parents: 36337
diff changeset
  5810
  apply (drule_tac x="\<lambda>i. (x i, f (x i))" in spec)
7808fbc9c3b4 generalize more constants and lemmas
huffman
parents: 36337
diff changeset
  5811
  apply simp
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5812
  using assms[unfolded convex]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5813
  apply simp
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5814
  apply (rule_tac y="\<Sum>i = 1..k. u i * f (fst (x i))" in order_trans)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5815
  defer
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5816
  apply (rule setsum_mono)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5817
  apply (erule_tac x=i in allE)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5818
  unfolding real_scaleR_def
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5819
  apply (rule mult_left_mono)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5820
  using assms[unfolded convex]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5821
  apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5822
  done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5823
36338
7808fbc9c3b4 generalize more constants and lemmas
huffman
parents: 36337
diff changeset
  5824
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5825
subsection \<open>Convexity of general and special intervals\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5826
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5827
lemma is_interval_convex:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5828
  fixes s :: "'a::euclidean_space set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5829
  assumes "is_interval s"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5830
  shows "convex s"
37732
6432bf0d7191 generalize type of is_interval to class euclidean_space
huffman
parents: 37673
diff changeset
  5831
proof (rule convexI)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5832
  fix x y and u v :: real
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5833
  assume as: "x \<in> s" "y \<in> s" "0 \<le> u" "0 \<le> v" "u + v = 1"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5834
  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
  5835
    by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5836
  {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5837
    fix a b
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5838
    assume "\<not> b \<le> u * a + v * b"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5839
    then have "u * a < (1 - v) * b"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5840
      unfolding not_le using as(4) by (auto simp add: field_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5841
    then have "a < b"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5842
      unfolding * using as(4) *(2)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5843
      apply (rule_tac mult_left_less_imp_less[of "1 - v"])
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5844
      apply (auto simp add: field_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5845
      done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5846
    then have "a \<le> u * a + v * b"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5847
      unfolding * using as(4)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5848
      by (auto simp add: field_simps intro!:mult_right_mono)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5849
  }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5850
  moreover
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5851
  {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5852
    fix a b
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5853
    assume "\<not> u * a + v * b \<le> a"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5854
    then have "v * b > (1 - u) * a"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5855
      unfolding not_le using as(4) by (auto simp add: field_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5856
    then have "a < b"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5857
      unfolding * using as(4)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5858
      apply (rule_tac mult_left_less_imp_less)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5859
      apply (auto simp add: field_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5860
      done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5861
    then have "u * a + v * b \<le> b"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5862
      unfolding **
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5863
      using **(2) as(3)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5864
      by (auto simp add: field_simps intro!:mult_right_mono)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5865
  }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5866
  ultimately show "u *\<^sub>R x + v *\<^sub>R y \<in> s"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5867
    apply -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5868
    apply (rule assms[unfolded is_interval_def, rule_format, OF as(1,2)])
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5869
    using as(3-) DIM_positive[where 'a='a]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5870
    apply (auto simp: inner_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5871
    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
  5872
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5873
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5874
lemma is_interval_connected:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5875
  fixes s :: "'a::euclidean_space set"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5876
  shows "is_interval s \<Longrightarrow> connected s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5877
  using is_interval_convex convex_connected by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5878
62618
f7f2467ab854 Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  5879
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
  5880
  apply (rule_tac[!] is_interval_convex)+
56189
c4daa97ac57a removed dependencies on theory Ordered_Euclidean_Space
immler
parents: 56188
diff changeset
  5881
  using is_interval_box is_interval_cbox
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5882
  apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5883
  done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5884
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61762
diff changeset
  5885
subsection \<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
  5886
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5887
lemma is_interval_1:
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  5888
  "is_interval (s::real set) \<longleftrightarrow> (\<forall>a\<in>s. \<forall>b\<in>s. \<forall> x. a \<le> x \<and> x \<le> b \<longrightarrow> x \<in> s)"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  5889
  unfolding is_interval_def by auto
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  5890
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5891
lemma is_interval_connected_1:
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5892
  fixes s :: "real set"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5893
  shows "is_interval s \<longleftrightarrow> connected s"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5894
  apply rule
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5895
  apply (rule is_interval_connected, assumption)
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5896
  unfolding is_interval_1
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5897
  apply rule
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5898
  apply rule
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5899
  apply rule
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5900
  apply rule
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5901
  apply (erule conjE)
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5902
  apply (rule ccontr)
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5903
proof -
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5904
  fix a b x
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5905
  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
  5906
  then have *: "a < x" "x < b"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5907
    unfolding not_le [symmetric] by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5908
  let ?halfl = "{..<x} "
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5909
  let ?halfr = "{x<..}"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5910
  {
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5911
    fix y
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5912
    assume "y \<in> s"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5913
    with \<open>x \<notin> s\<close> have "x \<noteq> y" by auto
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5914
    then have "y \<in> ?halfr \<union> ?halfl" by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5915
  }
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5916
  moreover have "a \<in> ?halfl" "b \<in> ?halfr" using * by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5917
  then have "?halfl \<inter> s \<noteq> {}" "?halfr \<inter> s \<noteq> {}"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5918
    using as(2-3) by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5919
  ultimately show False
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5920
    apply (rule_tac notE[OF as(1)[unfolded connected_def]])
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5921
    apply (rule_tac x = ?halfl in exI)
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5922
    apply (rule_tac x = ?halfr in exI)
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5923
    apply rule
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5924
    apply (rule open_lessThan)
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5925
    apply rule
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5926
    apply (rule open_greaterThan)
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5927
    apply auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5928
    done
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5929
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5930
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5931
lemma is_interval_convex_1:
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5932
  fixes s :: "real set"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5933
  shows "is_interval s \<longleftrightarrow> convex s"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5934
  by (metis is_interval_convex convex_connected is_interval_connected_1)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5935
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  5936
lemma connected_convex_1:
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5937
  fixes s :: "real set"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5938
  shows "connected s \<longleftrightarrow> convex s"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  5939
  by (metis is_interval_convex convex_connected is_interval_connected_1)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5940
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  5941
lemma connected_convex_1_gen:
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  5942
  fixes s :: "'a :: euclidean_space set"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  5943
  assumes "DIM('a) = 1"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  5944
  shows "connected s \<longleftrightarrow> convex s"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  5945
proof -
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  5946
  obtain f:: "'a \<Rightarrow> real" where linf: "linear f" and "inj f"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  5947
    using subspace_isomorphism [where 'a = 'a and 'b = real]
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  5948
    by (metis DIM_real dim_UNIV subspace_UNIV assms)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  5949
  then have "f -` (f ` s) = s"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  5950
    by (simp add: inj_vimage_image_eq)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  5951
  then show ?thesis
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  5952
    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
  5953
qed
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5954
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  5955
subsection \<open>Another intermediate value theorem formulation\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5956
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5957
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
  5958
  fixes f :: "real \<Rightarrow> 'a::euclidean_space"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5959
  assumes "a \<le> b"
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  5960
    and "continuous_on {a..b} f"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5961
    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
  5962
  shows "\<exists>x\<in>{a..b}. (f x)\<bullet>k = y"
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  5963
proof -
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  5964
  have "f a \<in> f ` cbox a b" "f b \<in> f ` cbox a b"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5965
    apply (rule_tac[!] imageI)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5966
    using assms(1)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5967
    apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5968
    done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5969
  then show ?thesis
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  5970
    using connected_ivt_component[of "f ` cbox a b" "f a" "f b" k y]
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  5971
    by (simp add: Topology_Euclidean_Space.connected_continuous_image assms)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5972
qed
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5973
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5974
lemma ivt_increasing_component_1:
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5975
  fixes f :: "real \<Rightarrow> 'a::euclidean_space"
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  5976
  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
  5977
    f a\<bullet>k \<le> y \<Longrightarrow> y \<le> f b\<bullet>k \<Longrightarrow> \<exists>x\<in>{a..b}. (f x)\<bullet>k = y"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5978
  by (rule ivt_increasing_component_on_1) (auto simp add: continuous_at_imp_continuous_on)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5979
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5980
lemma ivt_decreasing_component_on_1:
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5981
  fixes f :: "real \<Rightarrow> 'a::euclidean_space"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5982
  assumes "a \<le> b"
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  5983
    and "continuous_on {a..b} f"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5984
    and "(f b)\<bullet>k \<le> y"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5985
    and "y \<le> (f a)\<bullet>k"
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  5986
  shows "\<exists>x\<in>{a..b}. (f x)\<bullet>k = y"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5987
  apply (subst neg_equal_iff_equal[symmetric])
44531
1d477a2b1572 replace some continuous_on lemmas with more general versions
huffman
parents: 44525
diff changeset
  5988
  using ivt_increasing_component_on_1[of a b "\<lambda>x. - f x" k "- y"]
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5989
  using assms using continuous_on_minus
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5990
  apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5991
  done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5992
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5993
lemma ivt_decreasing_component_1:
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5994
  fixes f :: "real \<Rightarrow> 'a::euclidean_space"
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  5995
  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
  5996
    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
  5997
  by (rule ivt_decreasing_component_on_1) (auto simp: continuous_at_imp_continuous_on)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  5998
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  5999
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  6000
subsection \<open>A bound within a convex hull, and so an interval\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6001
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6002
lemma convex_on_convex_hull_bound:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6003
  assumes "convex_on (convex hull s) f"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6004
    and "\<forall>x\<in>s. f x \<le> b"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6005
  shows "\<forall>x\<in> convex hull s. f x \<le> b"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6006
proof
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6007
  fix x
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6008
  assume "x \<in> convex hull s"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6009
  then obtain k u v where
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6010
    obt: "\<forall>i\<in>{1..k::nat}. 0 \<le> u i \<and> v i \<in> s" "setsum u {1..k} = 1" "(\<Sum>i = 1..k. u i *\<^sub>R v i) = x"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6011
    unfolding convex_hull_indexed mem_Collect_eq by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6012
  have "(\<Sum>i = 1..k. u i * f (v i)) \<le> b"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6013
    using setsum_mono[of "{1..k}" "\<lambda>i. u i * f (v i)" "\<lambda>i. u i * b"]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6014
    unfolding setsum_left_distrib[symmetric] obt(2) mult_1
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6015
    apply (drule_tac meta_mp)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6016
    apply (rule mult_left_mono)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6017
    using assms(2) obt(1)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6018
    apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6019
    done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6020
  then show "f x \<le> b"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6021
    using assms(1)[unfolded convex_on[OF convex_convex_hull], rule_format, of k u v]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6022
    unfolding obt(2-3)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6023
    using obt(1) and hull_subset[unfolded subset_eq, rule_format, of _ s]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6024
    by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6025
qed
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6026
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6027
lemma inner_setsum_Basis[simp]: "i \<in> Basis \<Longrightarrow> (\<Sum>Basis) \<bullet> i = 1"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  6028
  by (simp add: inner_setsum_left setsum.If_cases inner_Basis)
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6029
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6030
lemma convex_set_plus:
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6031
  assumes "convex s" and "convex t" shows "convex (s + t)"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6032
proof -
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6033
  have "convex {x + y |x y. x \<in> s \<and> y \<in> t}"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6034
    using assms by (rule convex_sums)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6035
  moreover have "{x + y |x y. x \<in> s \<and> y \<in> t} = s + t"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6036
    unfolding set_plus_def by auto
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6037
  finally show "convex (s + t)" .
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6038
qed
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6039
55929
91f245c23bc5 remove lemmas in favor of more general ones: convex(_hull)_set_{plus,setsum}
huffman
parents: 55928
diff changeset
  6040
lemma convex_set_setsum:
91f245c23bc5 remove lemmas in favor of more general ones: convex(_hull)_set_{plus,setsum}
huffman
parents: 55928
diff changeset
  6041
  assumes "\<And>i. i \<in> A \<Longrightarrow> convex (B i)"
91f245c23bc5 remove lemmas in favor of more general ones: convex(_hull)_set_{plus,setsum}
huffman
parents: 55928
diff changeset
  6042
  shows "convex (\<Sum>i\<in>A. B i)"
91f245c23bc5 remove lemmas in favor of more general ones: convex(_hull)_set_{plus,setsum}
huffman
parents: 55928
diff changeset
  6043
proof (cases "finite A")
91f245c23bc5 remove lemmas in favor of more general ones: convex(_hull)_set_{plus,setsum}
huffman
parents: 55928
diff changeset
  6044
  case True then show ?thesis using assms
91f245c23bc5 remove lemmas in favor of more general ones: convex(_hull)_set_{plus,setsum}
huffman
parents: 55928
diff changeset
  6045
    by induct (auto simp: convex_set_plus)
91f245c23bc5 remove lemmas in favor of more general ones: convex(_hull)_set_{plus,setsum}
huffman
parents: 55928
diff changeset
  6046
qed auto
91f245c23bc5 remove lemmas in favor of more general ones: convex(_hull)_set_{plus,setsum}
huffman
parents: 55928
diff changeset
  6047
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6048
lemma finite_set_setsum:
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6049
  assumes "finite A" and "\<forall>i\<in>A. finite (B i)" shows "finite (\<Sum>i\<in>A. B i)"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6050
  using assms by (induct set: finite, simp, simp add: finite_set_plus)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6051
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6052
lemma set_setsum_eq:
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6053
  "finite A \<Longrightarrow> (\<Sum>i\<in>A. B i) = {\<Sum>i\<in>A. f i |f. \<forall>i\<in>A. f i \<in> B i}"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6054
  apply (induct set: finite)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6055
  apply simp
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6056
  apply simp
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6057
  apply (safe elim!: set_plus_elim)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6058
  apply (rule_tac x="fun_upd f x a" in exI)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6059
  apply simp
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6060
  apply (rule_tac f="\<lambda>x. a + x" in arg_cong)
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  6061
  apply (rule setsum.cong [OF refl])
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6062
  apply clarsimp
57865
dcfb33c26f50 tuned proofs -- fewer warnings;
wenzelm
parents: 57512
diff changeset
  6063
  apply fast
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6064
  done
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6065
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6066
lemma box_eq_set_setsum_Basis:
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6067
  shows "{x. \<forall>i\<in>Basis. x\<bullet>i \<in> B i} = (\<Sum>i\<in>Basis. image (\<lambda>x. x *\<^sub>R i) (B i))"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6068
  apply (subst set_setsum_eq [OF finite_Basis])
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6069
  apply safe
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6070
  apply (fast intro: euclidean_representation [symmetric])
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6071
  apply (subst inner_setsum_left)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6072
  apply (subgoal_tac "(\<Sum>x\<in>Basis. f x \<bullet> i) = f i \<bullet> i")
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6073
  apply (drule (1) bspec)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6074
  apply clarsimp
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  6075
  apply (frule setsum.remove [OF finite_Basis])
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6076
  apply (erule trans)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6077
  apply simp
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  6078
  apply (rule setsum.neutral)
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6079
  apply clarsimp
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6080
  apply (frule_tac x=i in bspec, assumption)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6081
  apply (drule_tac x=x in bspec, assumption)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6082
  apply clarsimp
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6083
  apply (cut_tac u=x and v=i in inner_Basis, assumption+)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6084
  apply (rule ccontr)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6085
  apply simp
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6086
  done
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6087
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6088
lemma convex_hull_set_setsum:
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6089
  "convex hull (\<Sum>i\<in>A. B i) = (\<Sum>i\<in>A. convex hull (B i))"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6090
proof (cases "finite A")
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6091
  assume "finite A" then show ?thesis
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6092
    by (induct set: finite, simp, simp add: convex_hull_set_plus)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6093
qed simp
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6094
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  6095
lemma convex_hull_eq_real_cbox:
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6096
  fixes x y :: real assumes "x \<le> y"
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  6097
  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
  6098
proof (rule hull_unique)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  6099
  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
  6100
  show "convex (cbox x y)"
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  6101
    by (rule convex_box)
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6102
next
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6103
  fix s assume "{x, y} \<subseteq> s" and "convex s"
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  6104
  then show "cbox x y \<subseteq> s"
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6105
    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
  6106
    by - (clarify, simp (no_asm_use), fast)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6107
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
  6108
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6109
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
  6110
  "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
  6111
  (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
  6112
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
  6113
  have One[simp]: "\<And>i. i \<in> Basis \<Longrightarrow> One \<bullet> i = 1"
62091
wenzelm
parents: 61952
diff changeset
  6114
    by (simp add: inner_setsum_left setsum.If_cases inner_Basis)
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  6115
  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
  6116
    by (auto simp: cbox_def)
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  6117
  also have "\<dots> = (\<Sum>i\<in>Basis. (\<lambda>x. x *\<^sub>R i) ` cbox 0 1)"
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6118
    by (simp only: box_eq_set_setsum_Basis)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6119
  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
  6120
    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
  6121
  also have "\<dots> = (\<Sum>i\<in>Basis. convex hull ((\<lambda>x. x *\<^sub>R i) ` {0, 1}))"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6122
    by (simp only: convex_hull_linear_image linear_scaleR_left)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6123
  also have "\<dots> = convex hull (\<Sum>i\<in>Basis. (\<lambda>x. x *\<^sub>R i) ` {0, 1})"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6124
    by (simp only: convex_hull_set_setsum)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6125
  also have "\<dots> = convex hull {x. \<forall>i\<in>Basis. x\<bullet>i \<in> {0, 1}}"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6126
    by (simp only: box_eq_set_setsum_Basis)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6127
  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
  6128
    by simp
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6129
  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
  6130
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6131
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  6132
text \<open>And this is a finite set of vertices.\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6133
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
  6134
lemma unit_cube_convex_hull:
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  6135
  obtains s :: "'a::euclidean_space set"
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  6136
    where "finite s" and "cbox 0 (\<Sum>Basis) = convex hull s"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6137
  apply (rule that[of "{x::'a. \<forall>i\<in>Basis. x\<bullet>i=0 \<or> x\<bullet>i=1}"])
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6138
  apply (rule finite_subset[of _ "(\<lambda>s. (\<Sum>i\<in>Basis. (if i\<in>s then 1 else 0) *\<^sub>R i)::'a) ` Pow Basis"])
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6139
  prefer 3
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6140
  apply (rule unit_interval_convex_hull)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6141
  apply rule
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6142
  unfolding mem_Collect_eq
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6143
proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6144
  fix x :: 'a
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6145
  assume as: "\<forall>i\<in>Basis. x \<bullet> i = 0 \<or> x \<bullet> i = 1"
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
  6146
  show "x \<in> (\<lambda>s. \<Sum>i\<in>Basis. (if i\<in>s then 1 else 0) *\<^sub>R i) ` Pow Basis"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6147
    apply (rule image_eqI[where x="{i. i\<in>Basis \<and> x\<bullet>i = 1}"])
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6148
    using as
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6149
    apply (subst euclidean_eq_iff)
57865
dcfb33c26f50 tuned proofs -- fewer warnings;
wenzelm
parents: 57512
diff changeset
  6150
    apply auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6151
    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
  6152
qed auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6153
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  6154
text \<open>Hence any cube (could do any nonempty interval).\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6155
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6156
lemma cube_convex_hull:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6157
  assumes "d > 0"
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  6158
  obtains s :: "'a::euclidean_space set" where
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  6159
    "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
  6160
proof -
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
  6161
  let ?d = "(\<Sum>i\<in>Basis. d*\<^sub>Ri)::'a"
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  6162
  have *: "cbox (x - ?d) (x + ?d) = (\<lambda>y. x - ?d + (2 * d) *\<^sub>R y) ` cbox 0 (\<Sum>Basis)"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6163
    apply (rule set_eqI, rule)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6164
    unfolding image_iff
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6165
    defer
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6166
    apply (erule bexE)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6167
  proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6168
    fix y
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  6169
    assume as: "y\<in>cbox (x - ?d) (x + ?d)"
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  6170
    then have "inverse (2 * d) *\<^sub>R (y - (x - ?d)) \<in> cbox 0 (\<Sum>Basis)"
58776
95e58e04e534 use NO_MATCH-simproc for distribution rules in field_simps, otherwise field_simps on '(a / (c + d)) * (e + f)' can be non-terminating
hoelzl
parents: 57865
diff changeset
  6171
      using assms by (simp add: mem_box field_simps inner_simps)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  6172
    with \<open>0 < d\<close> show "\<exists>z\<in>cbox 0 (\<Sum>Basis). y = x - ?d + (2 * d) *\<^sub>R z"
58776
95e58e04e534 use NO_MATCH-simproc for distribution rules in field_simps, otherwise field_simps on '(a / (c + d)) * (e + f)' can be non-terminating
hoelzl
parents: 57865
diff changeset
  6173
      by (intro bexI[of _ "inverse (2 * d) *\<^sub>R (y - (x - ?d))"]) auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6174
  next
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6175
    fix y z
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  6176
    assume as: "z\<in>cbox 0 (\<Sum>Basis)" "y = x - ?d + (2*d) *\<^sub>R z"
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
  6177
    have "\<And>i. i\<in>Basis \<Longrightarrow> 0 \<le> d * (z \<bullet> i) \<and> d * (z \<bullet> i) \<le> d"
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  6178
      using assms as(1)[unfolded mem_box]
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6179
      apply (erule_tac x=i in ballE)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6180
      apply rule
56536
aefb4a8da31f made mult_nonneg_nonneg a simp rule
nipkow
parents: 56480
diff changeset
  6181
      prefer 2
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6182
      apply (rule mult_right_le_one_le)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6183
      using assms
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6184
      apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6185
      done
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  6186
    then show "y \<in> cbox (x - ?d) (x + ?d)"
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  6187
      unfolding as(2) mem_box
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6188
      apply -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6189
      apply rule
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  6190
      using as(1)[unfolded mem_box]
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6191
      apply (erule_tac x=i in ballE)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6192
      using assms
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6193
      apply (auto simp: inner_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6194
      done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6195
  qed
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  6196
  obtain s where "finite s" "cbox 0 (\<Sum>Basis::'a) = convex hull s"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6197
    using unit_cube_convex_hull by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6198
  then show ?thesis
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6199
    apply (rule_tac that[of "(\<lambda>y. x - ?d + (2 * d) *\<^sub>R y)` s"])
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6200
    unfolding * and convex_hull_affinity
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6201
    apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6202
    done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6203
qed
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6204
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6205
subsubsection\<open>Representation of any interval as a finite convex hull\<close>
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6206
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6207
lemma image_stretch_interval:
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6208
  "(\<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
  6209
  (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
  6210
    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
  6211
     (\<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
  6212
proof cases
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6213
  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
  6214
  show ?thesis
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6215
    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
  6216
    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
  6217
    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
  6218
  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
  6219
    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
  6220
    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
  6221
      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
  6222
    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
  6223
        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
  6224
    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
  6225
      case True
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6226
      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
  6227
    next
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6228
      case False
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6229
      then have *: "\<And>a b. a = m i * b \<longleftrightarrow> b = a / m i"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6230
        by (auto simp add: field_simps)
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6231
      from False have
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6232
          "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
  6233
          "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
  6234
        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
  6235
      with False show ?thesis using a_le_b
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6236
        unfolding * by (auto simp add: le_divide_eq divide_le_eq ac_simps)
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6237
    qed
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6238
  qed
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6239
qed simp
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6240
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6241
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
  6242
  "\<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
  6243
  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
  6244
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6245
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
  6246
  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
  6247
  using box_ne_empty [of "a+c" "b+c"]  box_ne_empty [of a b]
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6248
  by (auto simp add: inner_left_distrib add.commute)
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6249
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6250
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
  6251
  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
  6252
  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
  6253
    shows "cbox a b =
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6254
           op + a ` (\<lambda>x. \<Sum>k\<in>Basis. ((b \<bullet> k - a \<bullet> k) * (x \<bullet> k)) *\<^sub>R k) ` cbox 0 One"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6255
using assms
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6256
apply (simp add: box_ne_empty image_stretch_interval cbox_translation [symmetric])
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6257
apply (simp add: min_def max_def algebra_simps setsum_subtractf euclidean_representation)
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6258
done
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6259
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6260
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
  6261
  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
  6262
  obtains s where "finite s" "cbox a b = convex hull s"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6263
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
  6264
  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
  6265
    by blast
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6266
next
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6267
  case False
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6268
  obtain s::"'a set" where "finite s" and eq: "cbox 0 One = convex hull s"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6269
    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
  6270
  have lin: "linear (\<lambda>x. \<Sum>k\<in>Basis. ((b \<bullet> k - a \<bullet> k) * (x \<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
  6271
    by (rule linear_compose_setsum) (auto simp: algebra_simps linearI)
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6272
  have "finite (op + a ` (\<lambda>x. \<Sum>k\<in>Basis. ((b \<bullet> k - a \<bullet> k) * (x \<bullet> k)) *\<^sub>R k) ` s)"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6273
    by (rule finite_imageI \<open>finite s\<close>)+
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6274
  then show ?thesis
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6275
    apply (rule that)
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6276
    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
  6277
    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
  6278
    done
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6279
qed
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6280
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6281
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  6282
subsection \<open>Bounded convex function on open set is continuous\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6283
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6284
lemma convex_on_bounded_continuous:
36338
7808fbc9c3b4 generalize more constants and lemmas
huffman
parents: 36337
diff changeset
  6285
  fixes s :: "('a::real_normed_vector) set"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6286
  assumes "open s"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6287
    and "convex_on s f"
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61880
diff changeset
  6288
    and "\<forall>x\<in>s. \<bar>f x\<bar> \<le> b"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6289
  shows "continuous_on s f"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6290
  apply (rule continuous_at_imp_continuous_on)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6291
  unfolding continuous_at_real_range
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6292
proof (rule,rule,rule)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6293
  fix x and e :: real
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6294
  assume "x \<in> s" "e > 0"
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  6295
  define B where "B = \<bar>b\<bar> + 1"
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61880
diff changeset
  6296
  have B: "0 < B" "\<And>x. x\<in>s \<Longrightarrow> \<bar>f x\<bar> \<le> B"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6297
    unfolding B_def
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6298
    defer
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6299
    apply (drule assms(3)[rule_format])
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6300
    apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6301
    done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6302
  obtain k where "k > 0" and k: "cball x k \<subseteq> s"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6303
    using assms(1)[unfolded open_contains_cball, THEN bspec[where x=x]]
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  6304
    using \<open>x\<in>s\<close> by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6305
  show "\<exists>d>0. \<forall>x'. norm (x' - x) < d \<longrightarrow> \<bar>f x' - f x\<bar> < e"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6306
    apply (rule_tac x="min (k / 2) (e / (2 * B) * k)" in exI)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6307
    apply rule
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6308
    defer
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6309
  proof (rule, rule)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6310
    fix y
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6311
    assume as: "norm (y - x) < min (k / 2) (e / (2 * B) * k)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6312
    show "\<bar>f y - f x\<bar> < e"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6313
    proof (cases "y = x")
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6314
      case False
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  6315
      define t where "t = k / norm (y - x)"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6316
      have "2 < t" "0<t"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  6317
        unfolding t_def using as False and \<open>k>0\<close>
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6318
        by (auto simp add:field_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6319
      have "y \<in> s"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6320
        apply (rule k[unfolded subset_eq,rule_format])
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6321
        unfolding mem_cball dist_norm
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6322
        apply (rule order_trans[of _ "2 * norm (x - y)"])
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6323
        using as
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6324
        by (auto simp add: field_simps norm_minus_commute)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6325
      {
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  6326
        define w where "w = x + t *\<^sub>R (y - x)"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6327
        have "w \<in> s"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6328
          unfolding w_def
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6329
          apply (rule k[unfolded subset_eq,rule_format])
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6330
          unfolding mem_cball dist_norm
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6331
          unfolding t_def
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  6332
          using \<open>k>0\<close>
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6333
          apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6334
          done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6335
        have "(1 / t) *\<^sub>R x + - x + ((t - 1) / t) *\<^sub>R x = (1 / t - 1 + (t - 1) / t) *\<^sub>R x"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6336
          by (auto simp add: algebra_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6337
        also have "\<dots> = 0"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  6338
          using \<open>t > 0\<close> by (auto simp add:field_simps)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6339
        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
  6340
          unfolding w_def using False and \<open>t > 0\<close>
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6341
          by (auto simp add: algebra_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6342
        have  "2 * B < e * t"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  6343
          unfolding t_def using \<open>0 < e\<close> \<open>0 < k\<close> \<open>B > 0\<close> and as and False
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6344
          by (auto simp add:field_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6345
        then have "(f w - f x) / t < e"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  6346
          using B(2)[OF \<open>w\<in>s\<close>] and B(2)[OF \<open>x\<in>s\<close>]
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  6347
          using \<open>t > 0\<close> by (auto simp add:field_simps)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6348
        then have th1: "f y - f x < e"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6349
          apply -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6350
          apply (rule le_less_trans)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6351
          defer
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6352
          apply assumption
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6353
          using assms(2)[unfolded convex_on_def,rule_format,of w x "1/t" "(t - 1)/t", unfolded w]
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  6354
          using \<open>0 < t\<close> \<open>2 < t\<close> and \<open>x \<in> s\<close> \<open>w \<in> s\<close>
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6355
          by (auto simp add:field_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6356
      }
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  6357
      moreover
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6358
      {
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  6359
        define w where "w = x - t *\<^sub>R (y - x)"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6360
        have "w \<in> s"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6361
          unfolding w_def
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6362
          apply (rule k[unfolded subset_eq,rule_format])
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6363
          unfolding mem_cball dist_norm
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6364
          unfolding t_def
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  6365
          using \<open>k > 0\<close>
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6366
          apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6367
          done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6368
        have "(1 / (1 + t)) *\<^sub>R x + (t / (1 + t)) *\<^sub>R x = (1 / (1 + t) + t / (1 + t)) *\<^sub>R x"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6369
          by (auto simp add: algebra_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6370
        also have "\<dots> = x"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  6371
          using \<open>t > 0\<close> by (auto simp add:field_simps)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6372
        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
  6373
          unfolding w_def using False and \<open>t > 0\<close>
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6374
          by (auto simp add: algebra_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6375
        have "2 * B < e * t"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6376
          unfolding t_def
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  6377
          using \<open>0 < e\<close> \<open>0 < k\<close> \<open>B > 0\<close> and as and False
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6378
          by (auto simp add:field_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6379
        then have *: "(f w - f y) / t < e"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  6380
          using B(2)[OF \<open>w\<in>s\<close>] and B(2)[OF \<open>y\<in>s\<close>]
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  6381
          using \<open>t > 0\<close>
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6382
          by (auto simp add:field_simps)
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  6383
        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
  6384
          using assms(2)[unfolded convex_on_def,rule_format,of w y "1/(1+t)" "t / (1+t)",unfolded w]
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  6385
          using \<open>0 < t\<close> \<open>2 < t\<close> and \<open>y \<in> s\<close> \<open>w \<in> s\<close>
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6386
          by (auto simp add:field_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6387
        also have "\<dots> = (f w + t * f y) / (1 + t)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  6388
          using \<open>t > 0\<close> by (auto simp add: divide_simps)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6389
        also have "\<dots> < e + f y"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  6390
          using \<open>t > 0\<close> * \<open>e > 0\<close> by (auto simp add: field_simps)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6391
        finally have "f x - f y < e" by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6392
      }
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  6393
      ultimately show ?thesis by auto
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  6394
    qed (insert \<open>0<e\<close>, auto)
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  6395
  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
  6396
qed
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  6397
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  6398
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  6399
subsection \<open>Upper bound on a ball implies upper and lower bounds\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6400
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6401
lemma convex_bounds_lemma:
36338
7808fbc9c3b4 generalize more constants and lemmas
huffman
parents: 36337
diff changeset
  6402
  fixes x :: "'a::real_normed_vector"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6403
  assumes "convex_on (cball x e) f"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6404
    and "\<forall>y \<in> cball x e. f y \<le> b"
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61880
diff changeset
  6405
  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
  6406
  apply rule
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6407
proof (cases "0 \<le> e")
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6408
  case True
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6409
  fix y
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6410
  assume y: "y \<in> cball x e"
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  6411
  define z where "z = 2 *\<^sub>R x - y"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6412
  have *: "x - (2 *\<^sub>R x - y) = y - x"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6413
    by (simp add: scaleR_2)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6414
  have z: "z \<in> cball x e"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6415
    using y unfolding z_def mem_cball dist_norm * by (auto simp add: norm_minus_commute)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6416
  have "(1 / 2) *\<^sub>R y + (1 / 2) *\<^sub>R z = x"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6417
    unfolding z_def by (auto simp add: algebra_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6418
  then show "\<bar>f y\<bar> \<le> b + 2 * \<bar>f x\<bar>"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6419
    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
  6420
    using assms(2)[rule_format,OF y] assms(2)[rule_format,OF z]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6421
    by (auto simp add:field_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6422
next
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6423
  case False
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6424
  fix y
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6425
  assume "y \<in> cball x e"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6426
  then have "dist x y < 0"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6427
    using False unfolding mem_cball not_le by (auto simp del: dist_not_less_zero)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6428
  then show "\<bar>f y\<bar> \<le> b + 2 * \<bar>f x\<bar>"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6429
    using zero_le_dist[of x y] by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6430
qed
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6431
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6432
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  6433
subsubsection \<open>Hence a convex function on an open set is continuous\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6434
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
  6435
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
  6436
  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
  6437
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6438
lemma convex_on_continuous:
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6439
  assumes "open (s::('a::euclidean_space) set)" "convex_on s f"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6440
  shows "continuous_on s f"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6441
  unfolding continuous_on_eq_continuous_at[OF assms(1)]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6442
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
  6443
  note dimge1 = DIM_positive[where 'a='a]
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6444
  fix x
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6445
  assume "x \<in> s"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6446
  then obtain e where e: "cball x e \<subseteq> s" "e > 0"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6447
    using assms(1) unfolding open_contains_cball by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  6448
  define d where "d = e / real DIM('a)"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6449
  have "0 < d"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  6450
    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
  6451
  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
  6452
  obtain c
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6453
    where c: "finite c"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6454
    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
  6455
    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
  6456
  proof
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  6457
    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
  6458
    show "finite c"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6459
      unfolding c_def by (simp add: finite_set_setsum)
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  6460
    have 1: "convex hull c = {a. \<forall>i\<in>Basis. a \<bullet> i \<in> cbox (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
  6461
      unfolding box_eq_set_setsum_Basis
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6462
      unfolding c_def convex_hull_set_setsum
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6463
      apply (subst convex_hull_linear_image [symmetric])
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6464
      apply (simp add: linear_iff scaleR_add_left)
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  6465
      apply (rule setsum.cong [OF refl])
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6466
      apply (rule image_cong [OF _ refl])
56188
0268784f60da use cbox to relax class constraints
immler
parents: 56154
diff changeset
  6467
      apply (rule convex_hull_eq_real_cbox)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  6468
      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
  6469
      done
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6470
    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
  6471
      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
  6472
    show "cball x d \<subseteq> convex hull c"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6473
      unfolding 2
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6474
      apply clarsimp
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6475
      apply (simp only: dist_norm)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6476
      apply (subst inner_diff_left [symmetric])
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6477
      apply simp
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6478
      apply (erule (1) order_trans [OF Basis_le_norm])
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6479
      done
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6480
    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
  6481
      by (simp add: d_def DIM_positive)
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6482
    show "convex hull c \<subseteq> cball x e"
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6483
      unfolding 2
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6484
      apply clarsimp
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6485
      apply (subst euclidean_dist_l2)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6486
      apply (rule order_trans [OF setL2_le_setsum])
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6487
      apply (rule zero_le_dist)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6488
      unfolding e'
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6489
      apply (rule setsum_mono)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6490
      apply simp
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6491
      done
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6492
  qed
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  6493
  define k where "k = Max (f ` c)"
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6494
  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
  6495
    apply(rule convex_on_subset[OF assms(2)])
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50104
diff changeset
  6496
    apply(rule subset_trans[OF _ e(1)])
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6497
    apply(rule c1)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6498
    done
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6499
  then have k: "\<forall>y\<in>convex hull c. f y \<le> k"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6500
    apply (rule_tac convex_on_convex_hull_bound)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6501
    apply assumption
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6502
    unfolding k_def
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6503
    apply (rule, rule Max_ge)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6504
    using c(1)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6505
    apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6506
    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
  6507
  have "d \<le> e"
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50104
diff changeset
  6508
    unfolding d_def
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6509
    apply (rule mult_imp_div_pos_le)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  6510
    using \<open>e > 0\<close>
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
  6511
    unfolding mult_le_cancel_left1
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50104
diff changeset
  6512
    apply (auto simp: real_of_nat_ge_one_iff Suc_le_eq DIM_positive)
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50104
diff changeset
  6513
    done
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6514
  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
  6515
    by (rule subset_cball)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6516
  have conv: "convex_on (cball x d) f"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6517
    apply (rule convex_on_subset)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6518
    apply (rule convex_on_subset[OF assms(2)])
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6519
    apply (rule e(1))
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6520
    apply (rule dsube)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6521
    done
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61880
diff changeset
  6522
  then have "\<forall>y\<in>cball x d. \<bar>f y\<bar> \<le> k + 2 * \<bar>f x\<bar>"
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6523
    apply (rule convex_bounds_lemma)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6524
    apply (rule ballI)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6525
    apply (rule k [rule_format])
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6526
    apply (erule rev_subsetD)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6527
    apply (rule c2)
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  6528
    done
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6529
  then have "continuous_on (ball x d) f"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6530
    apply (rule_tac convex_on_bounded_continuous)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6531
    apply (rule open_ball, rule convex_on_subset[OF conv])
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6532
    apply (rule ball_subset_cball)
33270
paulson
parents: 33175
diff changeset
  6533
    apply force
paulson
parents: 33175
diff changeset
  6534
    done
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6535
  then show "continuous (at x) f"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6536
    unfolding continuous_on_eq_continuous_at[OF open_ball]
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  6537
    using \<open>d > 0\<close> by auto
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  6538
qed
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  6539
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  6540
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  6541
subsection \<open>Line segments, Starlike Sets, etc.\<close>
33270
paulson
parents: 33175
diff changeset
  6542
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  6543
(* Use the same overloading tricks as for intervals, so that
33270
paulson
parents: 33175
diff changeset
  6544
   segment[a,b] is closed and segment(a,b) is open relative to affine hull. *)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6545
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6546
definition midpoint :: "'a::real_vector \<Rightarrow> 'a \<Rightarrow> 'a"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6547
  where "midpoint a b = (inverse (2::real)) *\<^sub>R (a + b)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6548
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6549
definition closed_segment :: "'a::real_vector \<Rightarrow> 'a \<Rightarrow> 'a set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6550
  where "closed_segment a b = {(1 - u) *\<^sub>R a + u *\<^sub>R b | u::real. 0 \<le> u \<and> u \<le> 1}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6551
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6552
definition open_segment :: "'a::real_vector \<Rightarrow> 'a \<Rightarrow> 'a set" where
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6553
  "open_segment a b \<equiv> closed_segment a b - {a,b}"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6554
62626
de25474ce728 Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
  6555
lemmas segment = open_segment_def closed_segment_def
de25474ce728 Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
  6556
62948
7700f467892b lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 62843
diff changeset
  6557
lemma in_segment:
7700f467892b lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 62843
diff changeset
  6558
    "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)"
7700f467892b lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 62843
diff changeset
  6559
    "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)"
7700f467892b lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 62843
diff changeset
  6560
  using less_eq_real_def by (auto simp: segment algebra_simps)
7700f467892b lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 62843
diff changeset
  6561
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6562
lemma open_segment_PairD:
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6563
  "(x, x') \<in> open_segment (a, a') (b, b')
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6564
  \<Longrightarrow> (x \<in> open_segment a b \<or> a = b) \<and> (x' \<in> open_segment a' b' \<or> a' = b')"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6565
  by (auto simp: in_segment)
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6566
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6567
lemma closed_segment_PairD:
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6568
  "(x, x') \<in> closed_segment (a, a') (b, b') \<Longrightarrow> x \<in> closed_segment a b \<and> x' \<in> closed_segment a' b'"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6569
  by (auto simp: closed_segment_def)
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6570
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6571
lemma closed_segment_translation_eq [simp]:
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6572
    "d + x \<in> closed_segment (d + a) (d + b) \<longleftrightarrow> x \<in> closed_segment a b"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6573
proof -
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6574
  have *: "\<And>d x a b. x \<in> closed_segment a b \<Longrightarrow> d + x \<in> closed_segment (d + a) (d + b)"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6575
    apply (simp add: closed_segment_def)
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6576
    apply (erule ex_forward)
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6577
    apply (simp add: algebra_simps)
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6578
    done
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6579
  show ?thesis
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6580
  using * [where d = "-d"] *
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6581
  by (fastforce simp add:)
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6582
qed
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6583
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6584
lemma open_segment_translation_eq [simp]:
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6585
    "d + x \<in> open_segment (d + a) (d + b) \<longleftrightarrow> x \<in> open_segment a b"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6586
  by (simp add: open_segment_def)
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  6587
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6588
definition "between = (\<lambda>(a,b) x. x \<in> closed_segment a b)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6589
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  6590
definition "starlike s \<longleftrightarrow> (\<exists>a\<in>s. \<forall>x\<in>s. closed_segment a x \<subseteq> s)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  6591
62626
de25474ce728 Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
  6592
lemma starlike_UNIV [simp]: "starlike UNIV"
de25474ce728 Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
  6593
  by (simp add: starlike_def)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6594
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6595
lemma midpoint_refl: "midpoint x x = x"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6596
  unfolding midpoint_def
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6597
  unfolding scaleR_right_distrib
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6598
  unfolding scaleR_left_distrib[symmetric]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6599
  by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6600
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6601
lemma midpoint_sym: "midpoint a b = midpoint b a"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6602
  unfolding midpoint_def by (auto simp add: scaleR_right_distrib)
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6603
36338
7808fbc9c3b4 generalize more constants and lemmas
huffman
parents: 36337
diff changeset
  6604
lemma midpoint_eq_iff: "midpoint a b = c \<longleftrightarrow> a + b = c + c"
7808fbc9c3b4 generalize more constants and lemmas
huffman
parents: 36337
diff changeset
  6605
proof -
7808fbc9c3b4 generalize more constants and lemmas
huffman
parents: 36337
diff changeset
  6606
  have "midpoint a b = c \<longleftrightarrow> scaleR 2 (midpoint a b) = scaleR 2 c"
7808fbc9c3b4 generalize more constants and lemmas
huffman
parents: 36337
diff changeset
  6607
    by simp
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6608
  then show ?thesis
36338
7808fbc9c3b4 generalize more constants and lemmas
huffman
parents: 36337
diff changeset
  6609
    unfolding midpoint_def scaleR_2 [symmetric] by simp
7808fbc9c3b4 generalize more constants and lemmas
huffman
parents: 36337
diff changeset
  6610
qed
7808fbc9c3b4 generalize more constants and lemmas
huffman
parents: 36337
diff changeset
  6611
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6612
lemma dist_midpoint:
36338
7808fbc9c3b4 generalize more constants and lemmas
huffman
parents: 36337
diff changeset
  6613
  fixes a b :: "'a::real_normed_vector" shows
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6614
  "dist a (midpoint a b) = (dist a b) / 2" (is ?t1)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6615
  "dist b (midpoint a b) = (dist a b) / 2" (is ?t2)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6616
  "dist (midpoint a b) a = (dist a b) / 2" (is ?t3)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6617
  "dist (midpoint a b) b = (dist a b) / 2" (is ?t4)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6618
proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6619
  have *: "\<And>x y::'a. 2 *\<^sub>R x = - y \<Longrightarrow> norm x = (norm y) / 2"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6620
    unfolding equation_minus_iff by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6621
  have **: "\<And>x y::'a. 2 *\<^sub>R x =   y \<Longrightarrow> norm x = (norm y) / 2"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6622
    by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6623
  note scaleR_right_distrib [simp]
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6624
  show ?t1
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6625
    unfolding midpoint_def dist_norm
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6626
    apply (rule **)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6627
    apply (simp add: scaleR_right_diff_distrib)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6628
    apply (simp add: scaleR_2)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6629
    done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6630
  show ?t2
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6631
    unfolding midpoint_def dist_norm
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6632
    apply (rule *)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6633
    apply (simp add: scaleR_right_diff_distrib)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6634
    apply (simp add: scaleR_2)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6635
    done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6636
  show ?t3
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
  6637
    unfolding midpoint_def dist_norm
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6638
    apply (rule *)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6639
    apply (simp add: scaleR_right_diff_distrib)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6640
    apply (simp add: scaleR_2)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6641
    done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6642
  show ?t4
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6643
    unfolding midpoint_def dist_norm
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6644
    apply (rule **)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6645
    apply (simp add: scaleR_right_diff_distrib)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6646
    apply (simp add: scaleR_2)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6647
    done
36338
7808fbc9c3b4 generalize more constants and lemmas
huffman
parents: 36337
diff changeset
  6648
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6649
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6650
lemma midpoint_eq_endpoint:
36338
7808fbc9c3b4 generalize more constants and lemmas
huffman
parents: 36337
diff changeset
  6651
  "midpoint a b = a \<longleftrightarrow> a = b"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6652
  "midpoint a b = b \<longleftrightarrow> a = b"
36338
7808fbc9c3b4 generalize more constants and lemmas
huffman
parents: 36337
diff changeset
  6653
  unfolding midpoint_eq_iff by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6654
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6655
lemma convex_contains_segment:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6656
  "convex s \<longleftrightarrow> (\<forall>a\<in>s. \<forall>b\<in>s. closed_segment a b \<subseteq> s)"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6657
  unfolding convex_alt closed_segment_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6658
61848
9250e546ab23 New complex analysis material
paulson <lp15@cam.ac.uk>
parents: 61808
diff changeset
  6659
lemma closed_segment_subset: "\<lbrakk>x \<in> s; y \<in> s; convex s\<rbrakk> \<Longrightarrow> closed_segment x y \<subseteq> s"
9250e546ab23 New complex analysis material
paulson <lp15@cam.ac.uk>
parents: 61808
diff changeset
  6660
  by (simp add: convex_contains_segment)
9250e546ab23 New complex analysis material
paulson <lp15@cam.ac.uk>
parents: 61808
diff changeset
  6661
60762
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
  6662
lemma closed_segment_subset_convex_hull:
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
  6663
    "\<lbrakk>x \<in> convex hull s; y \<in> convex hull s\<rbrakk> \<Longrightarrow> closed_segment x y \<subseteq> convex hull s"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
  6664
  using convex_contains_segment by blast
60762
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
  6665
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6666
lemma convex_imp_starlike:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6667
  "convex s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> starlike s"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6668
  unfolding convex_contains_segment starlike_def by auto
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6669
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6670
lemma segment_convex_hull:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6671
  "closed_segment a b = convex hull {a,b}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6672
proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6673
  have *: "\<And>x. {x} \<noteq> {}" by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6674
  show ?thesis
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6675
    unfolding segment convex_hull_insert[OF *] convex_hull_singleton
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6676
    by (safe; rule_tac x="1 - u" in exI; force)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6677
qed
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6678
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  6679
lemma open_closed_segment: "u \<in> open_segment w z \<Longrightarrow> u \<in> closed_segment w z"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  6680
  by (auto simp add: closed_segment_def open_segment_def)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  6681
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6682
lemma segment_open_subset_closed:
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6683
   "open_segment a b \<subseteq> closed_segment a b"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6684
  by (auto simp: closed_segment_def open_segment_def)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6685
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6686
lemma bounded_closed_segment:
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6687
    fixes a :: "'a::euclidean_space" shows "bounded (closed_segment a b)"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6688
  by (simp add: segment_convex_hull compact_convex_hull compact_imp_bounded)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6689
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6690
lemma bounded_open_segment:
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6691
    fixes a :: "'a::euclidean_space" shows "bounded (open_segment a b)"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6692
  by (rule bounded_subset [OF bounded_closed_segment segment_open_subset_closed])
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6693
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6694
lemmas bounded_segment = bounded_closed_segment open_closed_segment
61426
d53db136e8fd new material on path_component_sets, inside, outside, etc. And more default simprules
paulson <lp15@cam.ac.uk>
parents: 61222
diff changeset
  6695
d53db136e8fd new material on path_component_sets, inside, outside, etc. And more default simprules
paulson <lp15@cam.ac.uk>
parents: 61222
diff changeset
  6696
lemma ends_in_segment [iff]: "a \<in> closed_segment a b" "b \<in> closed_segment a b"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6697
  unfolding segment_convex_hull
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6698
  by (auto intro!: hull_subset[unfolded subset_eq, rule_format])
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6699
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6700
lemma segment_furthest_le:
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 36844
diff changeset
  6701
  fixes a b x y :: "'a::euclidean_space"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6702
  assumes "x \<in> closed_segment a b"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6703
  shows "norm (y - x) \<le> norm (y - a) \<or>  norm (y - x) \<le> norm (y - b)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6704
proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6705
  obtain z where "z \<in> {a, b}" "norm (x - y) \<le> norm (z - y)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6706
    using simplex_furthest_le[of "{a, b}" y]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6707
    using assms[unfolded segment_convex_hull]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6708
    by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6709
  then show ?thesis
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6710
    by (auto simp add:norm_minus_commute)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  6711
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  6712
60176
38b630409aa2 closures of intervals
immler
parents: 59807
diff changeset
  6713
lemma closed_segment_commute: "closed_segment a b = closed_segment b a"
38b630409aa2 closures of intervals
immler
parents: 59807
diff changeset
  6714
proof -
38b630409aa2 closures of intervals
immler
parents: 59807
diff changeset
  6715
  have "{a, b} = {b, a}" by auto
38b630409aa2 closures of intervals
immler
parents: 59807
diff changeset
  6716
  thus ?thesis
38b630409aa2 closures of intervals
immler
parents: 59807
diff changeset
  6717
    by (simp add: segment_convex_hull)
38b630409aa2 closures of intervals
immler
parents: 59807
diff changeset
  6718
qed
38b630409aa2 closures of intervals
immler
parents: 59807
diff changeset
  6719
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6720
lemma segment_bound1:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6721
  assumes "x \<in> closed_segment a b"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6722
  shows "norm (x - a) \<le> norm (b - a)"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6723
proof -
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6724
  obtain u where "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 \<le> u" "u \<le> 1"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6725
    using assms by (auto simp add: closed_segment_def)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6726
  then show "norm (x - a) \<le> norm (b - a)"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6727
    apply clarify
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6728
    apply (auto simp: algebra_simps)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6729
    apply (simp add: scaleR_diff_right [symmetric] mult_left_le_one_le)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6730
    done
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6731
qed
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6732
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6733
lemma segment_bound:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6734
  assumes "x \<in> closed_segment a b"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6735
  shows "norm (x - a) \<le> norm (b - a)" "norm (x - b) \<le> norm (b - a)"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6736
apply (simp add: assms segment_bound1)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6737
by (metis assms closed_segment_commute dist_commute dist_norm segment_bound1)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6738
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  6739
lemma open_segment_commute: "open_segment a b = open_segment b a"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  6740
proof -
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  6741
  have "{a, b} = {b, a}" by auto
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  6742
  thus ?thesis
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  6743
    by (simp add: closed_segment_commute open_segment_def)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  6744
qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  6745
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  6746
lemma closed_segment_idem [simp]: "closed_segment a a = {a}"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  6747
  unfolding segment by (auto simp add: algebra_simps)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  6748
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  6749
lemma open_segment_idem [simp]: "open_segment a a = {}"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  6750
  by (simp add: open_segment_def)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  6751
62626
de25474ce728 Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
  6752
lemma closed_segment_eq_open: "closed_segment a b = open_segment a b \<union> {a,b}"
de25474ce728 Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
  6753
  using open_segment_def by auto
de25474ce728 Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
  6754
  
60176
38b630409aa2 closures of intervals
immler
parents: 59807
diff changeset
  6755
lemma closed_segment_eq_real_ivl:
38b630409aa2 closures of intervals
immler
parents: 59807
diff changeset
  6756
  fixes a b::real
38b630409aa2 closures of intervals
immler
parents: 59807
diff changeset
  6757
  shows "closed_segment a b = (if a \<le> b then {a .. b} else {b .. a})"
38b630409aa2 closures of intervals
immler
parents: 59807
diff changeset
  6758
proof -
38b630409aa2 closures of intervals
immler
parents: 59807
diff changeset
  6759
  have "b \<le> a \<Longrightarrow> closed_segment b a = {b .. a}"
38b630409aa2 closures of intervals
immler
parents: 59807
diff changeset
  6760
    and "a \<le> b \<Longrightarrow> closed_segment a b = {a .. b}"
38b630409aa2 closures of intervals
immler
parents: 59807
diff changeset
  6761
    by (auto simp: convex_hull_eq_real_cbox segment_convex_hull)
38b630409aa2 closures of intervals
immler
parents: 59807
diff changeset
  6762
  thus ?thesis
38b630409aa2 closures of intervals
immler
parents: 59807
diff changeset
  6763
    by (auto simp: closed_segment_commute)
38b630409aa2 closures of intervals
immler
parents: 59807
diff changeset
  6764
qed
38b630409aa2 closures of intervals
immler
parents: 59807
diff changeset
  6765
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
  6766
lemma closed_segment_real_eq:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
  6767
  fixes u::real shows "closed_segment u v = (\<lambda>x. (v - u) * x + u) ` {0..1}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
  6768
  by (simp add: add.commute [of u] image_affinity_atLeastAtMost [where c=u] closed_segment_eq_real_ivl)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
  6769
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6770
subsubsection\<open>More lemmas, especially for working with the underlying formula\<close>
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6771
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6772
lemma segment_eq_compose:
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6773
  fixes a :: "'a :: real_vector"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6774
  shows "(\<lambda>u. (1 - u) *\<^sub>R a + u *\<^sub>R b) = (\<lambda>x. a + x) o (\<lambda>u. u *\<^sub>R (b - a))"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6775
    by (simp add: o_def algebra_simps)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6776
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6777
lemma segment_degen_1:
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6778
  fixes a :: "'a :: real_vector"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6779
  shows "(1 - u) *\<^sub>R a + u *\<^sub>R b = b \<longleftrightarrow> a=b \<or> u=1"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6780
proof -
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6781
  { assume "(1 - u) *\<^sub>R a + u *\<^sub>R b = b"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6782
    then have "(1 - u) *\<^sub>R a = (1 - u) *\<^sub>R b"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6783
      by (simp add: algebra_simps)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6784
    then have "a=b \<or> u=1"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6785
      by simp
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6786
  } then show ?thesis
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6787
      by (auto simp: algebra_simps)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6788
qed
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6789
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6790
lemma segment_degen_0:
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6791
    fixes a :: "'a :: real_vector"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6792
    shows "(1 - u) *\<^sub>R a + u *\<^sub>R b = a \<longleftrightarrow> a=b \<or> u=0"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6793
  using segment_degen_1 [of "1-u" b a]
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6794
  by (auto simp: algebra_simps)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6795
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6796
lemma closed_segment_image_interval:
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6797
     "closed_segment a b = (\<lambda>u. (1 - u) *\<^sub>R a + u *\<^sub>R b) ` {0..1}"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6798
  by (auto simp: set_eq_iff image_iff closed_segment_def)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6799
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6800
lemma open_segment_image_interval:
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6801
     "open_segment a b = (if a=b then {} else (\<lambda>u. (1 - u) *\<^sub>R a + u *\<^sub>R b) ` {0<..<1})"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6802
  by (auto simp:  open_segment_def closed_segment_def segment_degen_0 segment_degen_1)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6803
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6804
lemmas segment_image_interval = closed_segment_image_interval open_segment_image_interval
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6805
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6806
lemma open_segment_bound1:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6807
  assumes "x \<in> open_segment a b"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6808
  shows "norm (x - a) < norm (b - a)"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6809
proof -
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6810
  obtain u where "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 < u" "u < 1" "a \<noteq> b"
62618
f7f2467ab854 Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  6811
    using assms by (auto simp add: open_segment_image_interval split: if_split_asm)
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6812
  then show "norm (x - a) < norm (b - a)"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6813
    apply clarify
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6814
    apply (auto simp: algebra_simps)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6815
    apply (simp add: scaleR_diff_right [symmetric])
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6816
    done
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6817
qed
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6818
62950
c355b3223cbd generalized
immler
parents: 62948
diff changeset
  6819
lemma compact_segment [simp]:
c355b3223cbd generalized
immler
parents: 62948
diff changeset
  6820
  fixes a :: "'a::real_normed_vector"
c355b3223cbd generalized
immler
parents: 62948
diff changeset
  6821
  shows "compact (closed_segment a b)"
c355b3223cbd generalized
immler
parents: 62948
diff changeset
  6822
  by (auto simp: segment_image_interval intro!: compact_continuous_image continuous_intros)
c355b3223cbd generalized
immler
parents: 62948
diff changeset
  6823
c355b3223cbd generalized
immler
parents: 62948
diff changeset
  6824
lemma closed_segment [simp]:
c355b3223cbd generalized
immler
parents: 62948
diff changeset
  6825
  fixes a :: "'a::real_normed_vector"
c355b3223cbd generalized
immler
parents: 62948
diff changeset
  6826
  shows "closed (closed_segment a b)"
c355b3223cbd generalized
immler
parents: 62948
diff changeset
  6827
  by (simp add: compact_imp_closed)
c355b3223cbd generalized
immler
parents: 62948
diff changeset
  6828
c355b3223cbd generalized
immler
parents: 62948
diff changeset
  6829
lemma closure_closed_segment [simp]:
c355b3223cbd generalized
immler
parents: 62948
diff changeset
  6830
  fixes a :: "'a::real_normed_vector"
c355b3223cbd generalized
immler
parents: 62948
diff changeset
  6831
  shows "closure(closed_segment a b) = closed_segment a b"
c355b3223cbd generalized
immler
parents: 62948
diff changeset
  6832
  by simp
c355b3223cbd generalized
immler
parents: 62948
diff changeset
  6833
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6834
lemma open_segment_bound:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6835
  assumes "x \<in> open_segment a b"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6836
  shows "norm (x - a) < norm (b - a)" "norm (x - b) < norm (b - a)"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6837
apply (simp add: assms open_segment_bound1)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6838
by (metis assms norm_minus_commute open_segment_bound1 open_segment_commute)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  6839
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6840
lemma closure_open_segment [simp]:
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6841
    fixes a :: "'a::euclidean_space"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6842
    shows "closure(open_segment a b) = (if a = b then {} else closed_segment a b)"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6843
proof -
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6844
  have "closure ((\<lambda>u. u *\<^sub>R (b - a)) ` {0<..<1}) = (\<lambda>u. u *\<^sub>R (b - a)) ` closure {0<..<1}" if "a \<noteq> b"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6845
    apply (rule closure_injective_linear_image [symmetric])
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6846
    apply (simp add:)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6847
    using that by (simp add: inj_on_def)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6848
  then show ?thesis
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6849
    by (simp add: segment_image_interval segment_eq_compose closure_greaterThanLessThan [symmetric]
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6850
         closure_translation image_comp [symmetric] del: closure_greaterThanLessThan)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6851
qed
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6852
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6853
lemma closed_open_segment_iff [simp]:
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6854
    fixes a :: "'a::euclidean_space"  shows "closed(open_segment a b) \<longleftrightarrow> a = b"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6855
  by (metis open_segment_def DiffE closure_eq closure_open_segment ends_in_segment(1) insert_iff segment_image_interval(2))
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6856
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6857
lemma compact_open_segment_iff [simp]:
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6858
    fixes a :: "'a::euclidean_space"  shows "compact(open_segment a b) \<longleftrightarrow> a = b"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6859
  by (simp add: bounded_open_segment compact_eq_bounded_closed)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6860
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6861
lemma convex_closed_segment [iff]: "convex (closed_segment a b)"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6862
  unfolding segment_convex_hull by(rule convex_convex_hull)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6863
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6864
lemma convex_open_segment [iff]: "convex(open_segment a b)"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6865
proof -
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6866
  have "convex ((\<lambda>u. u *\<^sub>R (b-a)) ` {0<..<1})"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6867
    by (rule convex_linear_image) auto
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6868
  then show ?thesis
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6869
    apply (simp add: open_segment_image_interval segment_eq_compose)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6870
    by (metis image_comp convex_translation)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6871
qed
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6872
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6873
lemmas convex_segment = convex_closed_segment convex_open_segment
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6874
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6875
lemma connected_segment [iff]:
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6876
  fixes x :: "'a :: real_normed_vector"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6877
  shows "connected (closed_segment x y)"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6878
  by (simp add: convex_connected)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6879
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6880
lemma affine_hull_closed_segment [simp]:
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6881
     "affine hull (closed_segment a b) = affine hull {a,b}"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6882
  by (simp add: segment_convex_hull)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6883
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6884
lemma affine_hull_open_segment [simp]:
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6885
    fixes a :: "'a::euclidean_space"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6886
    shows "affine hull (open_segment a b) = (if a = b then {} else affine hull {a,b})"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6887
by (metis affine_hull_convex_hull affine_hull_empty closure_open_segment closure_same_affine_hull segment_convex_hull)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  6888
62618
f7f2467ab854 Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  6889
lemma rel_interior_closure_convex_segment:
f7f2467ab854 Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  6890
  fixes S :: "_::euclidean_space set"
f7f2467ab854 Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  6891
  assumes "convex S" "a \<in> rel_interior S" "b \<in> closure S"
f7f2467ab854 Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  6892
    shows "open_segment a b \<subseteq> rel_interior S"
f7f2467ab854 Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  6893
proof
f7f2467ab854 Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  6894
  fix x
f7f2467ab854 Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  6895
  have [simp]: "(1 - u) *\<^sub>R a + u *\<^sub>R b = b - (1 - u) *\<^sub>R (b - a)" for u
f7f2467ab854 Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  6896
    by (simp add: algebra_simps)
f7f2467ab854 Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  6897
  assume "x \<in> open_segment a b"
f7f2467ab854 Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  6898
  then show "x \<in> rel_interior S"
f7f2467ab854 Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  6899
    unfolding closed_segment_def open_segment_def  using assms
f7f2467ab854 Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  6900
    by (auto intro: rel_interior_closure_convex_shrink)
f7f2467ab854 Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  6901
qed
f7f2467ab854 Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  6902
62620
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6903
subsection\<open>More results about segments\<close>
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6904
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6905
lemma dist_half_times2:
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6906
  fixes a :: "'a :: real_normed_vector"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6907
  shows "dist ((1 / 2) *\<^sub>R (a + b)) x * 2 = dist (a+b) (2 *\<^sub>R x)"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6908
proof -
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6909
  have "norm ((1 / 2) *\<^sub>R (a + b) - x) * 2 = norm (2 *\<^sub>R ((1 / 2) *\<^sub>R (a + b) - x))"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6910
    by simp
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6911
  also have "... = norm ((a + b) - 2 *\<^sub>R x)"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6912
    by (simp add: real_vector.scale_right_diff_distrib)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6913
  finally show ?thesis
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6914
    by (simp only: dist_norm)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6915
qed
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6916
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6917
lemma closed_segment_as_ball:
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6918
    "closed_segment a b = affine hull {a,b} \<inter> cball(inverse 2 *\<^sub>R (a + b))(norm(b - a) / 2)"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6919
proof (cases "b = a")
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6920
  case True then show ?thesis by (auto simp: hull_inc)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6921
next
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6922
  case False
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6923
  then have *: "((\<exists>u v. x = u *\<^sub>R a + v *\<^sub>R b \<and> u + v = 1) \<and>
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6924
                  dist ((1 / 2) *\<^sub>R (a + b)) x * 2 \<le> norm (b - a)) =
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6925
                 (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> 0 \<le> u \<and> u \<le> 1)" for x
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6926
  proof -
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6927
    have "((\<exists>u v. x = u *\<^sub>R a + v *\<^sub>R b \<and> u + v = 1) \<and>
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6928
                  dist ((1 / 2) *\<^sub>R (a + b)) x * 2 \<le> norm (b - a)) =
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6929
          ((\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b) \<and>
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6930
                  dist ((1 / 2) *\<^sub>R (a + b)) x * 2 \<le> norm (b - a))"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6931
      unfolding eq_diff_eq [symmetric] by simp
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6932
    also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and>
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6933
                          norm ((a+b) - (2 *\<^sub>R x)) \<le> norm (b - a))"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6934
      by (simp add: dist_half_times2) (simp add: dist_norm)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6935
    also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and>
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6936
            norm ((a+b) - (2 *\<^sub>R ((1 - u) *\<^sub>R a + u *\<^sub>R b))) \<le> norm (b - a))"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6937
      by auto
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6938
    also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and>
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6939
                norm ((1 - u * 2) *\<^sub>R (b - a)) \<le> norm (b - a))"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6940
      by (simp add: algebra_simps scaleR_2)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6941
    also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and>
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6942
                          \<bar>1 - u * 2\<bar> * norm (b - a) \<le> norm (b - a))"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6943
      by simp
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6944
    also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> \<bar>1 - u * 2\<bar> \<le> 1)"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6945
      by (simp add: mult_le_cancel_right2 False)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6946
    also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> 0 \<le> u \<and> u \<le> 1)"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6947
      by auto
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6948
    finally show ?thesis .
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6949
  qed
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6950
  show ?thesis
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6951
    by (simp add: affine_hull_2 Set.set_eq_iff closed_segment_def *)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6952
qed
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6953
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6954
lemma open_segment_as_ball:
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6955
    "open_segment a b =
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6956
     affine hull {a,b} \<inter> ball(inverse 2 *\<^sub>R (a + b))(norm(b - a) / 2)"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6957
proof (cases "b = a")
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6958
  case True then show ?thesis by (auto simp: hull_inc)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6959
next
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6960
  case False
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6961
  then have *: "((\<exists>u v. x = u *\<^sub>R a + v *\<^sub>R b \<and> u + v = 1) \<and>
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6962
                  dist ((1 / 2) *\<^sub>R (a + b)) x * 2 < norm (b - a)) =
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6963
                 (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> 0 < u \<and> u < 1)" for x
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6964
  proof -
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6965
    have "((\<exists>u v. x = u *\<^sub>R a + v *\<^sub>R b \<and> u + v = 1) \<and>
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6966
                  dist ((1 / 2) *\<^sub>R (a + b)) x * 2 < norm (b - a)) =
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6967
          ((\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b) \<and>
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6968
                  dist ((1 / 2) *\<^sub>R (a + b)) x * 2 < norm (b - a))"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6969
      unfolding eq_diff_eq [symmetric] by simp
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6970
    also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and>
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6971
                          norm ((a+b) - (2 *\<^sub>R x)) < norm (b - a))"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6972
      by (simp add: dist_half_times2) (simp add: dist_norm)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6973
    also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and>
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6974
            norm ((a+b) - (2 *\<^sub>R ((1 - u) *\<^sub>R a + u *\<^sub>R b))) < norm (b - a))"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6975
      by auto
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6976
    also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and>
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6977
                norm ((1 - u * 2) *\<^sub>R (b - a)) < norm (b - a))"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6978
      by (simp add: algebra_simps scaleR_2)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6979
    also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and>
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6980
                          \<bar>1 - u * 2\<bar> * norm (b - a) < norm (b - a))"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6981
      by simp
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6982
    also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> \<bar>1 - u * 2\<bar> < 1)"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6983
      by (simp add: mult_le_cancel_right2 False)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6984
    also have "... = (\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> 0 < u \<and> u < 1)"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6985
      by auto
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6986
    finally show ?thesis .
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6987
  qed
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6988
  show ?thesis
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6989
    using False by (force simp: affine_hull_2 Set.set_eq_iff open_segment_image_interval *)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6990
qed
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6991
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6992
lemmas segment_as_ball = closed_segment_as_ball open_segment_as_ball
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6993
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6994
lemma closed_segment_neq_empty [simp]: "closed_segment a b \<noteq> {}"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6995
  by auto
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6996
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6997
lemma open_segment_eq_empty [simp]: "open_segment a b = {} \<longleftrightarrow> a = b"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6998
proof -
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  6999
  { assume a1: "open_segment a b = {}"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7000
    have "{} \<noteq> {0::real<..<1}"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7001
      by simp
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7002
    then have "a = b"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7003
      using a1 open_segment_image_interval by fastforce
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7004
  } then show ?thesis by auto
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7005
qed
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7006
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7007
lemma open_segment_eq_empty' [simp]: "{} = open_segment a b \<longleftrightarrow> a = b"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7008
  using open_segment_eq_empty by blast
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7009
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7010
lemmas segment_eq_empty = closed_segment_neq_empty open_segment_eq_empty
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7011
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7012
lemma inj_segment:
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7013
  fixes a :: "'a :: real_vector"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7014
  assumes "a \<noteq> b"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7015
    shows "inj_on (\<lambda>u. (1 - u) *\<^sub>R a + u *\<^sub>R b) I"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7016
proof
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7017
  fix x y
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7018
  assume "(1 - x) *\<^sub>R a + x *\<^sub>R b = (1 - y) *\<^sub>R a + y *\<^sub>R b"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7019
  then have "x *\<^sub>R (b - a) = y *\<^sub>R (b - a)"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7020
    by (simp add: algebra_simps)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7021
  with assms show "x = y"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7022
    by (simp add: real_vector.scale_right_imp_eq)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7023
qed
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7024
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7025
lemma finite_closed_segment [simp]: "finite(closed_segment a b) \<longleftrightarrow> a = b"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7026
  apply auto
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7027
  apply (rule ccontr)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7028
  apply (simp add: segment_image_interval)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7029
  using infinite_Icc [OF zero_less_one] finite_imageD [OF _ inj_segment] apply blast
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7030
  done
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7031
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7032
lemma finite_open_segment [simp]: "finite(open_segment a b) \<longleftrightarrow> a = b"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7033
  by (auto simp: open_segment_def)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7034
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7035
lemmas finite_segment = finite_closed_segment finite_open_segment
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7036
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7037
lemma closed_segment_eq_sing: "closed_segment a b = {c} \<longleftrightarrow> a = c \<and> b = c"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7038
  by auto
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7039
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7040
lemma open_segment_eq_sing: "open_segment a b \<noteq> {c}"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7041
  by (metis finite_insert finite_open_segment insert_not_empty open_segment_image_interval)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7042
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7043
lemmas segment_eq_sing = closed_segment_eq_sing open_segment_eq_sing
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7044
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7045
lemma subset_closed_segment:
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7046
    "closed_segment a b \<subseteq> closed_segment c d \<longleftrightarrow>
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7047
     a \<in> closed_segment c d \<and> b \<in> closed_segment c d"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7048
  by auto (meson contra_subsetD convex_closed_segment convex_contains_segment)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7049
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7050
lemma subset_co_segment:
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7051
    "closed_segment a b \<subseteq> open_segment c d \<longleftrightarrow>
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7052
     a \<in> open_segment c d \<and> b \<in> open_segment c d"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7053
using closed_segment_subset by blast
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7054
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7055
lemma subset_open_segment:
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7056
  fixes a :: "'a::euclidean_space"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7057
  shows "open_segment a b \<subseteq> open_segment c d \<longleftrightarrow>
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7058
         a = b \<or> a \<in> closed_segment c d \<and> b \<in> closed_segment c d"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7059
        (is "?lhs = ?rhs")
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7060
proof (cases "a = b")
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7061
  case True then show ?thesis by simp
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7062
next
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7063
  case False show ?thesis
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7064
  proof
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7065
    assume rhs: ?rhs
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7066
    with \<open>a \<noteq> b\<close> have "c \<noteq> d"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7067
      using closed_segment_idem singleton_iff by auto
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7068
    have "\<exists>uc. (1 - u) *\<^sub>R ((1 - ua) *\<^sub>R c + ua *\<^sub>R d) + u *\<^sub>R ((1 - ub) *\<^sub>R c + ub *\<^sub>R d) =
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7069
               (1 - uc) *\<^sub>R c + uc *\<^sub>R d \<and> 0 < uc \<and> uc < 1"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7070
        if neq: "(1 - ua) *\<^sub>R c + ua *\<^sub>R d \<noteq> (1 - ub) *\<^sub>R c + ub *\<^sub>R d" "c \<noteq> d"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7071
           and "a = (1 - ua) *\<^sub>R c + ua *\<^sub>R d" "b = (1 - ub) *\<^sub>R c + ub *\<^sub>R d"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7072
           and u: "0 < u" "u < 1" and uab: "0 \<le> ua" "ua \<le> 1" "0 \<le> ub" "ub \<le> 1"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7073
        for u ua ub
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7074
    proof -
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7075
      have "ua \<noteq> ub"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7076
        using neq by auto
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7077
      moreover have "(u - 1) * ua \<le> 0" using u uab
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7078
        by (simp add: mult_nonpos_nonneg)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7079
      ultimately have lt: "(u - 1) * ua < u * ub" using u uab
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7080
        by (metis antisym_conv diff_ge_0_iff_ge le_less_trans mult_eq_0_iff mult_le_0_iff not_less)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7081
      have "p * ua + q * ub < p+q" if p: "0 < p" and  q: "0 < q" for p q
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7082
      proof -
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7083
        have "\<not> p \<le> 0" "\<not> q \<le> 0"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7084
          using p q not_less by blast+
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7085
        then show ?thesis
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7086
          by (metis \<open>ua \<noteq> ub\<close> add_less_cancel_left add_less_cancel_right add_mono_thms_linordered_field(5)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7087
                    less_eq_real_def mult_cancel_left1 mult_less_cancel_left2 uab(2) uab(4))
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7088
      qed
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7089
      then have "(1 - u) * ua + u * ub < 1" using u \<open>ua \<noteq> ub\<close>
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7090
        by (metis diff_add_cancel diff_gt_0_iff_gt)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7091
      with lt show ?thesis
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7092
        by (rule_tac x="ua + u*(ub-ua)" in exI) (simp add: algebra_simps)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7093
    qed
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7094
    with rhs \<open>a \<noteq> b\<close> \<open>c \<noteq> d\<close> show ?lhs
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7095
      unfolding open_segment_image_interval closed_segment_def
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7096
      by (fastforce simp add:)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7097
  next
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7098
    assume lhs: ?lhs
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7099
    with \<open>a \<noteq> b\<close> have "c \<noteq> d"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7100
      by (meson finite_open_segment rev_finite_subset)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7101
    have "closure (open_segment a b) \<subseteq> closure (open_segment c d)"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7102
      using lhs closure_mono by blast
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7103
    then have "closed_segment a b \<subseteq> closed_segment c d"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7104
      by (simp add: \<open>a \<noteq> b\<close> \<open>c \<noteq> d\<close>)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7105
    then show ?rhs
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7106
      by (force simp: \<open>a \<noteq> b\<close>)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7107
  qed
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7108
qed
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7109
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7110
lemma subset_oc_segment:
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7111
  fixes a :: "'a::euclidean_space"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7112
  shows "open_segment a b \<subseteq> closed_segment c d \<longleftrightarrow>
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7113
         a = b \<or> a \<in> closed_segment c d \<and> b \<in> closed_segment c d"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7114
apply (simp add: subset_open_segment [symmetric])
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7115
apply (rule iffI)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7116
 apply (metis closure_closed_segment closure_mono closure_open_segment subset_closed_segment subset_open_segment)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7117
apply (meson dual_order.trans segment_open_subset_closed)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7118
done
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7119
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7120
lemmas subset_segment = subset_closed_segment subset_co_segment subset_oc_segment subset_open_segment
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7121
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  7122
62618
f7f2467ab854 Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  7123
subsection\<open>Betweenness\<close>
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  7124
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  7125
lemma between_mem_segment: "between (a,b) x \<longleftrightarrow> x \<in> closed_segment a b"
44170
510ac30f44c0 make Multivariate_Analysis work with separate set type
huffman
parents: 44142
diff changeset
  7126
  unfolding between_def by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  7127
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7128
lemma between: "between (a, b) (x::'a::euclidean_space) \<longleftrightarrow> dist a b = (dist a x) + (dist x b)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7129
proof (cases "a = b")
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7130
  case True
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7131
  then show ?thesis
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7132
    unfolding between_def split_conv
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  7133
    by (auto simp add: dist_commute)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7134
next
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7135
  case False
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7136
  then have Fal: "norm (a - b) \<noteq> 0" and Fal2: "norm (a - b) > 0"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7137
    by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7138
  have *: "\<And>u. a - ((1 - u) *\<^sub>R a + u *\<^sub>R b) = u *\<^sub>R (a - b)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7139
    by (auto simp add: algebra_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7140
  show ?thesis
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7141
    unfolding between_def split_conv closed_segment_def mem_Collect_eq
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7142
    apply rule
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7143
    apply (elim exE conjE)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7144
    apply (subst dist_triangle_eq)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7145
  proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7146
    fix u
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7147
    assume as: "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 \<le> u" "u \<le> 1"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7148
    then have *: "a - x = u *\<^sub>R (a - b)" "x - b = (1 - u) *\<^sub>R (a - b)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7149
      unfolding as(1) by (auto simp add:algebra_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7150
    show "norm (a - x) *\<^sub>R (x - b) = norm (x - b) *\<^sub>R (a - x)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7151
      unfolding norm_minus_commute[of x a] * using as(2,3)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7152
      by (auto simp add: field_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7153
  next
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7154
    assume as: "dist a b = dist a x + dist x b"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7155
    have "norm (a - x) / norm (a - b) \<le> 1"
56571
f4635657d66f added divide_nonneg_nonneg and co; made it a simp rule
hoelzl
parents: 56544
diff changeset
  7156
      using Fal2 unfolding as[unfolded dist_norm] norm_ge_zero by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7157
    then show "\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> 0 \<le> u \<and> u \<le> 1"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7158
      apply (rule_tac x="dist a x / dist a b" in exI)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7159
      unfolding dist_norm
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7160
      apply (subst euclidean_eq_iff)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7161
      apply rule
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7162
      defer
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7163
      apply rule
56571
f4635657d66f added divide_nonneg_nonneg and co; made it a simp rule
hoelzl
parents: 56544
diff changeset
  7164
      prefer 3
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7165
      apply rule
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7166
    proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7167
      fix i :: 'a
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7168
      assume i: "i \<in> Basis"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7169
      have "((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R b) \<bullet> i =
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7170
        ((norm (a - b) - norm (a - x)) * (a \<bullet> i) + norm (a - x) * (b \<bullet> i)) / norm (a - b)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7171
        using Fal by (auto simp add: field_simps inner_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7172
      also have "\<dots> = x\<bullet>i"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7173
        apply (rule divide_eq_imp[OF Fal])
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7174
        unfolding as[unfolded dist_norm]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7175
        using as[unfolded dist_triangle_eq]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7176
        apply -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7177
        apply (subst (asm) euclidean_eq_iff)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7178
        using i
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7179
        apply (erule_tac x=i in ballE)
57865
dcfb33c26f50 tuned proofs -- fewer warnings;
wenzelm
parents: 57512
diff changeset
  7180
        apply (auto simp add: field_simps inner_simps)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7181
        done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7182
      finally show "x \<bullet> i =
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7183
        ((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R b) \<bullet> i"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7184
        by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7185
    qed (insert Fal2, auto)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7186
  qed
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7187
qed
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7188
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7189
lemma between_midpoint:
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7190
  fixes a :: "'a::euclidean_space"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7191
  shows "between (a,b) (midpoint a b)" (is ?t1)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7192
    and "between (b,a) (midpoint a b)" (is ?t2)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7193
proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7194
  have *: "\<And>x y z. x = (1/2::real) *\<^sub>R z \<Longrightarrow> y = (1/2) *\<^sub>R z \<Longrightarrow> norm z = norm x + norm y"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7195
    by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7196
  show ?t1 ?t2
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7197
    unfolding between midpoint_def dist_norm
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7198
    apply(rule_tac[!] *)
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
  7199
    unfolding euclidean_eq_iff[where 'a='a]
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7200
    apply (auto simp add: field_simps inner_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7201
    done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7202
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  7203
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  7204
lemma between_mem_convex_hull:
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  7205
  "between (a,b) x \<longleftrightarrow> x \<in> convex hull {a,b}"
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  7206
  unfolding between_mem_segment segment_convex_hull ..
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  7207
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7208
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7209
subsection \<open>Shrinking towards the interior of a convex set\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  7210
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  7211
lemma mem_interior_convex_shrink:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7212
  fixes s :: "'a::euclidean_space set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7213
  assumes "convex s"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7214
    and "c \<in> interior s"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7215
    and "x \<in> s"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7216
    and "0 < e"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7217
    and "e \<le> 1"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  7218
  shows "x - e *\<^sub>R (x - c) \<in> interior s"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7219
proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7220
  obtain d where "d > 0" and d: "ball c d \<subseteq> s"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7221
    using assms(2) unfolding mem_interior by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7222
  show ?thesis
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7223
    unfolding mem_interior
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7224
    apply (rule_tac x="e*d" in exI)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7225
    apply rule
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7226
    defer
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7227
    unfolding subset_eq Ball_def mem_ball
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7228
  proof (rule, rule)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7229
    fix y
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7230
    assume as: "dist (x - e *\<^sub>R (x - c)) y < e * d"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7231
    have *: "y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7232
      using \<open>e > 0\<close> by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib)
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61880
diff changeset
  7233
    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)"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7234
      unfolding dist_norm
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7235
      unfolding norm_scaleR[symmetric]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7236
      apply (rule arg_cong[where f=norm])
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7237
      using \<open>e > 0\<close>
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7238
      by (auto simp add: euclidean_eq_iff[where 'a='a] field_simps inner_simps)
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61880
diff changeset
  7239
    also have "\<dots> = \<bar>1/e\<bar> * norm (x - e *\<^sub>R (x - c) - y)"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7240
      by (auto intro!:arg_cong[where f=norm] simp add: algebra_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7241
    also have "\<dots> < d"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7242
      using as[unfolded dist_norm] and \<open>e > 0\<close>
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7243
      by (auto simp add:pos_divide_less_eq[OF \<open>e > 0\<close>] mult.commute)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7244
    finally show "y \<in> s"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7245
      apply (subst *)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7246
      apply (rule assms(1)[unfolded convex_alt,rule_format])
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7247
      apply (rule d[unfolded subset_eq,rule_format])
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7248
      unfolding mem_ball
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7249
      using assms(3-5)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7250
      apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7251
      done
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7252
  qed (insert \<open>e>0\<close> \<open>d>0\<close>, auto)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7253
qed
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  7254
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  7255
lemma mem_interior_closure_convex_shrink:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7256
  fixes s :: "'a::euclidean_space set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7257
  assumes "convex s"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7258
    and "c \<in> interior s"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7259
    and "x \<in> closure s"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7260
    and "0 < e"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7261
    and "e \<le> 1"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  7262
  shows "x - e *\<^sub>R (x - c) \<in> interior s"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7263
proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7264
  obtain d where "d > 0" and d: "ball c d \<subseteq> s"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7265
    using assms(2) unfolding mem_interior by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7266
  have "\<exists>y\<in>s. norm (y - x) * (1 - e) < e * d"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7267
  proof (cases "x \<in> s")
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7268
    case True
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7269
    then show ?thesis
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7270
      using \<open>e > 0\<close> \<open>d > 0\<close>
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7271
      apply (rule_tac bexI[where x=x])
56544
b60d5d119489 made mult_pos_pos a simp rule
nipkow
parents: 56541
diff changeset
  7272
      apply (auto)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7273
      done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7274
  next
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7275
    case False
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7276
    then have x: "x islimpt s"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7277
      using assms(3)[unfolded closure_def] by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7278
    show ?thesis
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7279
    proof (cases "e = 1")
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7280
      case True
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7281
      obtain y where "y \<in> s" "y \<noteq> x" "dist y x < 1"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  7282
        using x[unfolded islimpt_approachable,THEN spec[where x=1]] by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7283
      then show ?thesis
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7284
        apply (rule_tac x=y in bexI)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7285
        unfolding True
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7286
        using \<open>d > 0\<close>
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7287
        apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7288
        done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7289
    next
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7290
      case False
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7291
      then have "0 < e * d / (1 - e)" and *: "1 - e > 0"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7292
        using \<open>e \<le> 1\<close> \<open>e > 0\<close> \<open>d > 0\<close> by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7293
      then obtain y where "y \<in> s" "y \<noteq> x" "dist y x < e * d / (1 - e)"
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  7294
        using x[unfolded islimpt_approachable,THEN spec[where x="e*d / (1 - e)"]] by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7295
      then show ?thesis
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7296
        apply (rule_tac x=y in bexI)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7297
        unfolding dist_norm
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7298
        using pos_less_divide_eq[OF *]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7299
        apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7300
        done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7301
    qed
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7302
  qed
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7303
  then obtain y where "y \<in> s" and y: "norm (y - x) * (1 - e) < e * d"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7304
    by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  7305
  define z where "z = c + ((1 - e) / e) *\<^sub>R (x - y)"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7306
  have *: "x - e *\<^sub>R (x - c) = y - e *\<^sub>R (y - z)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7307
    unfolding z_def using \<open>e > 0\<close>
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7308
    by (auto simp add: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7309
  have "z \<in> interior s"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7310
    apply (rule interior_mono[OF d,unfolded subset_eq,rule_format])
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  7311
    unfolding interior_open[OF open_ball] mem_ball z_def dist_norm using y and assms(4,5)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7312
    apply (auto simp add:field_simps norm_minus_commute)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7313
    done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7314
  then show ?thesis
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7315
    unfolding *
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7316
    apply -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7317
    apply (rule mem_interior_convex_shrink)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7318
    using assms(1,4-5) \<open>y\<in>s\<close>
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7319
    apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7320
    done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7321
qed
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7322
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  7323
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7324
subsection \<open>Some obvious but surprisingly hard simplex lemmas\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  7325
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  7326
lemma simplex:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7327
  assumes "finite s"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7328
    and "0 \<notin> s"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7329
  shows "convex hull (insert 0 s) =
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7330
    {y. (\<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s \<le> 1 \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y)}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7331
  unfolding convex_hull_finite[OF finite.insertI[OF assms(1)]]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7332
  apply (rule set_eqI, rule)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7333
  unfolding mem_Collect_eq
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7334
  apply (erule_tac[!] exE)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7335
  apply (erule_tac[!] conjE)+
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7336
  unfolding setsum_clauses(2)[OF assms(1)]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7337
  apply (rule_tac x=u in exI)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7338
  defer
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7339
  apply (rule_tac x="\<lambda>x. if x = 0 then 1 - setsum u s else u x" in exI)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7340
  using assms(2)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7341
  unfolding if_smult and setsum_delta_notmem[OF assms(2)]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7342
  apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7343
  done
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  7344
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
  7345
lemma substd_simplex:
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50104
diff changeset
  7346
  assumes d: "d \<subseteq> Basis"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7347
  shows "convex hull (insert 0 d) =
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  7348
    {x. (\<forall>i\<in>Basis. 0 \<le> x\<bullet>i) \<and> (\<Sum>i\<in>d. x\<bullet>i) \<le> 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0)}"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7349
  (is "convex hull (insert 0 ?p) = ?s")
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7350
proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7351
  let ?D = d
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7352
  have "0 \<notin> ?p"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7353
    using assms by (auto simp: image_def)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7354
  from d have "finite d"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7355
    by (blast intro: finite_subset finite_Basis)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7356
  show ?thesis
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7357
    unfolding simplex[OF \<open>finite d\<close> \<open>0 \<notin> ?p\<close>]
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7358
    apply (rule set_eqI)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7359
    unfolding mem_Collect_eq
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7360
    apply rule
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7361
    apply (elim exE conjE)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7362
    apply (erule_tac[2] conjE)+
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7363
  proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7364
    fix x :: "'a::euclidean_space"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7365
    fix u
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7366
    assume as: "\<forall>x\<in>?D. 0 \<le> u x" "setsum u ?D \<le> 1" "(\<Sum>x\<in>?D. u x *\<^sub>R x) = x"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7367
    have *: "\<forall>i\<in>Basis. i:d \<longrightarrow> u i = x\<bullet>i"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  7368
      and "(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7369
      using as(3)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7370
      unfolding substdbasis_expansion_unique[OF assms]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7371
      by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7372
    then have **: "setsum u ?D = setsum (op \<bullet> x) ?D"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7373
      apply -
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  7374
      apply (rule setsum.cong)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7375
      using assms
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7376
      apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7377
      done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7378
    have "(\<forall>i\<in>Basis. 0 \<le> x\<bullet>i) \<and> setsum (op \<bullet> x) ?D \<le> 1"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7379
    proof (rule,rule)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7380
      fix i :: 'a
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7381
      assume i: "i \<in> Basis"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7382
      have "i \<in> d \<Longrightarrow> 0 \<le> x\<bullet>i"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7383
        unfolding *[rule_format,OF i,symmetric]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7384
         apply (rule_tac as(1)[rule_format])
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7385
         apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7386
         done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7387
      moreover have "i \<notin> d \<Longrightarrow> 0 \<le> x\<bullet>i"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7388
        using \<open>(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)\<close>[rule_format, OF i] 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
  7389
      ultimately show "0 \<le> x\<bullet>i" by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7390
    qed (insert as(2)[unfolded **], auto)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7391
    then show "(\<forall>i\<in>Basis. 0 \<le> x\<bullet>i) \<and> setsum (op \<bullet> x) ?D \<le> 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7392
      using \<open>(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)\<close> by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7393
  next
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7394
    fix x :: "'a::euclidean_space"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7395
    assume as: "\<forall>i\<in>Basis. 0 \<le> x \<bullet> i" "setsum (op \<bullet> x) ?D \<le> 1" "(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)"
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
  7396
    show "\<exists>u. (\<forall>x\<in>?D. 0 \<le> u x) \<and> setsum u ?D \<le> 1 \<and> (\<Sum>x\<in>?D. u x *\<^sub>R x) = x"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7397
      using as d
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7398
      unfolding substdbasis_expansion_unique[OF assms]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7399
      apply (rule_tac x="inner x" in exI)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7400
      apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7401
      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
  7402
  qed
899c9c4e4a4c Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents: 50104
diff changeset
  7403
qed
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7404
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  7405
lemma std_simplex:
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
  7406
  "convex hull (insert 0 Basis) =
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7407
    {x::'a::euclidean_space. (\<forall>i\<in>Basis. 0 \<le> x\<bullet>i) \<and> setsum (\<lambda>i. x\<bullet>i) Basis \<le> 1}"
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
  7408
  using substd_simplex[of Basis] by auto
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  7409
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  7410
lemma interior_std_simplex:
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
  7411
  "interior (convex hull (insert 0 Basis)) =
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7412
    {x::'a::euclidean_space. (\<forall>i\<in>Basis. 0 < x\<bullet>i) \<and> setsum (\<lambda>i. x\<bullet>i) Basis < 1}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7413
  apply (rule set_eqI)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7414
  unfolding mem_interior std_simplex
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7415
  unfolding subset_eq mem_Collect_eq Ball_def mem_ball
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7416
  unfolding Ball_def[symmetric]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7417
  apply rule
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7418
  apply (elim exE conjE)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7419
  defer
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7420
  apply (erule conjE)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7421
proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7422
  fix x :: 'a
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7423
  fix e
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7424
  assume "e > 0" and as: "\<forall>xa. dist x xa < e \<longrightarrow> (\<forall>x\<in>Basis. 0 \<le> xa \<bullet> x) \<and> setsum (op \<bullet> xa) Basis \<le> 1"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7425
  show "(\<forall>xa\<in>Basis. 0 < x \<bullet> xa) \<and> setsum (op \<bullet> x) Basis < 1"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7426
    apply safe
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7427
  proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7428
    fix i :: 'a
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7429
    assume i: "i \<in> Basis"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7430
    then show "0 < x \<bullet> i"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7431
      using as[THEN spec[where x="x - (e / 2) *\<^sub>R i"]] and \<open>e > 0\<close>
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7432
      unfolding dist_norm
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7433
      by (auto elim!: ballE[where x=i] simp: inner_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7434
  next
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7435
    have **: "dist x (x + (e / 2) *\<^sub>R (SOME i. i\<in>Basis)) < e" using \<open>e > 0\<close>
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
  7436
      unfolding dist_norm
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7437
      by (auto intro!: mult_strict_left_mono simp: SOME_Basis)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7438
    have "\<And>i. i \<in> Basis \<Longrightarrow> (x + (e / 2) *\<^sub>R (SOME i. i\<in>Basis)) \<bullet> i =
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7439
      x\<bullet>i + (if i = (SOME i. i\<in>Basis) then e/2 else 0)"
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
  7440
      by (auto simp: SOME_Basis inner_Basis inner_simps)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7441
    then have *: "setsum (op \<bullet> (x + (e / 2) *\<^sub>R (SOME i. i\<in>Basis))) Basis =
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7442
      setsum (\<lambda>i. x\<bullet>i + (if (SOME i. i\<in>Basis) = i then e/2 else 0)) Basis"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  7443
      apply (rule_tac setsum.cong)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7444
      apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7445
      done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7446
    have "setsum (op \<bullet> x) Basis < setsum (op \<bullet> (x + (e / 2) *\<^sub>R (SOME i. i\<in>Basis))) Basis"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  7447
      unfolding * setsum.distrib
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7448
      using \<open>e > 0\<close> DIM_positive[where 'a='a]
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  7449
      apply (subst setsum.delta')
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7450
      apply (auto simp: SOME_Basis)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7451
      done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7452
    also have "\<dots> \<le> 1"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7453
      using **
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7454
      apply (drule_tac as[rule_format])
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7455
      apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7456
      done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7457
    finally show "setsum (op \<bullet> x) Basis < 1" by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7458
  qed
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7459
next
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7460
  fix x :: 'a
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7461
  assume as: "\<forall>i\<in>Basis. 0 < x \<bullet> i" "setsum (op \<bullet> x) Basis < 1"
55697
abec82f4e3e9 tuned proofs;
wenzelm
parents: 54780
diff changeset
  7462
  obtain a :: 'b where "a \<in> UNIV" using UNIV_witness ..
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
  7463
  let ?d = "(1 - setsum (op \<bullet> x) Basis) / real (DIM('a))"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7464
  have "Min ((op \<bullet> x) ` Basis) > 0"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7465
    apply (rule Min_grI)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7466
    using as(1)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7467
    apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7468
    done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7469
  moreover have "?d > 0"
56541
0e3abadbef39 made divide_pos_pos a simp rule
nipkow
parents: 56536
diff changeset
  7470
    using as(2) by (auto simp: Suc_le_eq DIM_positive)
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
  7471
  ultimately show "\<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> (\<forall>i\<in>Basis. 0 \<le> y \<bullet> i) \<and> setsum (op \<bullet> y) Basis \<le> 1"
59807
22bc39064290 prefer local fixes;
wenzelm
parents: 58877
diff changeset
  7472
    apply (rule_tac x="min (Min ((op \<bullet> x) ` Basis)) D" for D in exI)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7473
    apply rule
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7474
    defer
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7475
    apply (rule, rule)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7476
  proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7477
    fix y
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7478
    assume y: "dist x y < min (Min (op \<bullet> x ` Basis)) ?d"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7479
    have "setsum (op \<bullet> y) Basis \<le> setsum (\<lambda>i. x\<bullet>i + ?d) Basis"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7480
    proof (rule setsum_mono)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7481
      fix i :: 'a
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7482
      assume i: "i \<in> Basis"
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61880
diff changeset
  7483
      then have "\<bar>y\<bullet>i - x\<bullet>i\<bar> < ?d"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7484
        apply -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7485
        apply (rule le_less_trans)
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
  7486
        using Basis_le_norm[OF i, of "y - x"]
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7487
        using y[unfolded min_less_iff_conj dist_norm, THEN conjunct2]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7488
        apply (auto simp add: norm_minus_commute inner_diff_left)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7489
        done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7490
      then show "y \<bullet> i \<le> x \<bullet> i + ?d" by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7491
    qed
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7492
    also have "\<dots> \<le> 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
  7493
      unfolding setsum.distrib setsum_constant
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7494
      by (auto simp add: Suc_le_eq)
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
  7495
    finally show "(\<forall>i\<in>Basis. 0 \<le> y \<bullet> i) \<and> setsum (op \<bullet> y) Basis \<le> 1"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7496
    proof safe
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7497
      fix i :: 'a
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7498
      assume i: "i \<in> Basis"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7499
      have "norm (x - y) < x\<bullet>i"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7500
        apply (rule less_le_trans)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7501
        apply (rule y[unfolded min_less_iff_conj dist_norm, THEN conjunct1])
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7502
        using i
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7503
        apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7504
        done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7505
      then show "0 \<le> y\<bullet>i"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7506
        using Basis_le_norm[OF i, of "x - y"] and as(1)[rule_format, OF i]
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
  7507
        by (auto simp: inner_simps)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7508
    qed
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7509
  qed auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7510
qed
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7511
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7512
lemma interior_std_simplex_nonempty:
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7513
  obtains a :: "'a::euclidean_space" where
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7514
    "a \<in> interior(convex hull (insert 0 Basis))"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7515
proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7516
  let ?D = "Basis :: 'a set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7517
  let ?a = "setsum (\<lambda>b::'a. inverse (2 * real DIM('a)) *\<^sub>R b) Basis"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7518
  {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7519
    fix i :: 'a
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7520
    assume i: "i \<in> Basis"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7521
    have "?a \<bullet> i = inverse (2 * real DIM('a))"
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
  7522
      by (rule trans[of _ "setsum (\<lambda>j. if i = j then inverse (2 * real DIM('a)) else 0) ?D"])
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  7523
         (simp_all add: setsum.If_cases i) }
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
  7524
  note ** = this
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7525
  show ?thesis
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7526
    apply (rule that[of ?a])
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7527
    unfolding interior_std_simplex mem_Collect_eq
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7528
  proof safe
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7529
    fix i :: 'a
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7530
    assume i: "i \<in> Basis"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7531
    show "0 < ?a \<bullet> i"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7532
      unfolding **[OF i] by (auto simp add: Suc_le_eq DIM_positive)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7533
  next
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7534
    have "setsum (op \<bullet> ?a) ?D = setsum (\<lambda>i. inverse (2 * real DIM('a))) ?D"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  7535
      apply (rule setsum.cong)
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  7536
      apply rule
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7537
      apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7538
      done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7539
    also have "\<dots> < 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
  7540
      unfolding setsum_constant divide_inverse[symmetric]
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7541
      by (auto simp add: field_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7542
    finally show "setsum (op \<bullet> ?a) ?D < 1" by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7543
  qed
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7544
qed
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7545
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7546
lemma rel_interior_substd_simplex:
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7547
  assumes d: "d \<subseteq> Basis"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7548
  shows "rel_interior (convex hull (insert 0 d)) =
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7549
    {x::'a::euclidean_space. (\<forall>i\<in>d. 0 < x\<bullet>i) \<and> (\<Sum>i\<in>d. x\<bullet>i) < 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0)}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7550
  (is "rel_interior (convex hull (insert 0 ?p)) = ?s")
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7551
proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7552
  have "finite d"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7553
    apply (rule finite_subset)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7554
    using assms
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7555
    apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7556
    done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7557
  show ?thesis
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7558
  proof (cases "d = {}")
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7559
    case True
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7560
    then show ?thesis
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7561
      using rel_interior_sing using euclidean_eq_iff[of _ 0] by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7562
  next
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7563
    case False
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7564
    have h0: "affine hull (convex hull (insert 0 ?p)) =
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7565
      {x::'a::euclidean_space. (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0)}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7566
      using affine_hull_convex_hull affine_hull_substd_basis assms by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7567
    have aux: "\<And>x::'a. \<forall>i\<in>Basis. (\<forall>i\<in>d. 0 \<le> x\<bullet>i) \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0) \<longrightarrow> 0 \<le> x\<bullet>i"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7568
      by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7569
    {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7570
      fix x :: "'a::euclidean_space"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7571
      assume x: "x \<in> rel_interior (convex hull (insert 0 ?p))"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7572
      then obtain e where e0: "e > 0" and
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7573
        "ball x e \<inter> {xa. (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> xa\<bullet>i = 0)} \<subseteq> convex hull (insert 0 ?p)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7574
        using mem_rel_interior_ball[of x "convex hull (insert 0 ?p)"] h0 by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7575
      then have as: "\<forall>xa. dist x xa < e \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> xa\<bullet>i = 0) \<longrightarrow>
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7576
        (\<forall>i\<in>d. 0 \<le> xa \<bullet> i) \<and> setsum (op \<bullet> xa) d \<le> 1"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7577
        unfolding ball_def unfolding substd_simplex[OF assms] using assms by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7578
      have x0: "(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7579
        using x rel_interior_subset  substd_simplex[OF assms] by auto
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  7580
      have "(\<forall>i\<in>d. 0 < x \<bullet> i) \<and> setsum (op \<bullet> x) d < 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0)"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  7581
        apply rule
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  7582
        apply rule
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7583
      proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7584
        fix i :: 'a
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7585
        assume "i \<in> d"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7586
        then have "\<forall>ia\<in>d. 0 \<le> (x - (e / 2) *\<^sub>R i) \<bullet> ia"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7587
          apply -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7588
          apply (rule as[rule_format,THEN conjunct1])
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7589
          unfolding dist_norm
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7590
          using d \<open>e > 0\<close> x0
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7591
          apply (auto simp: inner_simps inner_Basis)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7592
          done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7593
        then show "0 < x \<bullet> i"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7594
          apply (erule_tac x=i in ballE)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7595
          using \<open>e > 0\<close> \<open>i \<in> d\<close> d
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7596
          apply (auto simp: inner_simps inner_Basis)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7597
          done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7598
      next
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7599
        obtain a where a: "a \<in> d"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7600
          using \<open>d \<noteq> {}\<close> by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7601
        then have **: "dist x (x + (e / 2) *\<^sub>R a) < e"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7602
          using \<open>e > 0\<close> norm_Basis[of a] d
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7603
          unfolding dist_norm
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7604
          by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7605
        have "\<And>i. i \<in> Basis \<Longrightarrow> (x + (e / 2) *\<^sub>R a) \<bullet> i = x\<bullet>i + (if i = a then e/2 else 0)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7606
          using a d by (auto simp: inner_simps inner_Basis)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7607
        then have *: "setsum (op \<bullet> (x + (e / 2) *\<^sub>R a)) d =
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7608
          setsum (\<lambda>i. x\<bullet>i + (if a = i then e/2 else 0)) d"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  7609
          using d by (intro setsum.cong) auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7610
        have "a \<in> Basis"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7611
          using \<open>a \<in> d\<close> d by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7612
        then have h1: "(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> (x + (e / 2) *\<^sub>R a) \<bullet> i = 0)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7613
          using x0 d \<open>a\<in>d\<close> by (auto simp add: inner_add_left inner_Basis)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7614
        have "setsum (op \<bullet> x) d < setsum (op \<bullet> (x + (e / 2) *\<^sub>R a)) d"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  7615
          unfolding * setsum.distrib
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7616
          using \<open>e > 0\<close> \<open>a \<in> d\<close>
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7617
          using \<open>finite d\<close>
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  7618
          by (auto simp add: setsum.delta')
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7619
        also have "\<dots> \<le> 1"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7620
          using ** h1 as[rule_format, of "x + (e / 2) *\<^sub>R a"]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7621
          by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7622
        finally show "setsum (op \<bullet> x) d < 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7623
          using x0 by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7624
      qed
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7625
    }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7626
    moreover
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7627
    {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7628
      fix x :: "'a::euclidean_space"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7629
      assume as: "x \<in> ?s"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7630
      have "\<forall>i. 0 < x\<bullet>i \<or> 0 = x\<bullet>i \<longrightarrow> 0 \<le> x\<bullet>i"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7631
        by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7632
      moreover have "\<forall>i. i \<in> d \<or> i \<notin> d" by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7633
      ultimately
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  7634
      have "\<forall>i. (\<forall>i\<in>d. 0 < x\<bullet>i) \<and> (\<forall>i. i \<notin> d \<longrightarrow> x\<bullet>i = 0) \<longrightarrow> 0 \<le> x\<bullet>i"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7635
        by metis
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7636
      then have h2: "x \<in> convex hull (insert 0 ?p)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7637
        using as assms
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7638
        unfolding substd_simplex[OF assms] by fastforce
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7639
      obtain a where a: "a \<in> d"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7640
        using \<open>d \<noteq> {}\<close> by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7641
      let ?d = "(1 - setsum (op \<bullet> x) d) / real (card d)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7642
      have "0 < card d" using \<open>d \<noteq> {}\<close> \<open>finite d\<close>
44466
0e5c27f07529 remove unnecessary lemma card_ge1
huffman
parents: 44465
diff changeset
  7643
        by (simp add: card_gt_0_iff)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7644
      have "Min ((op \<bullet> x) ` d) > 0"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7645
        using as \<open>d \<noteq> {}\<close> \<open>finite d\<close> by (simp add: Min_grI)
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7646
      moreover have "?d > 0" using as using \<open>0 < card d\<close> by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7647
      ultimately have h3: "min (Min ((op \<bullet> x) ` d)) ?d > 0"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7648
        by auto
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  7649
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7650
      have "x \<in> rel_interior (convex hull (insert 0 ?p))"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7651
        unfolding rel_interior_ball mem_Collect_eq h0
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7652
        apply (rule,rule h2)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7653
        unfolding substd_simplex[OF assms]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7654
        apply (rule_tac x="min (Min ((op \<bullet> x) ` d)) ?d" in exI)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7655
        apply (rule, rule h3)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7656
        apply safe
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7657
        unfolding mem_ball
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7658
      proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7659
        fix y :: 'a
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7660
        assume y: "dist x y < min (Min (op \<bullet> x ` d)) ?d"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7661
        assume y2: "\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> y\<bullet>i = 0"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7662
        have "setsum (op \<bullet> y) d \<le> setsum (\<lambda>i. x\<bullet>i + ?d) d"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7663
        proof (rule setsum_mono)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7664
          fix i
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7665
          assume "i \<in> d"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7666
          with d have i: "i \<in> Basis"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7667
            by auto
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61880
diff changeset
  7668
          have "\<bar>y\<bullet>i - x\<bullet>i\<bar> < ?d"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7669
            apply (rule le_less_trans)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7670
            using Basis_le_norm[OF i, of "y - x"]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7671
            using y[unfolded min_less_iff_conj dist_norm, THEN conjunct2]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7672
            apply (auto simp add: norm_minus_commute inner_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7673
            done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7674
          then show "y \<bullet> i \<le> x \<bullet> i + ?d" by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7675
        qed
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7676
        also have "\<dots> \<le> 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
  7677
          unfolding setsum.distrib setsum_constant  using \<open>0 < card d\<close>
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7678
          by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7679
        finally show "setsum (op \<bullet> y) d \<le> 1" .
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  7680
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7681
        fix i :: 'a
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7682
        assume i: "i \<in> Basis"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7683
        then show "0 \<le> y\<bullet>i"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7684
        proof (cases "i\<in>d")
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7685
          case True
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7686
          have "norm (x - y) < x\<bullet>i"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7687
            using y[unfolded min_less_iff_conj dist_norm, THEN conjunct1]
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7688
            using Min_gr_iff[of "op \<bullet> x ` d" "norm (x - y)"] \<open>0 < card d\<close> \<open>i:d\<close>
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7689
            by (simp add: card_gt_0_iff)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7690
          then show "0 \<le> y\<bullet>i"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7691
            using Basis_le_norm[OF i, of "x - y"] and as(1)[rule_format]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7692
            by (auto simp: inner_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7693
        qed (insert y2, auto)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7694
      qed
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7695
    }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7696
    ultimately have
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7697
      "\<And>x. x \<in> rel_interior (convex hull insert 0 d) \<longleftrightarrow>
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7698
        x \<in> {x. (\<forall>i\<in>d. 0 < x \<bullet> i) \<and> setsum (op \<bullet> x) d < 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7699
      by blast
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7700
    then show ?thesis by (rule set_eqI)
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7701
  qed
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7702
qed
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7703
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7704
lemma rel_interior_substd_simplex_nonempty:
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7705
  assumes "d \<noteq> {}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7706
    and "d \<subseteq> Basis"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7707
  obtains a :: "'a::euclidean_space"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7708
    where "a \<in> rel_interior (convex hull (insert 0 d))"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7709
proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7710
  let ?D = d
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7711
  let ?a = "setsum (\<lambda>b::'a::euclidean_space. inverse (2 * real (card d)) *\<^sub>R b) ?D"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7712
  have "finite d"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7713
    apply (rule finite_subset)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7714
    using assms(2)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7715
    apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7716
    done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7717
  then have d1: "0 < real (card d)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7718
    using \<open>d \<noteq> {}\<close> by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7719
  {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7720
    fix i
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7721
    assume "i \<in> d"
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
  7722
    have "?a \<bullet> i = inverse (2 * real (card d))"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7723
      apply (rule trans[of _ "setsum (\<lambda>j. if i = j then inverse (2 * real (card d)) else 0) ?D"])
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
  7724
      unfolding inner_setsum_left
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  7725
      apply (rule setsum.cong)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7726
      using \<open>i \<in> d\<close> \<open>finite d\<close> setsum.delta'[of d i "(\<lambda>k. inverse (2 * real (card d)))"]
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7727
        d1 assms(2)
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  7728
      by (auto simp: inner_Basis set_rev_mp[OF _ assms(2)])
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7729
  }
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7730
  note ** = this
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7731
  show ?thesis
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7732
    apply (rule that[of ?a])
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7733
    unfolding rel_interior_substd_simplex[OF assms(2)] mem_Collect_eq
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7734
  proof safe
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7735
    fix i
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7736
    assume "i \<in> d"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7737
    have "0 < inverse (2 * real (card d))"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7738
      using d1 by auto
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7739
    also have "\<dots> = ?a \<bullet> i" using **[of i] \<open>i \<in> d\<close>
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7740
      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
  7741
    finally show "0 < ?a \<bullet> i" by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7742
  next
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7743
    have "setsum (op \<bullet> ?a) ?D = setsum (\<lambda>i. inverse (2 * real (card d))) ?D"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  7744
      by (rule setsum.cong) (rule refl, rule **)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7745
    also have "\<dots> < 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
  7746
      unfolding setsum_constant divide_real_def[symmetric]
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7747
      by (auto simp add: field_simps)
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
  7748
    finally show "setsum (op \<bullet> ?a) ?D < 1" by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7749
  next
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7750
    fix i
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7751
    assume "i \<in> Basis" and "i \<notin> d"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7752
    have "?a \<in> span d"
56196
32b7eafc5a52 remove unnecessary finiteness assumptions from lemmas about setsum
huffman
parents: 56189
diff changeset
  7753
    proof (rule span_setsum[of d "(\<lambda>b. b /\<^sub>R (2 * real (card d)))" d])
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7754
      {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7755
        fix x :: "'a::euclidean_space"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7756
        assume "x \<in> d"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7757
        then have "x \<in> span d"
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
  7758
          using span_superset[of _ "d"] by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7759
        then have "x /\<^sub>R (2 * real (card d)) \<in> span d"
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
  7760
          using span_mul[of x "d" "(inverse (real (card d)) / 2)"] by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7761
      }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7762
      then show "\<forall>x\<in>d. x /\<^sub>R (2 * real (card d)) \<in> span d"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7763
        by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7764
    qed
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7765
    then show "?a \<bullet> i = 0 "
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7766
      using \<open>i \<notin> d\<close> unfolding span_substd_basis[OF assms(2)] using \<open>i \<in> Basis\<close> by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7767
  qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7768
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7769
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7770
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7771
subsection \<open>Relative interior of convex set\<close>
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7772
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  7773
lemma rel_interior_convex_nonempty_aux:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7774
  fixes S :: "'n::euclidean_space set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7775
  assumes "convex S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7776
    and "0 \<in> S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7777
  shows "rel_interior S \<noteq> {}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7778
proof (cases "S = {0}")
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7779
  case True
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7780
  then show ?thesis using rel_interior_sing by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7781
next
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7782
  case False
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7783
  obtain B where B: "independent B \<and> B \<le> S \<and> S \<le> span B \<and> card B = dim S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7784
    using basis_exists[of S] by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7785
  then have "B \<noteq> {}"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7786
    using B assms \<open>S \<noteq> {0}\<close> span_empty by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7787
  have "insert 0 B \<le> span B"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7788
    using subspace_span[of B] subspace_0[of "span B"] span_inc by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7789
  then have "span (insert 0 B) \<le> span B"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7790
    using span_span[of B] span_mono[of "insert 0 B" "span B"] by blast
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7791
  then have "convex hull insert 0 B \<le> span B"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7792
    using convex_hull_subset_span[of "insert 0 B"] by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7793
  then have "span (convex hull insert 0 B) \<le> span B"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7794
    using span_span[of B] span_mono[of "convex hull insert 0 B" "span B"] by blast
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7795
  then have *: "span (convex hull insert 0 B) = span B"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7796
    using span_mono[of B "convex hull insert 0 B"] hull_subset[of "insert 0 B"] by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7797
  then have "span (convex hull insert 0 B) = span S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7798
    using B span_mono[of B S] span_mono[of S "span B"] span_span[of B] by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7799
  moreover have "0 \<in> affine hull (convex hull insert 0 B)"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7800
    using hull_subset[of "convex hull insert 0 B"] hull_subset[of "insert 0 B"] by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7801
  ultimately have **: "affine hull (convex hull insert 0 B) = affine hull S"
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  7802
    using affine_hull_span_0[of "convex hull insert 0 B"] affine_hull_span_0[of "S"]
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7803
      assms hull_subset[of S]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7804
    by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7805
  obtain d and f :: "'n \<Rightarrow> 'n" where
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7806
    fd: "card d = card B" "linear f" "f ` B = d"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7807
      "f ` span B = {x. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = (0::real)} \<and> inj_on f (span B)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7808
    and d: "d \<subseteq> Basis"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7809
    using basis_to_substdbasis_subspace_isomorphism[of B,OF _ ] B by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7810
  then have "bounded_linear f"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7811
    using linear_conv_bounded_linear by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7812
  have "d \<noteq> {}"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7813
    using fd B \<open>B \<noteq> {}\<close> by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7814
  have "insert 0 d = f ` (insert 0 B)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7815
    using fd linear_0 by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7816
  then have "(convex hull (insert 0 d)) = f ` (convex hull (insert 0 B))"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7817
    using convex_hull_linear_image[of f "(insert 0 d)"]
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7818
      convex_hull_linear_image[of f "(insert 0 B)"] \<open>linear f\<close>
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7819
    by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7820
  moreover have "rel_interior (f ` (convex hull insert 0 B)) =
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7821
    f ` rel_interior (convex hull insert 0 B)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7822
    apply (rule  rel_interior_injective_on_span_linear_image[of f "(convex hull insert 0 B)"])
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7823
    using \<open>bounded_linear f\<close> fd *
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7824
    apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7825
    done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7826
  ultimately have "rel_interior (convex hull insert 0 B) \<noteq> {}"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7827
    using rel_interior_substd_simplex_nonempty[OF \<open>d \<noteq> {}\<close> d]
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7828
    apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7829
    apply blast
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7830
    done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7831
  moreover have "convex hull (insert 0 B) \<subseteq> S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7832
    using B assms hull_mono[of "insert 0 B" "S" "convex"] convex_hull_eq
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7833
    by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7834
  ultimately show ?thesis
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7835
    using subset_rel_interior[of "convex hull insert 0 B" S] ** by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7836
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7837
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61609
diff changeset
  7838
lemma rel_interior_eq_empty:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7839
  fixes S :: "'n::euclidean_space set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7840
  assumes "convex S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7841
  shows "rel_interior S = {} \<longleftrightarrow> S = {}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7842
proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7843
  {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7844
    assume "S \<noteq> {}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7845
    then obtain a where "a \<in> S" by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7846
    then have "0 \<in> op + (-a) ` S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7847
      using assms exI[of "(\<lambda>x. x \<in> S \<and> - a + x = 0)" a] by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7848
    then have "rel_interior (op + (-a) ` S) \<noteq> {}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7849
      using rel_interior_convex_nonempty_aux[of "op + (-a) ` S"]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7850
        convex_translation[of S "-a"] assms
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7851
      by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7852
    then have "rel_interior S \<noteq> {}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7853
      using rel_interior_translation by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7854
  }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7855
  then show ?thesis
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7856
    using rel_interior_empty by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7857
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7858
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7859
lemma convex_rel_interior:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7860
  fixes S :: "'n::euclidean_space set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7861
  assumes "convex S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7862
  shows "convex (rel_interior S)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7863
proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7864
  {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7865
    fix x y and u :: real
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7866
    assume assm: "x \<in> rel_interior S" "y \<in> rel_interior S" "0 \<le> u" "u \<le> 1"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7867
    then have "x \<in> S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7868
      using rel_interior_subset by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7869
    have "x - u *\<^sub>R (x-y) \<in> rel_interior S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7870
    proof (cases "0 = u")
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7871
      case False
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7872
      then have "0 < u" using assm by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7873
      then show ?thesis
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7874
        using assm rel_interior_convex_shrink[of S y x u] assms \<open>x \<in> S\<close> by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7875
    next
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7876
      case True
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7877
      then show ?thesis using assm by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7878
    qed
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7879
    then have "(1 - u) *\<^sub>R x + u *\<^sub>R y \<in> rel_interior S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7880
      by (simp add: algebra_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7881
  }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7882
  then show ?thesis
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7883
    unfolding convex_alt by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7884
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7885
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  7886
lemma convex_closure_rel_interior:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7887
  fixes S :: "'n::euclidean_space set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7888
  assumes "convex S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7889
  shows "closure (rel_interior S) = closure S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7890
proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7891
  have h1: "closure (rel_interior S) \<le> closure S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7892
    using closure_mono[of "rel_interior S" S] rel_interior_subset[of S] by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7893
  show ?thesis
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7894
  proof (cases "S = {}")
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7895
    case False
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7896
    then obtain a where a: "a \<in> rel_interior S"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61609
diff changeset
  7897
      using rel_interior_eq_empty assms by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7898
    { fix x
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7899
      assume x: "x \<in> closure S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7900
      {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7901
        assume "x = a"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7902
        then have "x \<in> closure (rel_interior S)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7903
          using a unfolding closure_def by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7904
      }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7905
      moreover
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7906
      {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7907
        assume "x \<noteq> a"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7908
         {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7909
           fix e :: real
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7910
           assume "e > 0"
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  7911
           define e1 where "e1 = min 1 (e/norm (x - a))"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7912
           then have e1: "e1 > 0" "e1 \<le> 1" "e1 * norm (x - a) \<le> e"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7913
             using \<open>x \<noteq> a\<close> \<open>e > 0\<close> le_divide_eq[of e1 e "norm (x - a)"]
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7914
             by simp_all
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7915
           then have *: "x - e1 *\<^sub>R (x - a) : rel_interior S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7916
             using rel_interior_closure_convex_shrink[of S a x e1] assms x a e1_def
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7917
             by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7918
           have "\<exists>y. y \<in> rel_interior S \<and> y \<noteq> x \<and> dist y x \<le> e"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7919
              apply (rule_tac x="x - e1 *\<^sub>R (x - a)" in exI)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  7920
              using * e1 dist_norm[of "x - e1 *\<^sub>R (x - a)" x] \<open>x \<noteq> a\<close>
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7921
              apply simp
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7922
              done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7923
        }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7924
        then have "x islimpt rel_interior S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7925
          unfolding islimpt_approachable_le by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7926
        then have "x \<in> closure(rel_interior S)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7927
          unfolding closure_def by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7928
      }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7929
      ultimately have "x \<in> closure(rel_interior S)" by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7930
    }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7931
    then show ?thesis using h1 by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7932
  next
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7933
    case True
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7934
    then have "rel_interior S = {}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7935
      using rel_interior_empty by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7936
    then have "closure (rel_interior S) = {}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7937
      using closure_empty by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7938
    with True show ?thesis by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7939
  qed
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7940
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7941
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7942
lemma rel_interior_same_affine_hull:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7943
  fixes S :: "'n::euclidean_space set"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7944
  assumes "convex S"
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7945
  shows "affine hull (rel_interior S) = affine hull S"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7946
  by (metis assms closure_same_affine_hull convex_closure_rel_interior)
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7947
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  7948
lemma rel_interior_aff_dim:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7949
  fixes S :: "'n::euclidean_space set"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7950
  assumes "convex S"
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7951
  shows "aff_dim (rel_interior S) = aff_dim S"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7952
  by (metis aff_dim_affine_hull2 assms rel_interior_same_affine_hull)
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7953
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7954
lemma rel_interior_rel_interior:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7955
  fixes S :: "'n::euclidean_space set"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7956
  assumes "convex S"
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7957
  shows "rel_interior (rel_interior S) = rel_interior S"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7958
proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7959
  have "openin (subtopology euclidean (affine hull (rel_interior S))) (rel_interior S)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7960
    using opein_rel_interior[of S] rel_interior_same_affine_hull[of S] assms by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7961
  then show ?thesis
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7962
    using rel_interior_def by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7963
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7964
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7965
lemma rel_interior_rel_open:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7966
  fixes S :: "'n::euclidean_space set"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7967
  assumes "convex S"
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7968
  shows "rel_open (rel_interior S)"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7969
  unfolding rel_open_def using rel_interior_rel_interior assms by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7970
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7971
lemma convex_rel_interior_closure_aux:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7972
  fixes x y z :: "'n::euclidean_space"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7973
  assumes "0 < a" "0 < b" "(a + b) *\<^sub>R z = a *\<^sub>R x + b *\<^sub>R y"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7974
  obtains e where "0 < e" "e \<le> 1" "z = y - e *\<^sub>R (y - x)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7975
proof -
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  7976
  define e where "e = a / (a + b)"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7977
  have "z = (1 / (a + b)) *\<^sub>R ((a + b) *\<^sub>R z)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7978
    apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7979
    using assms
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7980
    apply simp
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7981
    done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7982
  also have "\<dots> = (1 / (a + b)) *\<^sub>R (a *\<^sub>R x + b *\<^sub>R y)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7983
    using assms scaleR_cancel_left[of "1/(a+b)" "(a + b) *\<^sub>R z" "a *\<^sub>R x + b *\<^sub>R y"]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7984
    by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7985
  also have "\<dots> = y - e *\<^sub>R (y-x)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7986
    using e_def
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7987
    apply (simp add: algebra_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7988
    using scaleR_left_distrib[of "a/(a+b)" "b/(a+b)" y] assms add_divide_distrib[of a b "a+b"]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7989
    apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7990
    done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7991
  finally have "z = y - e *\<^sub>R (y-x)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7992
    by auto
56541
0e3abadbef39 made divide_pos_pos a simp rule
nipkow
parents: 56536
diff changeset
  7993
  moreover have "e > 0" using e_def assms by auto
0e3abadbef39 made divide_pos_pos a simp rule
nipkow
parents: 56536
diff changeset
  7994
  moreover have "e \<le> 1" using e_def assms by auto
0e3abadbef39 made divide_pos_pos a simp rule
nipkow
parents: 56536
diff changeset
  7995
  ultimately show ?thesis using that[of e] by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7996
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  7997
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  7998
lemma convex_rel_interior_closure:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  7999
  fixes S :: "'n::euclidean_space set"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8000
  assumes "convex S"
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8001
  shows "rel_interior (closure S) = rel_interior S"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8002
proof (cases "S = {}")
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8003
  case True
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8004
  then show ?thesis
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61609
diff changeset
  8005
    using assms rel_interior_eq_empty by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8006
next
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8007
  case False
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8008
  have "rel_interior (closure S) \<supseteq> rel_interior S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8009
    using subset_rel_interior[of S "closure S"] closure_same_affine_hull closure_subset
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8010
    by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8011
  moreover
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8012
  {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8013
    fix z
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  8014
    assume z: "z \<in> rel_interior (closure S)"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8015
    obtain x where x: "x \<in> rel_interior S"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61609
diff changeset
  8016
      using \<open>S \<noteq> {}\<close> assms rel_interior_eq_empty by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8017
    have "z \<in> rel_interior S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8018
    proof (cases "x = z")
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8019
      case True
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8020
      then show ?thesis using x by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8021
    next
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8022
      case False
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  8023
      obtain e where e: "e > 0" "cball z e \<inter> affine hull closure S \<le> closure S"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8024
        using z rel_interior_cball[of "closure S"] by auto
56541
0e3abadbef39 made divide_pos_pos a simp rule
nipkow
parents: 56536
diff changeset
  8025
      hence *: "0 < e/norm(z-x)" using e False by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  8026
      define y where "y = z + (e/norm(z-x)) *\<^sub>R (z-x)"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8027
      have yball: "y \<in> cball z e"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8028
        using mem_cball y_def dist_norm[of z y] e by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8029
      have "x \<in> affine hull closure S"
62948
7700f467892b lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 62843
diff changeset
  8030
        using x rel_interior_subset_closure hull_inc[of x "closure S"] by blast
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8031
      moreover have "z \<in> affine hull closure S"
62948
7700f467892b lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 62843
diff changeset
  8032
        using z rel_interior_subset hull_subset[of "closure S"] by blast
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8033
      ultimately have "y \<in> affine hull closure S"
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  8034
        using y_def affine_affine_hull[of "closure S"]
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8035
          mem_affine_3_minus [of "affine hull closure S" z z x "e/norm(z-x)"] by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8036
      then have "y \<in> closure S" using e yball by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8037
      have "(1 + (e/norm(z-x))) *\<^sub>R z = (e/norm(z-x)) *\<^sub>R x + y"
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  8038
        using y_def by (simp add: algebra_simps)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8039
      then obtain e1 where "0 < e1" "e1 \<le> 1" "z = y - e1 *\<^sub>R (y - x)"
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  8040
        using * convex_rel_interior_closure_aux[of "e / norm (z - x)" 1 z x y]
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8041
        by (auto simp add: algebra_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8042
      then show ?thesis
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  8043
        using rel_interior_closure_convex_shrink assms x \<open>y \<in> closure S\<close>
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8044
        by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8045
    qed
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8046
  }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8047
  ultimately show ?thesis by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8048
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8049
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  8050
lemma convex_interior_closure:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8051
  fixes S :: "'n::euclidean_space set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8052
  assumes "convex S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8053
  shows "interior (closure S) = interior S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8054
  using closure_aff_dim[of S] interior_rel_interior_gen[of S]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8055
    interior_rel_interior_gen[of "closure S"]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8056
    convex_rel_interior_closure[of S] assms
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8057
  by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8058
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8059
lemma closure_eq_rel_interior_eq:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8060
  fixes S1 S2 :: "'n::euclidean_space set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8061
  assumes "convex S1"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8062
    and "convex S2"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8063
  shows "closure S1 = closure S2 \<longleftrightarrow> rel_interior S1 = rel_interior S2"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8064
  by (metis convex_rel_interior_closure convex_closure_rel_interior assms)
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8065
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8066
lemma closure_eq_between:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8067
  fixes S1 S2 :: "'n::euclidean_space set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8068
  assumes "convex S1"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8069
    and "convex S2"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8070
  shows "closure S1 = closure S2 \<longleftrightarrow> rel_interior S1 \<le> S2 \<and> S2 \<subseteq> closure S1"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  8071
  (is "?A \<longleftrightarrow> ?B")
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8072
proof
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8073
  assume ?A
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8074
  then show ?B
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8075
    by (metis assms closure_subset convex_rel_interior_closure rel_interior_subset)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8076
next
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8077
  assume ?B
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8078
  then have "closure S1 \<subseteq> closure S2"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8079
    by (metis assms(1) convex_closure_rel_interior closure_mono)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  8080
  moreover from \<open>?B\<close> have "closure S1 \<supseteq> closure S2"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8081
    by (metis closed_closure closure_minimal)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8082
  ultimately show ?A ..
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8083
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8084
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8085
lemma open_inter_closure_rel_interior:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8086
  fixes S A :: "'n::euclidean_space set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8087
  assumes "convex S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8088
    and "open A"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8089
  shows "A \<inter> closure S = {} \<longleftrightarrow> A \<inter> rel_interior S = {}"
62843
313d3b697c9a Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents: 62626
diff changeset
  8090
  by (metis assms convex_closure_rel_interior open_Int_closure_eq_empty)
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8091
62620
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8092
lemma rel_interior_open_segment:
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8093
  fixes a :: "'a :: euclidean_space"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8094
  shows "rel_interior(open_segment a b) = open_segment a b"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8095
proof (cases "a = b")
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8096
  case True then show ?thesis by auto
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8097
next
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8098
  case False then show ?thesis
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8099
    apply (simp add: rel_interior_eq openin_open)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8100
    apply (rule_tac x="ball (inverse 2 *\<^sub>R (a + b)) (norm(b - a) / 2)" in exI)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8101
    apply (simp add: open_segment_as_ball)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8102
    done
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8103
qed
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8104
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8105
lemma rel_interior_closed_segment:
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8106
  fixes a :: "'a :: euclidean_space"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8107
  shows "rel_interior(closed_segment a b) =
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8108
         (if a = b then {a} else open_segment a b)"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8109
proof (cases "a = b")
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8110
  case True then show ?thesis by auto
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8111
next
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8112
  case False then show ?thesis
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8113
    by simp
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8114
       (metis closure_open_segment convex_open_segment convex_rel_interior_closure
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8115
              rel_interior_open_segment)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8116
qed
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8117
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8118
lemmas rel_interior_segment = rel_interior_closed_segment rel_interior_open_segment
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8119
62626
de25474ce728 Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
  8120
lemma starlike_convex_tweak_boundary_points:
de25474ce728 Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
  8121
  fixes S :: "'a::euclidean_space set"
de25474ce728 Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
  8122
  assumes "convex S" "S \<noteq> {}" and ST: "rel_interior S \<subseteq> T" and TS: "T \<subseteq> closure S"
de25474ce728 Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
  8123
  shows "starlike T"
de25474ce728 Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
  8124
proof -
de25474ce728 Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
  8125
  have "rel_interior S \<noteq> {}"
de25474ce728 Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
  8126
    by (simp add: assms rel_interior_eq_empty)
de25474ce728 Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
  8127
  then obtain a where a: "a \<in> rel_interior S"  by blast 
de25474ce728 Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
  8128
  with ST have "a \<in> T"  by blast 
de25474ce728 Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
  8129
  have *: "\<And>x. x \<in> T \<Longrightarrow> open_segment a x \<subseteq> rel_interior S"
de25474ce728 Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
  8130
    apply (rule rel_interior_closure_convex_segment [OF \<open>convex S\<close> a])
de25474ce728 Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
  8131
    using assms by blast
de25474ce728 Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
  8132
  show ?thesis
de25474ce728 Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
  8133
    unfolding starlike_def
de25474ce728 Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
  8134
    apply (rule bexI [OF _ \<open>a \<in> T\<close>])
de25474ce728 Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
  8135
    apply (simp add: closed_segment_eq_open)
de25474ce728 Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
  8136
    apply (intro conjI ballI a \<open>a \<in> T\<close> rel_interior_closure_convex_segment [OF \<open>convex S\<close> a])
de25474ce728 Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
  8137
    apply (simp add: order_trans [OF * ST])
de25474ce728 Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
  8138
    done
de25474ce728 Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
  8139
qed
de25474ce728 Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
  8140
de25474ce728 Contractible sets. Also removal of obsolete theorems and refactoring
paulson <lp15@cam.ac.uk>
parents: 62620
diff changeset
  8141
subsection\<open>The relative frontier of a set\<close>
62620
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8142
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8143
definition "rel_frontier S = closure S - rel_interior S"
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8144
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8145
lemma closed_affine_hull:
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8146
  fixes S :: "'n::euclidean_space set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8147
  shows "closed (affine hull S)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8148
  by (metis affine_affine_hull affine_closed)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8149
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  8150
lemma rel_frontier_nonempty_interior:
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  8151
  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
  8152
  shows "interior S \<noteq> {} \<Longrightarrow> rel_frontier S = frontier S"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  8153
by (metis frontier_def interior_rel_interior_gen rel_frontier_def)
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  8154
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  8155
lemma rel_frontier_frontier:
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  8156
  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
  8157
  shows "affine hull S = UNIV \<Longrightarrow> rel_frontier S = frontier S"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  8158
by (simp add: frontier_def rel_frontier_def rel_interior_interior)
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  8159
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8160
lemma closed_rel_frontier:
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8161
  fixes S :: "'n::euclidean_space set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8162
  shows "closed (rel_frontier S)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8163
proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8164
  have *: "closedin (subtopology euclidean (affine hull S)) (closure S - rel_interior S)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8165
    apply (rule closedin_diff[of "subtopology euclidean (affine hull S)""closure S" "rel_interior S"])
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8166
    using closed_closedin_trans[of "affine hull S" "closure S"] closed_affine_hull[of S]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8167
      closure_affine_hull[of S] opein_rel_interior[of S]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8168
    apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8169
    done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8170
  show ?thesis
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8171
    apply (rule closedin_closed_trans[of "affine hull S" "rel_frontier S"])
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8172
    unfolding rel_frontier_def
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8173
    using * closed_affine_hull
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8174
    apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8175
    done
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  8176
qed
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  8177
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8178
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8179
lemma convex_rel_frontier_aff_dim:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8180
  fixes S1 S2 :: "'n::euclidean_space set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8181
  assumes "convex S1"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8182
    and "convex S2"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8183
    and "S2 \<noteq> {}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8184
    and "S1 \<le> rel_frontier S2"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8185
  shows "aff_dim S1 < aff_dim S2"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8186
proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8187
  have "S1 \<subseteq> closure S2"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8188
    using assms unfolding rel_frontier_def by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8189
  then have *: "affine hull S1 \<subseteq> affine hull S2"
62948
7700f467892b lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 62843
diff changeset
  8190
    using hull_mono[of "S1" "closure S2"] closure_same_affine_hull[of S2] by blast
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8191
  then have "aff_dim S1 \<le> aff_dim S2"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8192
    using * aff_dim_affine_hull[of S1] aff_dim_affine_hull[of S2]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8193
      aff_dim_subset[of "affine hull S1" "affine hull S2"]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8194
    by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8195
  moreover
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8196
  {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8197
    assume eq: "aff_dim S1 = aff_dim S2"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8198
    then have "S1 \<noteq> {}"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  8199
      using aff_dim_empty[of S1] aff_dim_empty[of S2] \<open>S2 \<noteq> {}\<close> by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8200
    have **: "affine hull S1 = affine hull S2"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8201
       apply (rule affine_dim_equal)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8202
       using * affine_affine_hull
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8203
       apply auto
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  8204
       using \<open>S1 \<noteq> {}\<close> hull_subset[of S1]
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8205
       apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8206
       using eq aff_dim_affine_hull[of S1] aff_dim_affine_hull[of S2]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8207
       apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8208
       done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8209
    obtain a where a: "a \<in> rel_interior S1"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61609
diff changeset
  8210
      using \<open>S1 \<noteq> {}\<close> rel_interior_eq_empty assms by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8211
    obtain T where T: "open T" "a \<in> T \<inter> S1" "T \<inter> affine hull S1 \<subseteq> S1"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8212
       using mem_rel_interior[of a S1] a by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8213
    then have "a \<in> T \<inter> closure S2"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8214
      using a assms unfolding rel_frontier_def by auto
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  8215
    then obtain b where b: "b \<in> T \<inter> rel_interior S2"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8216
      using open_inter_closure_rel_interior[of S2 T] assms T by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8217
    then have "b \<in> affine hull S1"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8218
      using rel_interior_subset hull_subset[of S2] ** by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8219
    then have "b \<in> S1"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8220
      using T b by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8221
    then have False
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8222
      using b assms unfolding rel_frontier_def by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8223
  }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8224
  ultimately show ?thesis
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8225
    using less_le by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8226
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8227
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8228
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8229
lemma convex_rel_interior_if:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8230
  fixes S ::  "'n::euclidean_space set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8231
  assumes "convex S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8232
    and "z \<in> rel_interior S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8233
  shows "\<forall>x\<in>affine hull S. \<exists>m. m > 1 \<and> (\<forall>e. e > 1 \<and> e \<le> m \<longrightarrow> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8234
proof -
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  8235
  obtain e1 where e1: "e1 > 0 \<and> cball z e1 \<inter> affine hull S \<subseteq> S"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  8236
    using mem_rel_interior_cball[of z S] assms by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8237
  {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8238
    fix x
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8239
    assume x: "x \<in> affine hull S"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  8240
    {
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  8241
      assume "x \<noteq> z"
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  8242
      define m where "m = 1 + e1/norm(x-z)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  8243
      hence "m > 1" using e1 \<open>x \<noteq> z\<close> by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8244
      {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8245
        fix e
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8246
        assume e: "e > 1 \<and> e \<le> m"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8247
        have "z \<in> affine hull S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8248
          using assms rel_interior_subset hull_subset[of S] by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8249
        then have *: "(1 - e)*\<^sub>R x + e *\<^sub>R z \<in> affine hull S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8250
          using mem_affine[of "affine hull S" x z "(1-e)" e] affine_affine_hull[of S] x
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8251
          by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8252
        have "norm (z + e *\<^sub>R x - (x + e *\<^sub>R z)) = norm ((e - 1) *\<^sub>R (x - z))"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8253
          by (simp add: algebra_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8254
        also have "\<dots> = (e - 1) * norm (x-z)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8255
          using norm_scaleR e by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8256
        also have "\<dots> \<le> (m - 1) * norm (x - z)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8257
          using e mult_right_mono[of _ _ "norm(x-z)"] by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8258
        also have "\<dots> = (e1 / norm (x - z)) * norm (x - z)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8259
          using m_def by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8260
        also have "\<dots> = e1"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  8261
          using \<open>x \<noteq> z\<close> e1 by simp
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8262
        finally have **: "norm (z + e *\<^sub>R x - (x + e *\<^sub>R z)) \<le> e1"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8263
          by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8264
        have "(1 - e)*\<^sub>R x+ e *\<^sub>R z \<in> cball z e1"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8265
          using m_def **
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8266
          unfolding cball_def dist_norm
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8267
          by (auto simp add: algebra_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8268
        then have "(1 - e) *\<^sub>R x+ e *\<^sub>R z \<in> S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8269
          using e * e1 by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8270
      }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8271
      then have "\<exists>m. m > 1 \<and> (\<forall>e. e > 1 \<and> e \<le> m \<longrightarrow> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S )"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  8272
        using \<open>m> 1 \<close> by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8273
    }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8274
    moreover
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8275
    {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8276
      assume "x = z"
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  8277
      define m where "m = 1 + e1"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8278
      then have "m > 1"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8279
        using e1 by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8280
      {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8281
        fix e
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8282
        assume e: "e > 1 \<and> e \<le> m"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8283
        then have "(1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  8284
          using e1 x \<open>x = z\<close> by (auto simp add: algebra_simps)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8285
        then have "(1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8286
          using e by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8287
      }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8288
      then have "\<exists>m. m > 1 \<and> (\<forall>e. e > 1 \<and> e \<le> m \<longrightarrow> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  8289
        using \<open>m > 1\<close> by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8290
    }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8291
    ultimately have "\<exists>m. m > 1 \<and> (\<forall>e. e > 1 \<and> e \<le> m \<longrightarrow> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S )"
62948
7700f467892b lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 62843
diff changeset
  8292
      by blast
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8293
  }
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8294
  then show ?thesis by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8295
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8296
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8297
lemma convex_rel_interior_if2:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8298
  fixes S :: "'n::euclidean_space set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8299
  assumes "convex S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8300
  assumes "z \<in> rel_interior S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8301
  shows "\<forall>x\<in>affine hull S. \<exists>e. e > 1 \<and> (1 - e)*\<^sub>R x + e *\<^sub>R z \<in> S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8302
  using convex_rel_interior_if[of S z] assms by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8303
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8304
lemma convex_rel_interior_only_if:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8305
  fixes S :: "'n::euclidean_space set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8306
  assumes "convex S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8307
    and "S \<noteq> {}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8308
  assumes "\<forall>x\<in>S. \<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8309
  shows "z \<in> rel_interior S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8310
proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8311
  obtain x where x: "x \<in> rel_interior S"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61609
diff changeset
  8312
    using rel_interior_eq_empty assms by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8313
  then have "x \<in> S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8314
    using rel_interior_subset by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8315
  then obtain e where e: "e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8316
    using assms by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  8317
  define y where [abs_def]: "y = (1 - e) *\<^sub>R x + e *\<^sub>R z"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8318
  then have "y \<in> S" using e by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  8319
  define e1 where "e1 = 1/e"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8320
  then have "0 < e1 \<and> e1 < 1" using e by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8321
  then have "z  =y - (1 - e1) *\<^sub>R (y - x)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8322
    using e1_def y_def by (auto simp add: algebra_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8323
  then show ?thesis
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  8324
    using rel_interior_convex_shrink[of S x y "1-e1"] \<open>0 < e1 \<and> e1 < 1\<close> \<open>y \<in> S\<close> x assms
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8325
    by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8326
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8327
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8328
lemma convex_rel_interior_iff:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8329
  fixes S :: "'n::euclidean_space set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8330
  assumes "convex S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8331
    and "S \<noteq> {}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8332
  shows "z \<in> rel_interior S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8333
  using assms hull_subset[of S "affine"]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8334
    convex_rel_interior_if[of S z] convex_rel_interior_only_if[of S z]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8335
  by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8336
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8337
lemma convex_rel_interior_iff2:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8338
  fixes S :: "'n::euclidean_space set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8339
  assumes "convex S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8340
    and "S \<noteq> {}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8341
  shows "z \<in> rel_interior S \<longleftrightarrow> (\<forall>x\<in>affine hull S. \<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8342
  using assms hull_subset[of S] convex_rel_interior_if2[of S z] convex_rel_interior_only_if[of S z]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8343
  by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8344
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8345
lemma convex_interior_iff:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8346
  fixes S :: "'n::euclidean_space set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8347
  assumes "convex S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8348
  shows "z \<in> interior S \<longleftrightarrow> (\<forall>x. \<exists>e. e > 0 \<and> z + e *\<^sub>R x \<in> S)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8349
proof (cases "aff_dim S = int DIM('n)")
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8350
  case False
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8351
  {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8352
    assume "z \<in> interior S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8353
    then have False
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8354
      using False interior_rel_interior_gen[of S] by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8355
  }
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8356
  moreover
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8357
  {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8358
    assume r: "\<forall>x. \<exists>e. e > 0 \<and> z + e *\<^sub>R x \<in> S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8359
    {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8360
      fix x
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  8361
      obtain e1 where e1: "e1 > 0 \<and> z + e1 *\<^sub>R (x - z) \<in> S"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8362
        using r by auto
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  8363
      obtain e2 where e2: "e2 > 0 \<and> z + e2 *\<^sub>R (z - x) \<in> S"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8364
        using r by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  8365
      define x1 where [abs_def]: "x1 = z + e1 *\<^sub>R (x - z)"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8366
      then have x1: "x1 \<in> affine hull S"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  8367
        using e1 hull_subset[of S] by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  8368
      define x2 where [abs_def]: "x2 = z + e2 *\<^sub>R (z - x)"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8369
      then have x2: "x2 \<in> affine hull S"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  8370
        using e2 hull_subset[of S] by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8371
      have *: "e1/(e1+e2) + e2/(e1+e2) = 1"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  8372
        using add_divide_distrib[of e1 e2 "e1+e2"] e1 e2 by simp
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8373
      then have "z = (e2/(e1+e2)) *\<^sub>R x1 + (e1/(e1+e2)) *\<^sub>R x2"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8374
        using x1_def x2_def
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8375
        apply (auto simp add: algebra_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8376
        using scaleR_left_distrib[of "e1/(e1+e2)" "e2/(e1+e2)" z]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8377
        apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8378
        done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8379
      then have z: "z \<in> affine hull S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8380
        using mem_affine[of "affine hull S" x1 x2 "e2/(e1+e2)" "e1/(e1+e2)"]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8381
          x1 x2 affine_affine_hull[of S] *
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8382
        by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8383
      have "x1 - x2 = (e1 + e2) *\<^sub>R (x - z)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8384
        using x1_def x2_def by (auto simp add: algebra_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8385
      then have "x = z+(1/(e1+e2)) *\<^sub>R (x1-x2)"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  8386
        using e1 e2 by simp
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8387
      then have "x \<in> affine hull S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8388
        using mem_affine_3_minus[of "affine hull S" z x1 x2 "1/(e1+e2)"]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8389
          x1 x2 z affine_affine_hull[of S]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8390
        by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8391
    }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8392
    then have "affine hull S = UNIV"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8393
      by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8394
    then have "aff_dim S = int DIM('n)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8395
      using aff_dim_affine_hull[of S] by (simp add: aff_dim_univ)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8396
    then have False
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8397
      using False by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8398
  }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8399
  ultimately show ?thesis by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8400
next
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8401
  case True
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8402
  then have "S \<noteq> {}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8403
    using aff_dim_empty[of S] by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8404
  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
  8405
    using True affine_hull_UNIV by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8406
  {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8407
    assume "z \<in> interior S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8408
    then have "z \<in> rel_interior S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8409
      using True interior_rel_interior_gen[of S] by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8410
    then have **: "\<forall>x. \<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  8411
      using convex_rel_interior_iff2[of S z] assms \<open>S \<noteq> {}\<close> * by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8412
    fix x
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8413
    obtain e1 where e1: "e1 > 1" "(1 - e1) *\<^sub>R (z - x) + e1 *\<^sub>R z \<in> S"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8414
      using **[rule_format, of "z-x"] by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  8415
    define e where [abs_def]: "e = e1 - 1"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8416
    then have "(1 - e1) *\<^sub>R (z - x) + e1 *\<^sub>R z = z + e *\<^sub>R x"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8417
      by (simp add: algebra_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8418
    then have "e > 0" "z + e *\<^sub>R x \<in> S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8419
      using e1 e_def by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8420
    then have "\<exists>e. e > 0 \<and> z + e *\<^sub>R x \<in> S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8421
      by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8422
  }
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8423
  moreover
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8424
  {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8425
    assume r: "\<forall>x. \<exists>e. e > 0 \<and> z + e *\<^sub>R x \<in> S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8426
    {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8427
      fix x
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8428
      obtain e1 where e1: "e1 > 0" "z + e1 *\<^sub>R (z - x) \<in> S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8429
        using r[rule_format, of "z-x"] by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  8430
      define e where "e = e1 + 1"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8431
      then have "z + e1 *\<^sub>R (z - x) = (1 - e) *\<^sub>R x + e *\<^sub>R z"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8432
        by (simp add: algebra_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8433
      then have "e > 1" "(1 - e)*\<^sub>R x + e *\<^sub>R z \<in> S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8434
        using e1 e_def by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8435
      then have "\<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S" by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8436
    }
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8437
    then have "z \<in> rel_interior S"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  8438
      using convex_rel_interior_iff2[of S z] assms \<open>S \<noteq> {}\<close> by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8439
    then have "z \<in> interior S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8440
      using True interior_rel_interior_gen[of S] by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8441
  }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8442
  ultimately show ?thesis by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8443
qed
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8444
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8445
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  8446
subsubsection \<open>Relative interior and closure under common operations\<close>
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8447
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8448
lemma rel_interior_inter_aux: "\<Inter>{rel_interior S |S. S : I} \<subseteq> \<Inter>I"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8449
proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8450
  {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8451
    fix y
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8452
    assume "y \<in> \<Inter>{rel_interior S |S. S : I}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8453
    then have y: "\<forall>S \<in> I. y \<in> rel_interior S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8454
      by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8455
    {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8456
      fix S
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8457
      assume "S \<in> I"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8458
      then have "y \<in> S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8459
        using rel_interior_subset y by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8460
    }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8461
    then have "y \<in> \<Inter>I" by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8462
  }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8463
  then show ?thesis by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8464
qed
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8465
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8466
lemma closure_inter: "closure (\<Inter>I) \<le> \<Inter>{closure S |S. S \<in> I}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8467
proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8468
  {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8469
    fix y
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8470
    assume "y \<in> \<Inter>I"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8471
    then have y: "\<forall>S \<in> I. y \<in> S" by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8472
    {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8473
      fix S
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8474
      assume "S \<in> I"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8475
      then have "y \<in> closure S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8476
        using closure_subset y by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8477
    }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8478
    then have "y \<in> \<Inter>{closure S |S. S \<in> I}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8479
      by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8480
  }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8481
  then have "\<Inter>I \<subseteq> \<Inter>{closure S |S. S \<in> I}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8482
    by auto
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  8483
  moreover have "closed (\<Inter>{closure S |S. S \<in> I})"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8484
    unfolding closed_Inter closed_closure by auto
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  8485
  ultimately show ?thesis using closure_hull[of "\<Inter>I"]
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8486
    hull_minimal[of "\<Inter>I" "\<Inter>{closure S |S. S \<in> I}" "closed"] by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8487
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8488
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  8489
lemma convex_closure_rel_interior_inter:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8490
  assumes "\<forall>S\<in>I. convex (S :: 'n::euclidean_space set)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8491
    and "\<Inter>{rel_interior S |S. S \<in> I} \<noteq> {}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8492
  shows "\<Inter>{closure S |S. S \<in> I} \<le> closure (\<Inter>{rel_interior S |S. S \<in> I})"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8493
proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8494
  obtain x where x: "\<forall>S\<in>I. x \<in> rel_interior S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8495
    using assms by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8496
  {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8497
    fix y
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8498
    assume "y \<in> \<Inter>{closure S |S. S \<in> I}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8499
    then have y: "\<forall>S \<in> I. y \<in> closure S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8500
      by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8501
    {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8502
      assume "y = x"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8503
      then have "y \<in> closure (\<Inter>{rel_interior S |S. S \<in> I})"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  8504
        using x closure_subset[of "\<Inter>{rel_interior S |S. S \<in> I}"] by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8505
    }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8506
    moreover
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8507
    {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8508
      assume "y \<noteq> x"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8509
      { fix e :: real
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8510
        assume e: "e > 0"
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  8511
        define e1 where "e1 = min 1 (e/norm (y - x))"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8512
        then have e1: "e1 > 0" "e1 \<le> 1" "e1 * norm (y - x) \<le> e"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  8513
          using \<open>y \<noteq> x\<close> \<open>e > 0\<close> le_divide_eq[of e1 e "norm (y - x)"]
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8514
          by simp_all
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  8515
        define z where "z = y - e1 *\<^sub>R (y - x)"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8516
        {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8517
          fix S
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8518
          assume "S \<in> I"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8519
          then have "z \<in> rel_interior S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8520
            using rel_interior_closure_convex_shrink[of S x y e1] assms x y e1 z_def
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8521
            by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8522
        }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8523
        then have *: "z \<in> \<Inter>{rel_interior S |S. S \<in> I}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8524
          by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8525
        have "\<exists>z. z \<in> \<Inter>{rel_interior S |S. S \<in> I} \<and> z \<noteq> y \<and> dist z y \<le> e"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8526
          apply (rule_tac x="z" in exI)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  8527
          using \<open>y \<noteq> x\<close> z_def * e1 e dist_norm[of z y]
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8528
          apply simp
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8529
          done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8530
      }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8531
      then have "y islimpt \<Inter>{rel_interior S |S. S \<in> I}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8532
        unfolding islimpt_approachable_le by blast
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8533
      then have "y \<in> closure (\<Inter>{rel_interior S |S. S \<in> I})"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8534
        unfolding closure_def by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8535
    }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8536
    ultimately have "y \<in> closure (\<Inter>{rel_interior S |S. S \<in> I})"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8537
      by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8538
  }
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8539
  then show ?thesis by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8540
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8541
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  8542
lemma convex_closure_inter:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8543
  assumes "\<forall>S\<in>I. convex (S :: 'n::euclidean_space set)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8544
    and "\<Inter>{rel_interior S |S. S \<in> I} \<noteq> {}"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  8545
  shows "closure (\<Inter>I) = \<Inter>{closure S |S. S \<in> I}"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8546
proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8547
  have "\<Inter>{closure S |S. S \<in> I} \<le> closure (\<Inter>{rel_interior S |S. S \<in> I})"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8548
    using convex_closure_rel_interior_inter assms by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8549
  moreover
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60420
diff changeset
  8550
  have "closure (\<Inter>{rel_interior S |S. S \<in> I}) \<le> closure (\<Inter>I)"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  8551
    using rel_interior_inter_aux closure_mono[of "\<Inter>{rel_interior S |S. S \<in> I}" "\<Inter>I"]
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8552
    by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8553
  ultimately show ?thesis
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8554
    using closure_inter[of I] by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8555
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8556
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  8557
lemma convex_inter_rel_interior_same_closure:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8558
  assumes "\<forall>S\<in>I. convex (S :: 'n::euclidean_space set)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8559
    and "\<Inter>{rel_interior S |S. S \<in> I} \<noteq> {}"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  8560
  shows "closure (\<Inter>{rel_interior S |S. S \<in> I}) = closure (\<Inter>I)"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8561
proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8562
  have "\<Inter>{closure S |S. S \<in> I} \<le> closure (\<Inter>{rel_interior S |S. S \<in> I})"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8563
    using convex_closure_rel_interior_inter assms by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8564
  moreover
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8565
  have "closure (\<Inter>{rel_interior S |S. S \<in> I}) \<le> closure (\<Inter>I)"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  8566
    using rel_interior_inter_aux closure_mono[of "\<Inter>{rel_interior S |S. S \<in> I}" "\<Inter>I"]
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8567
    by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8568
  ultimately show ?thesis
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8569
    using closure_inter[of I] by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8570
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8571
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  8572
lemma convex_rel_interior_inter:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8573
  assumes "\<forall>S\<in>I. convex (S :: 'n::euclidean_space set)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8574
    and "\<Inter>{rel_interior S |S. S \<in> I} \<noteq> {}"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  8575
  shows "rel_interior (\<Inter>I) \<subseteq> \<Inter>{rel_interior S |S. S \<in> I}"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8576
proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8577
  have "convex (\<Inter>I)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8578
    using assms convex_Inter by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8579
  moreover
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  8580
  have "convex (\<Inter>{rel_interior S |S. S \<in> I})"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8581
    apply (rule convex_Inter)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8582
    using assms convex_rel_interior
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8583
    apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8584
    done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8585
  ultimately
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8586
  have "rel_interior (\<Inter>{rel_interior S |S. S \<in> I}) = rel_interior (\<Inter>I)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8587
    using convex_inter_rel_interior_same_closure assms
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8588
      closure_eq_rel_interior_eq[of "\<Inter>{rel_interior S |S. S \<in> I}" "\<Inter>I"]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8589
    by blast
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8590
  then show ?thesis
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8591
    using rel_interior_subset[of "\<Inter>{rel_interior S |S. S \<in> I}"] by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8592
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8593
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  8594
lemma convex_rel_interior_finite_inter:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8595
  assumes "\<forall>S\<in>I. convex (S :: 'n::euclidean_space set)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8596
    and "\<Inter>{rel_interior S |S. S \<in> I} \<noteq> {}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8597
    and "finite I"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8598
  shows "rel_interior (\<Inter>I) = \<Inter>{rel_interior S |S. S \<in> I}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8599
proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8600
  have "\<Inter>I \<noteq> {}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8601
    using assms rel_interior_inter_aux[of I] by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8602
  have "convex (\<Inter>I)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8603
    using convex_Inter assms by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8604
  show ?thesis
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8605
  proof (cases "I = {}")
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8606
    case True
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8607
    then show ?thesis
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  8608
      using Inter_empty rel_interior_UNIV by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8609
  next
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8610
    case False
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8611
    {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8612
      fix z
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8613
      assume z: "z \<in> \<Inter>{rel_interior S |S. S \<in> I}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8614
      {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8615
        fix x
61952
546958347e05 prefer symbols for "Union", "Inter";
wenzelm
parents: 61945
diff changeset
  8616
        assume x: "x \<in> \<Inter>I"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8617
        {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8618
          fix S
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8619
          assume S: "S \<in> I"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8620
          then have "z \<in> rel_interior S" "x \<in> S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8621
            using z x by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8622
          then have "\<exists>m. m > 1 \<and> (\<forall>e. e > 1 \<and> e \<le> m \<longrightarrow> (1 - e)*\<^sub>R x + e *\<^sub>R z \<in> S)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8623
            using convex_rel_interior_if[of S z] S assms hull_subset[of S] by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8624
        }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8625
        then obtain mS where
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8626
          mS: "\<forall>S\<in>I. mS S > 1 \<and> (\<forall>e. e > 1 \<and> e \<le> mS S \<longrightarrow> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> S)" by metis
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  8627
        define e where "e = Min (mS ` I)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  8628
        then have "e \<in> mS ` I" using assms \<open>I \<noteq> {}\<close> by simp
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8629
        then have "e > 1" using mS by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8630
        moreover have "\<forall>S\<in>I. e \<le> mS S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8631
          using e_def assms by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8632
        ultimately have "\<exists>e > 1. (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> \<Inter>I"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8633
          using mS by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8634
      }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8635
      then have "z \<in> rel_interior (\<Inter>I)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  8636
        using convex_rel_interior_iff[of "\<Inter>I" z] \<open>\<Inter>I \<noteq> {}\<close> \<open>convex (\<Inter>I)\<close> by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8637
    }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8638
    then show ?thesis
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8639
      using convex_rel_interior_inter[of I] assms by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8640
  qed
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8641
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8642
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  8643
lemma convex_closure_inter_two:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8644
  fixes S T :: "'n::euclidean_space set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8645
  assumes "convex S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8646
    and "convex T"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8647
  assumes "rel_interior S \<inter> rel_interior T \<noteq> {}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8648
  shows "closure (S \<inter> T) = closure S \<inter> closure T"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8649
  using convex_closure_inter[of "{S,T}"] assms by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8650
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  8651
lemma convex_rel_interior_inter_two:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8652
  fixes S T :: "'n::euclidean_space set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8653
  assumes "convex S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8654
    and "convex T"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8655
    and "rel_interior S \<inter> rel_interior T \<noteq> {}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8656
  shows "rel_interior (S \<inter> T) = rel_interior S \<inter> rel_interior T"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8657
  using convex_rel_interior_finite_inter[of "{S,T}"] assms by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8658
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  8659
lemma convex_affine_closure_inter:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8660
  fixes S T :: "'n::euclidean_space set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8661
  assumes "convex S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8662
    and "affine T"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8663
    and "rel_interior S \<inter> T \<noteq> {}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8664
  shows "closure (S \<inter> T) = closure S \<inter> T"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8665
proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8666
  have "affine hull T = T"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8667
    using assms by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8668
  then have "rel_interior T = T"
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  8669
    using rel_interior_affine_hull[of T] by metis
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8670
  moreover have "closure T = T"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8671
    using assms affine_closed[of T] by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8672
  ultimately show ?thesis
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8673
    using convex_closure_inter_two[of S T] assms affine_imp_convex by auto
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  8674
qed
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  8675
62620
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8676
lemma connected_component_1_gen:
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8677
  fixes S :: "'a :: euclidean_space set"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8678
  assumes "DIM('a) = 1"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8679
  shows "connected_component S a b \<longleftrightarrow> closed_segment a b \<subseteq> S"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8680
unfolding connected_component_def
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8681
by (metis (no_types, lifting) assms subsetD subsetI convex_contains_segment convex_segment(1)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8682
            ends_in_segment connected_convex_1_gen)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8683
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8684
lemma connected_component_1:
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8685
  fixes S :: "real set"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8686
  shows "connected_component S a b \<longleftrightarrow> closed_segment a b \<subseteq> S"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8687
by (simp add: connected_component_1_gen)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  8688
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  8689
lemma convex_affine_rel_interior_inter:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8690
  fixes S T :: "'n::euclidean_space set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8691
  assumes "convex S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8692
    and "affine T"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8693
    and "rel_interior S \<inter> T \<noteq> {}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8694
  shows "rel_interior (S \<inter> T) = rel_interior S \<inter> T"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8695
proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8696
  have "affine hull T = T"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8697
    using assms by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8698
  then have "rel_interior T = T"
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  8699
    using rel_interior_affine_hull[of T] by metis
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8700
  moreover have "closure T = T"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8701
    using assms affine_closed[of T] by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8702
  ultimately show ?thesis
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8703
    using convex_rel_interior_inter_two[of S T] assms affine_imp_convex by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8704
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8705
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8706
lemma subset_rel_interior_convex:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8707
  fixes S T :: "'n::euclidean_space set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8708
  assumes "convex S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8709
    and "convex T"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8710
    and "S \<le> closure T"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8711
    and "\<not> S \<subseteq> rel_frontier T"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8712
  shows "rel_interior S \<subseteq> rel_interior T"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8713
proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8714
  have *: "S \<inter> closure T = S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8715
    using assms by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8716
  have "\<not> rel_interior S \<subseteq> rel_frontier T"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8717
    using closure_mono[of "rel_interior S" "rel_frontier T"] closed_rel_frontier[of T]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8718
      closure_closed[of S] convex_closure_rel_interior[of S] closure_subset[of S] assms
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8719
    by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8720
  then have "rel_interior S \<inter> rel_interior (closure T) \<noteq> {}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8721
    using assms rel_frontier_def[of T] rel_interior_subset convex_rel_interior_closure[of T]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8722
    by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8723
  then have "rel_interior S \<inter> rel_interior T = rel_interior (S \<inter> closure T)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8724
    using assms convex_closure convex_rel_interior_inter_two[of S "closure T"]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8725
      convex_rel_interior_closure[of T]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8726
    by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8727
  also have "\<dots> = rel_interior S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8728
    using * by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8729
  finally show ?thesis
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8730
    by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8731
qed
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8732
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8733
lemma rel_interior_convex_linear_image:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8734
  fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8735
  assumes "linear f"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8736
    and "convex S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8737
  shows "f ` (rel_interior S) = rel_interior (f ` S)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8738
proof (cases "S = {}")
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8739
  case True
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8740
  then show ?thesis
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61609
diff changeset
  8741
    using assms rel_interior_empty rel_interior_eq_empty by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8742
next
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8743
  case False
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8744
  have *: "f ` (rel_interior S) \<subseteq> f ` S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8745
    unfolding image_mono using rel_interior_subset by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8746
  have "f ` S \<subseteq> f ` (closure S)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8747
    unfolding image_mono using closure_subset by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8748
  also have "\<dots> = f ` (closure (rel_interior S))"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8749
    using convex_closure_rel_interior assms by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8750
  also have "\<dots> \<subseteq> closure (f ` (rel_interior S))"
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  8751
    using closure_linear_image_subset assms by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8752
  finally have "closure (f ` S) = closure (f ` rel_interior S)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8753
    using closure_mono[of "f ` S" "closure (f ` rel_interior S)"] closure_closure
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8754
      closure_mono[of "f ` rel_interior S" "f ` S"] *
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8755
    by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8756
  then have "rel_interior (f ` S) = rel_interior (f ` rel_interior S)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8757
    using assms convex_rel_interior
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  8758
      linear_conv_bounded_linear[of f] convex_linear_image[of _ S]
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  8759
      convex_linear_image[of _ "rel_interior S"]
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8760
      closure_eq_rel_interior_eq[of "f ` S" "f ` rel_interior S"]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8761
    by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8762
  then have "rel_interior (f ` S) \<subseteq> f ` rel_interior S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8763
    using rel_interior_subset by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8764
  moreover
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8765
  {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8766
    fix z
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8767
    assume "z \<in> f ` rel_interior S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8768
    then obtain z1 where z1: "z1 \<in> rel_interior S" "f z1 = z" by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8769
    {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8770
      fix x
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8771
      assume "x \<in> f ` S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8772
      then obtain x1 where x1: "x1 \<in> S" "f x1 = x" by auto
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  8773
      then obtain e where e: "e > 1" "(1 - e) *\<^sub>R x1 + e *\<^sub>R z1 : S"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  8774
        using convex_rel_interior_iff[of S z1] \<open>convex S\<close> x1 z1 by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8775
      moreover have "f ((1 - e) *\<^sub>R x1 + e *\<^sub>R z1) = (1 - e) *\<^sub>R x + e *\<^sub>R z"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  8776
        using x1 z1 \<open>linear f\<close> by (simp add: linear_add_cmul)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8777
      ultimately have "(1 - e) *\<^sub>R x + e *\<^sub>R z : f ` S"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8778
        using imageI[of "(1 - e) *\<^sub>R x1 + e *\<^sub>R z1" S f] by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8779
      then have "\<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z : f ` S"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  8780
        using e by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8781
    }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8782
    then have "z \<in> rel_interior (f ` S)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  8783
      using convex_rel_interior_iff[of "f ` S" z] \<open>convex S\<close>
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  8784
        \<open>linear f\<close> \<open>S \<noteq> {}\<close> convex_linear_image[of f S]  linear_conv_bounded_linear[of f]
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8785
      by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8786
  }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8787
  ultimately show ?thesis by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8788
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8789
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8790
lemma rel_interior_convex_linear_preimage:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8791
  fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8792
  assumes "linear f"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8793
    and "convex S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8794
    and "f -` (rel_interior S) \<noteq> {}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8795
  shows "rel_interior (f -` S) = f -` (rel_interior S)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8796
proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8797
  have "S \<noteq> {}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8798
    using assms rel_interior_empty by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8799
  have nonemp: "f -` S \<noteq> {}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8800
    by (metis assms(3) rel_interior_subset subset_empty vimage_mono)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8801
  then have "S \<inter> (range f) \<noteq> {}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8802
    by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8803
  have conv: "convex (f -` S)"
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  8804
    using convex_linear_vimage assms by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8805
  then have "convex (S \<inter> range f)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8806
    by (metis assms(1) assms(2) convex_Int subspace_UNIV subspace_imp_convex subspace_linear_image)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8807
  {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8808
    fix z
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8809
    assume "z \<in> f -` (rel_interior S)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8810
    then have z: "f z : rel_interior S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8811
      by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8812
    {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8813
      fix x
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8814
      assume "x \<in> f -` S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8815
      then have "f x \<in> S" by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8816
      then obtain e where e: "e > 1" "(1 - e) *\<^sub>R f x + e *\<^sub>R f z \<in> S"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  8817
        using convex_rel_interior_iff[of S "f z"] z assms \<open>S \<noteq> {}\<close> by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8818
      moreover have "(1 - e) *\<^sub>R f x + e *\<^sub>R f z = f ((1 - e) *\<^sub>R x + e *\<^sub>R z)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  8819
        using \<open>linear f\<close> by (simp add: linear_iff)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8820
      ultimately have "\<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R z \<in> f -` S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8821
        using e by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8822
    }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8823
    then have "z \<in> rel_interior (f -` S)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8824
      using convex_rel_interior_iff[of "f -` S" z] conv nonemp by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8825
  }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8826
  moreover
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  8827
  {
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8828
    fix z
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8829
    assume z: "z \<in> rel_interior (f -` S)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8830
    {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8831
      fix x
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8832
      assume "x \<in> S \<inter> range f"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8833
      then obtain y where y: "f y = x" "y \<in> f -` S" by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8834
      then obtain e where e: "e > 1" "(1 - e) *\<^sub>R y + e *\<^sub>R z \<in> f -` S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8835
        using convex_rel_interior_iff[of "f -` S" z] z conv by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8836
      moreover have "(1 - e) *\<^sub>R x + e *\<^sub>R f z = f ((1 - e) *\<^sub>R y + e *\<^sub>R z)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  8837
        using \<open>linear f\<close> y by (simp add: linear_iff)
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8838
      ultimately have "\<exists>e. e > 1 \<and> (1 - e) *\<^sub>R x + e *\<^sub>R f z \<in> S \<inter> range f"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8839
        using e by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8840
    }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8841
    then have "f z \<in> rel_interior (S \<inter> range f)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  8842
      using \<open>convex (S \<inter> (range f))\<close> \<open>S \<inter> range f \<noteq> {}\<close>
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8843
        convex_rel_interior_iff[of "S \<inter> (range f)" "f z"]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8844
      by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8845
    moreover have "affine (range f)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8846
      by (metis assms(1) subspace_UNIV subspace_imp_affine subspace_linear_image)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8847
    ultimately have "f z \<in> rel_interior S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8848
      using convex_affine_rel_interior_inter[of S "range f"] assms by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8849
    then have "z \<in> f -` (rel_interior S)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8850
      by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8851
  }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8852
  ultimately show ?thesis by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8853
qed
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  8854
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  8855
lemma rel_interior_Times:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8856
  fixes S :: "'n::euclidean_space set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8857
    and T :: "'m::euclidean_space set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8858
  assumes "convex S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8859
    and "convex T"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8860
  shows "rel_interior (S \<times> T) = rel_interior S \<times> rel_interior T"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8861
proof -
60303
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60176
diff changeset
  8862
  { assume "S = {}"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8863
    then have ?thesis
60303
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60176
diff changeset
  8864
      by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8865
  }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8866
  moreover
60303
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60176
diff changeset
  8867
  { assume "T = {}"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8868
    then have ?thesis
60303
00c06f1315d0 New material about paths, and some lemmas
paulson
parents: 60176
diff changeset
  8869
       by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8870
  }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8871
  moreover
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8872
  {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8873
    assume "S \<noteq> {}" "T \<noteq> {}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8874
    then have ri: "rel_interior S \<noteq> {}" "rel_interior T \<noteq> {}"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61609
diff changeset
  8875
      using rel_interior_eq_empty assms by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8876
    then have "fst -` rel_interior S \<noteq> {}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8877
      using fst_vimage_eq_Times[of "rel_interior S"] by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8878
    then have "rel_interior ((fst :: 'n * 'm \<Rightarrow> 'n) -` S) = fst -` rel_interior S"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  8879
      using fst_linear \<open>convex S\<close> rel_interior_convex_linear_preimage[of fst S] by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8880
    then have s: "rel_interior (S \<times> (UNIV :: 'm set)) = rel_interior S \<times> UNIV"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8881
      by (simp add: fst_vimage_eq_Times)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8882
    from ri have "snd -` rel_interior T \<noteq> {}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8883
      using snd_vimage_eq_Times[of "rel_interior T"] by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8884
    then have "rel_interior ((snd :: 'n * 'm \<Rightarrow> 'm) -` T) = snd -` rel_interior T"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  8885
      using snd_linear \<open>convex T\<close> rel_interior_convex_linear_preimage[of snd T] by auto
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8886
    then have t: "rel_interior ((UNIV :: 'n set) \<times> T) = UNIV \<times> rel_interior T"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8887
      by (simp add: snd_vimage_eq_Times)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8888
    from s t have *: "rel_interior (S \<times> (UNIV :: 'm set)) \<inter> rel_interior ((UNIV :: 'n set) \<times> T) =
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8889
      rel_interior S \<times> rel_interior T" by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8890
    have "S \<times> T = S \<times> (UNIV :: 'm set) \<inter> (UNIV :: 'n set) \<times> T"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8891
      by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8892
    then have "rel_interior (S \<times> T) = rel_interior ((S \<times> (UNIV :: 'm set)) \<inter> ((UNIV :: 'n set) \<times> T))"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8893
      by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8894
    also have "\<dots> = rel_interior (S \<times> (UNIV :: 'm set)) \<inter> rel_interior ((UNIV :: 'n set) \<times> T)"
55787
41a73a41f6c8 more symbols;
wenzelm
parents: 55697
diff changeset
  8895
       apply (subst convex_rel_interior_inter_two[of "S \<times> (UNIV :: 'm set)" "(UNIV :: 'n set) \<times> T"])
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  8896
       using * ri assms convex_Times
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8897
       apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8898
       done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8899
    finally have ?thesis using * by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8900
  }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8901
  ultimately show ?thesis by blast
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8902
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8903
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  8904
lemma rel_interior_scaleR:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8905
  fixes S :: "'n::euclidean_space set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8906
  assumes "c \<noteq> 0"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8907
  shows "(op *\<^sub>R c) ` (rel_interior S) = rel_interior ((op *\<^sub>R c) ` S)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8908
  using rel_interior_injective_linear_image[of "(op *\<^sub>R c)" S]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8909
    linear_conv_bounded_linear[of "op *\<^sub>R c"] linear_scaleR injective_scaleR[of c] assms
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8910
  by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8911
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  8912
lemma rel_interior_convex_scaleR:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8913
  fixes S :: "'n::euclidean_space set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8914
  assumes "convex S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8915
  shows "(op *\<^sub>R c) ` (rel_interior S) = rel_interior ((op *\<^sub>R c) ` S)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8916
  by (metis assms linear_scaleR rel_interior_convex_linear_image)
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8917
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  8918
lemma convex_rel_open_scaleR:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8919
  fixes S :: "'n::euclidean_space set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8920
  assumes "convex S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8921
    and "rel_open S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8922
  shows "convex ((op *\<^sub>R c) ` S) \<and> rel_open ((op *\<^sub>R c) ` S)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8923
  by (metis assms convex_scaling rel_interior_convex_scaleR rel_open_def)
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8924
49531
8d68162b7826 tuned whitespace;
wenzelm
parents: 49530
diff changeset
  8925
lemma convex_rel_open_finite_inter:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8926
  assumes "\<forall>S\<in>I. convex (S :: 'n::euclidean_space set) \<and> rel_open S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8927
    and "finite I"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8928
  shows "convex (\<Inter>I) \<and> rel_open (\<Inter>I)"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  8929
proof (cases "\<Inter>{rel_interior S |S. S \<in> I} = {}")
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8930
  case True
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8931
  then have "\<Inter>I = {}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8932
    using assms unfolding rel_open_def by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8933
  then show ?thesis
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8934
    unfolding rel_open_def using rel_interior_empty by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8935
next
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8936
  case False
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  8937
  then have "rel_open (\<Inter>I)"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8938
    using assms unfolding rel_open_def
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8939
    using convex_rel_interior_finite_inter[of I]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8940
    by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8941
  then show ?thesis
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8942
    using convex_Inter assms by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8943
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8944
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8945
lemma convex_rel_open_linear_image:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8946
  fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8947
  assumes "linear f"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8948
    and "convex S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8949
    and "rel_open S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8950
  shows "convex (f ` S) \<and> rel_open (f ` S)"
57865
dcfb33c26f50 tuned proofs -- fewer warnings;
wenzelm
parents: 57512
diff changeset
  8951
  by (metis assms convex_linear_image rel_interior_convex_linear_image rel_open_def)
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8952
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8953
lemma convex_rel_open_linear_preimage:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8954
  fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8955
  assumes "linear f"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8956
    and "convex S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8957
    and "rel_open S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8958
  shows "convex (f -` S) \<and> rel_open (f -` S)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8959
proof (cases "f -` (rel_interior S) = {}")
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8960
  case True
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8961
  then have "f -` S = {}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8962
    using assms unfolding rel_open_def by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8963
  then show ?thesis
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8964
    unfolding rel_open_def using rel_interior_empty by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8965
next
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8966
  case False
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8967
  then have "rel_open (f -` S)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8968
    using assms unfolding rel_open_def
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8969
    using rel_interior_convex_linear_preimage[of f S]
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8970
    by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8971
  then show ?thesis
53620
3c7f5e7926dc generalized and simplified proofs of several theorems about convex sets
huffman
parents: 53600
diff changeset
  8972
    using convex_linear_vimage assms
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8973
    by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8974
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8975
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8976
lemma rel_interior_projection:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8977
  fixes S :: "('m::euclidean_space \<times> 'n::euclidean_space) set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8978
    and f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8979
  assumes "convex S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8980
    and "f = (\<lambda>y. {z. (y, z) \<in> S})"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8981
  shows "(y, z) \<in> rel_interior S \<longleftrightarrow> (y \<in> rel_interior {y. (f y \<noteq> {})} \<and> z \<in> rel_interior (f y))"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8982
proof -
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8983
  {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8984
    fix y
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8985
    assume "y \<in> {y. f y \<noteq> {}}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8986
    then obtain z where "(y, z) \<in> S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8987
      using assms by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8988
    then have "\<exists>x. x \<in> S \<and> y = fst x"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8989
      apply (rule_tac x="(y, z)" in exI)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8990
      apply auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8991
      done
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8992
    then obtain x where "x \<in> S" "y = fst x"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8993
      by blast
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8994
    then have "y \<in> fst ` S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8995
      unfolding image_def by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  8996
  }
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8997
  then have "fst ` S = {y. f y \<noteq> {}}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8998
    unfolding fst_def using assms by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  8999
  then have h1: "fst ` rel_interior S = rel_interior {y. f y \<noteq> {}}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9000
    using rel_interior_convex_linear_image[of fst S] assms fst_linear by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9001
  {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9002
    fix y
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9003
    assume "y \<in> rel_interior {y. f y \<noteq> {}}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9004
    then have "y \<in> fst ` rel_interior S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9005
      using h1 by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9006
    then have *: "rel_interior S \<inter> fst -` {y} \<noteq> {}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9007
      by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9008
    moreover have aff: "affine (fst -` {y})"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9009
      unfolding affine_alt by (simp add: algebra_simps)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9010
    ultimately have **: "rel_interior (S \<inter> fst -` {y}) = rel_interior S \<inter> fst -` {y}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9011
      using convex_affine_rel_interior_inter[of S "fst -` {y}"] assms by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9012
    have conv: "convex (S \<inter> fst -` {y})"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9013
      using convex_Int assms aff affine_imp_convex by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9014
    {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9015
      fix x
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9016
      assume "x \<in> f y"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9017
      then have "(y, x) \<in> S \<inter> (fst -` {y})"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9018
        using assms by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9019
      moreover have "x = snd (y, x)" by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9020
      ultimately have "x \<in> snd ` (S \<inter> fst -` {y})"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9021
        by blast
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9022
    }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9023
    then have "snd ` (S \<inter> fst -` {y}) = f y"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9024
      using assms by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9025
    then have ***: "rel_interior (f y) = snd ` rel_interior (S \<inter> fst -` {y})"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9026
      using rel_interior_convex_linear_image[of snd "S \<inter> fst -` {y}"] snd_linear conv
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9027
      by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9028
    {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9029
      fix z
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9030
      assume "z \<in> rel_interior (f y)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9031
      then have "z \<in> snd ` rel_interior (S \<inter> fst -` {y})"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9032
        using *** by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9033
      moreover have "{y} = fst ` rel_interior (S \<inter> fst -` {y})"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9034
        using * ** rel_interior_subset by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9035
      ultimately have "(y, z) \<in> rel_interior (S \<inter> fst -` {y})"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9036
        by force
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9037
      then have "(y,z) \<in> rel_interior S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9038
        using ** by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9039
    }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9040
    moreover
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9041
    {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9042
      fix z
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9043
      assume "(y, z) \<in> rel_interior S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9044
      then have "(y, z) \<in> rel_interior (S \<inter> fst -` {y})"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9045
        using ** by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9046
      then have "z \<in> snd ` rel_interior (S \<inter> fst -` {y})"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9047
        by (metis Range_iff snd_eq_Range)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9048
      then have "z \<in> rel_interior (f y)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9049
        using *** by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9050
    }
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9051
    ultimately have "\<And>z. (y, z) \<in> rel_interior S \<longleftrightarrow> z \<in> rel_interior (f y)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9052
      by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  9053
  }
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9054
  then have h2: "\<And>y z. y \<in> rel_interior {t. f t \<noteq> {}} \<Longrightarrow>
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9055
    (y, z) \<in> rel_interior S \<longleftrightarrow> z \<in> rel_interior (f y)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9056
    by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9057
  {
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9058
    fix y z
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9059
    assume asm: "(y, z) \<in> rel_interior S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9060
    then have "y \<in> fst ` rel_interior S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9061
      by (metis Domain_iff fst_eq_Domain)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9062
    then have "y \<in> rel_interior {t. f t \<noteq> {}}"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9063
      using h1 by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9064
    then have "y \<in> rel_interior {t. f t \<noteq> {}}" and "(z : rel_interior (f y))"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9065
      using h2 asm by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  9066
  }
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9067
  then show ?thesis using h2 by blast
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9068
qed
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9069
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  9070
lemma rel_frontier_Times:
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  9071
  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
  9072
    and T :: "'m::euclidean_space set"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  9073
  assumes "convex S"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  9074
    and "convex T"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  9075
  shows "rel_frontier S \<times> rel_frontier T \<subseteq> rel_frontier (S \<times> T)"
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  9076
    by (force simp: rel_frontier_def rel_interior_Times assms closure_Times)
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  9077
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  9078
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  9079
subsubsection \<open>Relative interior of convex cone\<close>
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  9080
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  9081
lemma cone_rel_interior:
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9082
  fixes S :: "'m::euclidean_space set"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9083
  assumes "cone S"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9084
  shows "cone ({0} \<union> rel_interior S)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9085
proof (cases "S = {}")
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9086
  case True
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9087
  then show ?thesis
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9088
    by (simp add: rel_interior_empty cone_0)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9089
next
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9090
  case False
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9091
  then have *: "0 \<in> S \<and> (\<forall>c. c > 0 \<longrightarrow> op *\<^sub>R c ` S = S)"
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9092
    using cone_iff[of S] assms by auto
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9093
  then have *: "0 \<in> ({0} \<union> rel_interior S)"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9094
    and "\<forall>c. c > 0 \<longrightarrow> op *\<^sub>R c ` ({0} \<union> rel_interior S) = ({0} \<union> rel_interior S)"
53348
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9095
    by (auto simp add: rel_interior_scaleR)
0b467fc4e597 tuned proofs;
wenzelm
parents: 53347
diff changeset
  9096
  then show ?thesis
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9097
    using cone_iff[of "{0} \<union> rel_interior S"] by auto
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  9098
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  9099
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  9100
lemma rel_interior_convex_cone_aux:
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9101
  fixes S :: "'m::euclidean_space set"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9102
  assumes "convex S"
55787
41a73a41f6c8 more symbols;
wenzelm
parents: 55697
diff changeset
  9103
  shows "(c, x) \<in> rel_interior (cone hull ({(1 :: real)} \<times> S)) \<longleftrightarrow>
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9104
    c > 0 \<and> x \<in> ((op *\<^sub>R c) ` (rel_interior S))"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9105
proof (cases "S = {}")
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9106
  case True
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9107
  then show ?thesis
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9108
    by (simp add: rel_interior_empty cone_hull_empty)
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9109
next
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9110
  case False
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9111
  then obtain s where "s \<in> S" by auto
55787
41a73a41f6c8 more symbols;
wenzelm
parents: 55697
diff changeset
  9112
  have conv: "convex ({(1 :: real)} \<times> S)"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9113
    using convex_Times[of "{(1 :: real)}" S] assms convex_singleton[of "1 :: real"]
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9114
    by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  9115
  define f where "f y = {z. (y, z) \<in> cone hull ({1 :: real} \<times> S)}" for y
55787
41a73a41f6c8 more symbols;
wenzelm
parents: 55697
diff changeset
  9116
  then have *: "(c, x) \<in> rel_interior (cone hull ({(1 :: real)} \<times> S)) =
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9117
    (c \<in> rel_interior {y. f y \<noteq> {}} \<and> x \<in> rel_interior (f c))"
55787
41a73a41f6c8 more symbols;
wenzelm
parents: 55697
diff changeset
  9118
    apply (subst rel_interior_projection[of "cone hull ({(1 :: real)} \<times> S)" f c x])
41a73a41f6c8 more symbols;
wenzelm
parents: 55697
diff changeset
  9119
    using convex_cone_hull[of "{(1 :: real)} \<times> S"] conv
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9120
    apply auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9121
    done
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9122
  {
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9123
    fix y :: real
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9124
    assume "y \<ge> 0"
55787
41a73a41f6c8 more symbols;
wenzelm
parents: 55697
diff changeset
  9125
    then have "y *\<^sub>R (1,s) \<in> cone hull ({1 :: real} \<times> S)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  9126
      using cone_hull_expl[of "{(1 :: real)} \<times> S"] \<open>s \<in> S\<close> by auto
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9127
    then have "f y \<noteq> {}"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9128
      using f_def by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9129
  }
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9130
  then have "{y. f y \<noteq> {}} = {0..}"
55787
41a73a41f6c8 more symbols;
wenzelm
parents: 55697
diff changeset
  9131
    using f_def cone_hull_expl[of "{1 :: real} \<times> S"] by auto
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9132
  then have **: "rel_interior {y. f y \<noteq> {}} = {0<..}"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9133
    using rel_interior_real_semiline by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9134
  {
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9135
    fix c :: real
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9136
    assume "c > 0"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9137
    then have "f c = (op *\<^sub>R c ` S)"
55787
41a73a41f6c8 more symbols;
wenzelm
parents: 55697
diff changeset
  9138
      using f_def cone_hull_expl[of "{1 :: real} \<times> S"] by auto
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9139
    then have "rel_interior (f c) = op *\<^sub>R c ` rel_interior S"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9140
      using rel_interior_convex_scaleR[of S c] assms by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9141
  }
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9142
  then show ?thesis using * ** by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9143
qed
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  9144
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  9145
lemma rel_interior_convex_cone:
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9146
  fixes S :: "'m::euclidean_space set"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9147
  assumes "convex S"
55787
41a73a41f6c8 more symbols;
wenzelm
parents: 55697
diff changeset
  9148
  shows "rel_interior (cone hull ({1 :: real} \<times> S)) =
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9149
    {(c, c *\<^sub>R x) | c x. c > 0 \<and> x \<in> rel_interior S}"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9150
  (is "?lhs = ?rhs")
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9151
proof -
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9152
  {
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9153
    fix z
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9154
    assume "z \<in> ?lhs"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9155
    have *: "z = (fst z, snd z)"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9156
      by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9157
    have "z \<in> ?rhs"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  9158
      using rel_interior_convex_cone_aux[of S "fst z" "snd z"] assms \<open>z \<in> ?lhs\<close>
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9159
      apply auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9160
      apply (rule_tac x = "fst z" in exI)
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9161
      apply (rule_tac x = x in exI)
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9162
      using *
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9163
      apply auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9164
      done
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9165
  }
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9166
  moreover
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9167
  {
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9168
    fix z
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9169
    assume "z \<in> ?rhs"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9170
    then have "z \<in> ?lhs"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9171
      using rel_interior_convex_cone_aux[of S "fst z" "snd z"] assms
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9172
      by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9173
  }
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9174
  ultimately show ?thesis by blast
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  9175
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  9176
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  9177
lemma convex_hull_finite_union:
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9178
  assumes "finite I"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9179
  assumes "\<forall>i\<in>I. convex (S i) \<and> (S i) \<noteq> {}"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9180
  shows "convex hull (\<Union>(S ` I)) =
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9181
    {setsum (\<lambda>i. c i *\<^sub>R s i) I | c s. (\<forall>i\<in>I. c i \<ge> 0) \<and> setsum c I = 1 \<and> (\<forall>i\<in>I. s i \<in> S i)}"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  9182
  (is "?lhs = ?rhs")
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9183
proof -
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9184
  have "?lhs \<supseteq> ?rhs"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9185
  proof
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9186
    fix x
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9187
    assume "x : ?rhs"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9188
    then obtain c s where *: "setsum (\<lambda>i. c i *\<^sub>R s i) I = x" "setsum c I = 1"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9189
      "(\<forall>i\<in>I. c i \<ge> 0) \<and> (\<forall>i\<in>I. s i \<in> S i)" by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9190
    then have "\<forall>i\<in>I. s i \<in> convex hull (\<Union>(S ` I))"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9191
      using hull_subset[of "\<Union>(S ` I)" convex] by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9192
    then show "x \<in> ?lhs"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9193
      unfolding *(1)[symmetric]
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9194
      apply (subst convex_setsum[of I "convex hull \<Union>(S ` I)" c s])
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9195
      using * assms convex_convex_hull
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9196
      apply auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9197
      done
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9198
  qed
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9199
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9200
  {
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9201
    fix i
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9202
    assume "i \<in> I"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9203
    with assms have "\<exists>p. p \<in> S i" by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9204
  }
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9205
  then obtain p where p: "\<forall>i\<in>I. p i \<in> S i" by metis
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9206
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9207
  {
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9208
    fix i
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9209
    assume "i \<in> I"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9210
    {
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9211
      fix x
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9212
      assume "x \<in> S i"
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  9213
      define c where "c j = (if j = i then 1::real else 0)" for j
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9214
      then have *: "setsum c I = 1"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  9215
        using \<open>finite I\<close> \<open>i \<in> I\<close> setsum.delta[of I i "\<lambda>j::'a. 1::real"]
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9216
        by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  9217
      define s where "s j = (if j = i then x else p j)" for j
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9218
      then have "\<forall>j. c j *\<^sub>R s j = (if j = i then x else 0)"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9219
        using c_def by (auto simp add: algebra_simps)
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9220
      then have "x = setsum (\<lambda>i. c i *\<^sub>R s i) I"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  9221
        using s_def c_def \<open>finite I\<close> \<open>i \<in> I\<close> setsum.delta[of I i "\<lambda>j::'a. x"]
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9222
        by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9223
      then have "x \<in> ?rhs"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9224
        apply auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9225
        apply (rule_tac x = c in exI)
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9226
        apply (rule_tac x = s in exI)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  9227
        using * c_def s_def p \<open>x \<in> S i\<close>
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9228
        apply auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9229
        done
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  9230
    }
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9231
    then have "?rhs \<supseteq> S i" by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9232
  }
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9233
  then have *: "?rhs \<supseteq> \<Union>(S ` I)" by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9234
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9235
  {
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9236
    fix u v :: real
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9237
    assume uv: "u \<ge> 0 \<and> v \<ge> 0 \<and> u + v = 1"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9238
    fix x y
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9239
    assume xy: "x \<in> ?rhs \<and> y \<in> ?rhs"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9240
    from xy obtain c s where
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9241
      xc: "x = setsum (\<lambda>i. c i *\<^sub>R s i) I \<and> (\<forall>i\<in>I. c i \<ge> 0) \<and> setsum c I = 1 \<and> (\<forall>i\<in>I. s i \<in> S i)"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9242
      by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9243
    from xy obtain d t where
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9244
      yc: "y = setsum (\<lambda>i. d i *\<^sub>R t i) I \<and> (\<forall>i\<in>I. d i \<ge> 0) \<and> setsum d I = 1 \<and> (\<forall>i\<in>I. t i \<in> S i)"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9245
      by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  9246
    define e where "e i = u * c i + v * d i" for i
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9247
    have ge0: "\<forall>i\<in>I. e i \<ge> 0"
56536
aefb4a8da31f made mult_nonneg_nonneg a simp rule
nipkow
parents: 56480
diff changeset
  9248
      using e_def xc yc uv by simp
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9249
    have "setsum (\<lambda>i. u * c i) I = u * setsum c I"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9250
      by (simp add: setsum_right_distrib)
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9251
    moreover have "setsum (\<lambda>i. v * d i) I = v * setsum d I"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9252
      by (simp add: setsum_right_distrib)
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9253
    ultimately have sum1: "setsum e I = 1"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  9254
      using e_def xc yc uv by (simp add: setsum.distrib)
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  9255
    define q where "q i = (if e i = 0 then p i else (u * c i / e i) *\<^sub>R s i + (v * d i / e i) *\<^sub>R t i)"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  9256
      for i
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9257
    {
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9258
      fix i
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9259
      assume i: "i \<in> I"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9260
      have "q i \<in> S i"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9261
      proof (cases "e i = 0")
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9262
        case True
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9263
        then show ?thesis using i p q_def by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9264
      next
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9265
        case False
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9266
        then show ?thesis
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9267
          using mem_convex_alt[of "S i" "s i" "t i" "u * (c i)" "v * (d i)"]
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9268
            mult_nonneg_nonneg[of u "c i"] mult_nonneg_nonneg[of v "d i"]
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9269
            assms q_def e_def i False xc yc uv
56536
aefb4a8da31f made mult_nonneg_nonneg a simp rule
nipkow
parents: 56480
diff changeset
  9270
          by (auto simp del: mult_nonneg_nonneg)
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9271
      qed
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9272
    }
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9273
    then have qs: "\<forall>i\<in>I. q i \<in> S i" by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9274
    {
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9275
      fix i
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9276
      assume i: "i \<in> I"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9277
      have "(u * c i) *\<^sub>R s i + (v * d i) *\<^sub>R t i = e i *\<^sub>R q i"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9278
      proof (cases "e i = 0")
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9279
        case True
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9280
        have ge: "u * (c i) \<ge> 0 \<and> v * d i \<ge> 0"
56536
aefb4a8da31f made mult_nonneg_nonneg a simp rule
nipkow
parents: 56480
diff changeset
  9281
          using xc yc uv i by simp
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9282
        moreover from ge have "u * c i \<le> 0 \<and> v * d i \<le> 0"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9283
          using True e_def i by simp
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9284
        ultimately have "u * c i = 0 \<and> v * d i = 0" by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9285
        with True show ?thesis by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9286
      next
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9287
        case False
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9288
        then have "(u * (c i)/(e i))*\<^sub>R (s i)+(v * (d i)/(e i))*\<^sub>R (t i) = q i"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9289
          using q_def by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9290
        then have "e i *\<^sub>R ((u * (c i)/(e i))*\<^sub>R (s i)+(v * (d i)/(e i))*\<^sub>R (t i))
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9291
               = (e i) *\<^sub>R (q i)" by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9292
        with False show ?thesis by (simp add: algebra_simps)
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9293
      qed
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9294
    }
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9295
    then have *: "\<forall>i\<in>I. (u * c i) *\<^sub>R s i + (v * d i) *\<^sub>R t i = e i *\<^sub>R q i"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9296
      by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9297
    have "u *\<^sub>R x + v *\<^sub>R y = setsum (\<lambda>i. (u * c i) *\<^sub>R s i + (v * d i) *\<^sub>R t i) I"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56889
diff changeset
  9298
      using xc yc by (simp add: algebra_simps scaleR_right.setsum setsum.distrib)
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9299
    also have "\<dots> = setsum (\<lambda>i. e i *\<^sub>R q i) I"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9300
      using * by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9301
    finally have "u *\<^sub>R x + v *\<^sub>R y = setsum (\<lambda>i. (e i) *\<^sub>R (q i)) I"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9302
      by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9303
    then have "u *\<^sub>R x + v *\<^sub>R y \<in> ?rhs"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9304
      using ge0 sum1 qs by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9305
  }
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9306
  then have "convex ?rhs" unfolding convex_def by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9307
  then show ?thesis
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  9308
    using \<open>?lhs \<supseteq> ?rhs\<close> * hull_minimal[of "\<Union>(S ` I)" ?rhs convex]
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9309
    by blast
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  9310
qed
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  9311
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  9312
lemma convex_hull_union_two:
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9313
  fixes S T :: "'m::euclidean_space set"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9314
  assumes "convex S"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9315
    and "S \<noteq> {}"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9316
    and "convex T"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9317
    and "T \<noteq> {}"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9318
  shows "convex hull (S \<union> T) =
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9319
    {u *\<^sub>R s + v *\<^sub>R t | u v s t. u \<ge> 0 \<and> v \<ge> 0 \<and> u + v = 1 \<and> s \<in> S \<and> t \<in> T}"
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  9320
  (is "?lhs = ?rhs")
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9321
proof
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  9322
  define I :: "nat set" where "I = {1, 2}"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  9323
  define s where "s i = (if i = (1::nat) then S else T)" for i
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9324
  have "\<Union>(s ` I) = S \<union> T"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9325
    using s_def I_def by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9326
  then have "convex hull (\<Union>(s ` I)) = convex hull (S \<union> T)"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9327
    by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9328
  moreover have "convex hull \<Union>(s ` I) =
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9329
    {\<Sum> i\<in>I. c i *\<^sub>R sa i | c sa. (\<forall>i\<in>I. 0 \<le> c i) \<and> setsum c I = 1 \<and> (\<forall>i\<in>I. sa i \<in> s i)}"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9330
      apply (subst convex_hull_finite_union[of I s])
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9331
      using assms s_def I_def
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9332
      apply auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9333
      done
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9334
  moreover have
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9335
    "{\<Sum>i\<in>I. c i *\<^sub>R sa i | c sa. (\<forall>i\<in>I. 0 \<le> c i) \<and> setsum c I = 1 \<and> (\<forall>i\<in>I. sa i \<in> s i)} \<le> ?rhs"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9336
    using s_def I_def by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9337
  ultimately show "?lhs \<subseteq> ?rhs" by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9338
  {
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9339
    fix x
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9340
    assume "x \<in> ?rhs"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9341
    then obtain u v s t where *: "x = u *\<^sub>R s + v *\<^sub>R t \<and> u \<ge> 0 \<and> v \<ge> 0 \<and> u + v = 1 \<and> s \<in> S \<and> t \<in> T"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9342
      by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9343
    then have "x \<in> convex hull {s, t}"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9344
      using convex_hull_2[of s t] by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9345
    then have "x \<in> convex hull (S \<union> T)"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9346
      using * hull_mono[of "{s, t}" "S \<union> T"] by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9347
  }
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9348
  then show "?lhs \<supseteq> ?rhs" by blast
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9349
qed
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9350
40377
0e5d48096f58 Extend convex analysis by Bogdan Grechuk
hoelzl
parents: 39302
diff changeset
  9351
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  9352
subsection \<open>Convexity on direct sums\<close>
40887
ee8d0548c148 Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents: 40719
diff changeset
  9353
ee8d0548c148 Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents: 40719
diff changeset
  9354
lemma closure_sum:
55928
2d7582309d73 generalize lemma closure_sum
huffman
parents: 55787
diff changeset
  9355
  fixes S T :: "'a::real_normed_vector set"
47445
69e96e5500df Set_Algebras: removed syntax \<oplus> and \<otimes>, in favour of plain + and *
krauss
parents: 47444
diff changeset
  9356
  shows "closure S + closure T \<subseteq> closure (S + T)"
55928
2d7582309d73 generalize lemma closure_sum
huffman
parents: 55787
diff changeset
  9357
  unfolding set_plus_image closure_Times [symmetric] split_def
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  9358
  by (intro closure_bounded_linear_image_subset bounded_linear_add
55928
2d7582309d73 generalize lemma closure_sum
huffman
parents: 55787
diff changeset
  9359
    bounded_linear_fst bounded_linear_snd)
40887
ee8d0548c148 Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents: 40719
diff changeset
  9360
ee8d0548c148 Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents: 40719
diff changeset
  9361
lemma rel_interior_sum:
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9362
  fixes S T :: "'n::euclidean_space set"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9363
  assumes "convex S"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9364
    and "convex T"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9365
  shows "rel_interior (S + T) = rel_interior S + rel_interior T"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9366
proof -
55787
41a73a41f6c8 more symbols;
wenzelm
parents: 55697
diff changeset
  9367
  have "rel_interior S + rel_interior T = (\<lambda>(x,y). x + y) ` (rel_interior S \<times> rel_interior T)"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9368
    by (simp add: set_plus_image)
55787
41a73a41f6c8 more symbols;
wenzelm
parents: 55697
diff changeset
  9369
  also have "\<dots> = (\<lambda>(x,y). x + y) ` rel_interior (S \<times> T)"
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  9370
    using rel_interior_Times assms by auto
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9371
  also have "\<dots> = rel_interior (S + T)"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9372
    using fst_snd_linear convex_Times assms
55787
41a73a41f6c8 more symbols;
wenzelm
parents: 55697
diff changeset
  9373
      rel_interior_convex_linear_image[of "(\<lambda>(x,y). x + y)" "S \<times> T"]
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9374
    by (auto simp add: set_plus_image)
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9375
  finally show ?thesis ..
40887
ee8d0548c148 Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents: 40719
diff changeset
  9376
qed
ee8d0548c148 Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents: 40719
diff changeset
  9377
ee8d0548c148 Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents: 40719
diff changeset
  9378
lemma rel_interior_sum_gen:
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9379
  fixes S :: "'a \<Rightarrow> 'n::euclidean_space set"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9380
  assumes "\<forall>i\<in>I. convex (S i)"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9381
  shows "rel_interior (setsum S I) = setsum (\<lambda>i. rel_interior (S i)) I"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9382
  apply (subst setsum_set_cond_linear[of convex])
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9383
  using rel_interior_sum rel_interior_sing[of "0"] assms
55929
91f245c23bc5 remove lemmas in favor of more general ones: convex(_hull)_set_{plus,setsum}
huffman
parents: 55928
diff changeset
  9384
  apply (auto simp add: convex_set_plus)
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9385
  done
40887
ee8d0548c148 Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents: 40719
diff changeset
  9386
ee8d0548c148 Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents: 40719
diff changeset
  9387
lemma convex_rel_open_direct_sum:
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9388
  fixes S T :: "'n::euclidean_space set"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9389
  assumes "convex S"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9390
    and "rel_open S"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9391
    and "convex T"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9392
    and "rel_open T"
55787
41a73a41f6c8 more symbols;
wenzelm
parents: 55697
diff changeset
  9393
  shows "convex (S \<times> T) \<and> rel_open (S \<times> T)"
63007
aa894a49f77d new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents: 62950
diff changeset
  9394
  by (metis assms convex_Times rel_interior_Times rel_open_def)
40887
ee8d0548c148 Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents: 40719
diff changeset
  9395
ee8d0548c148 Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents: 40719
diff changeset
  9396
lemma convex_rel_open_sum:
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9397
  fixes S T :: "'n::euclidean_space set"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9398
  assumes "convex S"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9399
    and "rel_open S"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9400
    and "convex T"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9401
    and "rel_open T"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9402
  shows "convex (S + T) \<and> rel_open (S + T)"
55929
91f245c23bc5 remove lemmas in favor of more general ones: convex(_hull)_set_{plus,setsum}
huffman
parents: 55928
diff changeset
  9403
  by (metis assms convex_set_plus rel_interior_sum rel_open_def)
40887
ee8d0548c148 Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents: 40719
diff changeset
  9404
ee8d0548c148 Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents: 40719
diff changeset
  9405
lemma convex_hull_finite_union_cones:
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9406
  assumes "finite I"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9407
    and "I \<noteq> {}"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9408
  assumes "\<forall>i\<in>I. convex (S i) \<and> cone (S i) \<and> S i \<noteq> {}"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9409
  shows "convex hull (\<Union>(S ` I)) = setsum S I"
40887
ee8d0548c148 Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents: 40719
diff changeset
  9410
  (is "?lhs = ?rhs")
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9411
proof -
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9412
  {
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9413
    fix x
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9414
    assume "x \<in> ?lhs"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9415
    then obtain c xs where
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9416
      x: "x = setsum (\<lambda>i. c i *\<^sub>R xs i) I \<and> (\<forall>i\<in>I. c i \<ge> 0) \<and> setsum c I = 1 \<and> (\<forall>i\<in>I. xs i \<in> S i)"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9417
      using convex_hull_finite_union[of I S] assms by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  9418
    define s where "s i = c i *\<^sub>R xs i" for i
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9419
    {
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9420
      fix i
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9421
      assume "i \<in> I"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9422
      then have "s i \<in> S i"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9423
        using s_def x assms mem_cone[of "S i" "xs i" "c i"] by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9424
    }
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9425
    then have "\<forall>i\<in>I. s i \<in> S i" by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9426
    moreover have "x = setsum s I" using x s_def by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9427
    ultimately have "x \<in> ?rhs"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9428
      using set_setsum_alt[of I S] assms by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9429
  }
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9430
  moreover
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9431
  {
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9432
    fix x
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9433
    assume "x \<in> ?rhs"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9434
    then obtain s where x: "x = setsum s I \<and> (\<forall>i\<in>I. s i \<in> S i)"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9435
      using set_setsum_alt[of I S] assms by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  9436
    define xs where "xs i = of_nat(card I) *\<^sub>R s i" for i
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9437
    then have "x = setsum (\<lambda>i. ((1 :: real) / of_nat(card I)) *\<^sub>R xs i) I"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9438
      using x assms by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9439
    moreover have "\<forall>i\<in>I. xs i \<in> S i"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9440
      using x xs_def assms by (simp add: cone_def)
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9441
    moreover have "\<forall>i\<in>I. (1 :: real) / of_nat (card I) \<ge> 0"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9442
      by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9443
    moreover have "setsum (\<lambda>i. (1 :: real) / of_nat (card I)) I = 1"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9444
      using assms by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9445
    ultimately have "x \<in> ?lhs"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9446
      apply (subst convex_hull_finite_union[of I S])
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9447
      using assms
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9448
      apply blast
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9449
      using assms
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9450
      apply blast
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9451
      apply rule
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9452
      apply (rule_tac x = "(\<lambda>i. (1 :: real) / of_nat (card I))" in exI)
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9453
      apply auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9454
      done
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9455
  }
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9456
  ultimately show ?thesis by auto
40887
ee8d0548c148 Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents: 40719
diff changeset
  9457
qed
ee8d0548c148 Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents: 40719
diff changeset
  9458
ee8d0548c148 Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents: 40719
diff changeset
  9459
lemma convex_hull_union_cones_two:
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9460
  fixes S T :: "'m::euclidean_space set"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9461
  assumes "convex S"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9462
    and "cone S"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9463
    and "S \<noteq> {}"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9464
  assumes "convex T"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9465
    and "cone T"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9466
    and "T \<noteq> {}"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9467
  shows "convex hull (S \<union> T) = S + T"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9468
proof -
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  9469
  define I :: "nat set" where "I = {1, 2}"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  9470
  define A where "A i = (if i = (1::nat) then S else T)" for i
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9471
  have "\<Union>(A ` I) = S \<union> T"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9472
    using A_def I_def by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9473
  then have "convex hull (\<Union>(A ` I)) = convex hull (S \<union> T)"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9474
    by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9475
  moreover have "convex hull \<Union>(A ` I) = setsum A I"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9476
    apply (subst convex_hull_finite_union_cones[of I A])
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9477
    using assms A_def I_def
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9478
    apply auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9479
    done
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9480
  moreover have "setsum A I = S + T"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9481
    using A_def I_def
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9482
    unfolding set_plus_def
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9483
    apply auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9484
    unfolding set_plus_def
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9485
    apply auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9486
    done
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9487
  ultimately show ?thesis by auto
40887
ee8d0548c148 Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents: 40719
diff changeset
  9488
qed
ee8d0548c148 Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents: 40719
diff changeset
  9489
ee8d0548c148 Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents: 40719
diff changeset
  9490
lemma rel_interior_convex_hull_union:
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9491
  fixes S :: "'a \<Rightarrow> 'n::euclidean_space set"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9492
  assumes "finite I"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9493
    and "\<forall>i\<in>I. convex (S i) \<and> S i \<noteq> {}"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9494
  shows "rel_interior (convex hull (\<Union>(S ` I))) =
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9495
    {setsum (\<lambda>i. c i *\<^sub>R s i) I | c s. (\<forall>i\<in>I. c i > 0) \<and> setsum c I = 1 \<and>
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9496
      (\<forall>i\<in>I. s i \<in> rel_interior(S i))}"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9497
  (is "?lhs = ?rhs")
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9498
proof (cases "I = {}")
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9499
  case True
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9500
  then show ?thesis
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9501
    using convex_hull_empty rel_interior_empty by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9502
next
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9503
  case False
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  9504
  define C0 where "C0 = convex hull (\<Union>(S ` I))"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9505
  have "\<forall>i\<in>I. C0 \<ge> S i"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9506
    unfolding C0_def using hull_subset[of "\<Union>(S ` I)"] by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  9507
  define K0 where "K0 = cone hull ({1 :: real} \<times> C0)"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  9508
  define K where "K i = cone hull ({1 :: real} \<times> S i)" for i
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9509
  have "\<forall>i\<in>I. K i \<noteq> {}"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9510
    unfolding K_def using assms
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9511
    by (simp add: cone_hull_empty_iff[symmetric])
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9512
  {
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9513
    fix i
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9514
    assume "i \<in> I"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9515
    then have "convex (K i)"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9516
      unfolding K_def
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9517
      apply (subst convex_cone_hull)
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9518
      apply (subst convex_Times)
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9519
      using assms
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9520
      apply auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9521
      done
40887
ee8d0548c148 Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents: 40719
diff changeset
  9522
  }
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9523
  then have convK: "\<forall>i\<in>I. convex (K i)"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9524
    by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9525
  {
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9526
    fix i
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9527
    assume "i \<in> I"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9528
    then have "K0 \<supseteq> K i"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9529
      unfolding K0_def K_def
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9530
      apply (subst hull_mono)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  9531
      using \<open>\<forall>i\<in>I. C0 \<ge> S i\<close>
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9532
      apply auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9533
      done
40887
ee8d0548c148 Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents: 40719
diff changeset
  9534
  }
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9535
  then have "K0 \<supseteq> \<Union>(K ` I)" by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9536
  moreover have "convex K0"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9537
    unfolding K0_def
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9538
    apply (subst convex_cone_hull)
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9539
    apply (subst convex_Times)
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9540
    unfolding C0_def
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9541
    using convex_convex_hull
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9542
    apply auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9543
    done
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9544
  ultimately have geq: "K0 \<supseteq> convex hull (\<Union>(K ` I))"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9545
    using hull_minimal[of _ "K0" "convex"] by blast
55787
41a73a41f6c8 more symbols;
wenzelm
parents: 55697
diff changeset
  9546
  have "\<forall>i\<in>I. K i \<supseteq> {1 :: real} \<times> S i"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9547
    using K_def by (simp add: hull_subset)
55787
41a73a41f6c8 more symbols;
wenzelm
parents: 55697
diff changeset
  9548
  then have "\<Union>(K ` I) \<supseteq> {1 :: real} \<times> \<Union>(S ` I)"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9549
    by auto
55787
41a73a41f6c8 more symbols;
wenzelm
parents: 55697
diff changeset
  9550
  then have "convex hull \<Union>(K ` I) \<supseteq> convex hull ({1 :: real} \<times> \<Union>(S ` I))"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9551
    by (simp add: hull_mono)
55787
41a73a41f6c8 more symbols;
wenzelm
parents: 55697
diff changeset
  9552
  then have "convex hull \<Union>(K ` I) \<supseteq> {1 :: real} \<times> C0"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9553
    unfolding C0_def
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9554
    using convex_hull_Times[of "{(1 :: real)}" "\<Union>(S ` I)"] convex_hull_singleton
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9555
    by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9556
  moreover have "cone (convex hull (\<Union>(K ` I)))"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9557
    apply (subst cone_convex_hull)
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9558
    using cone_Union[of "K ` I"]
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9559
    apply auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9560
    unfolding K_def
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9561
    using cone_cone_hull
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9562
    apply auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9563
    done
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9564
  ultimately have "convex hull (\<Union>(K ` I)) \<supseteq> K0"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9565
    unfolding K0_def
60585
48fdff264eb2 tuned whitespace;
wenzelm
parents: 60420
diff changeset
  9566
    using hull_minimal[of _ "convex hull (\<Union>(K ` I))" "cone"]
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9567
    by blast
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9568
  then have "K0 = convex hull (\<Union>(K ` I))"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9569
    using geq by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9570
  also have "\<dots> = setsum K I"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9571
    apply (subst convex_hull_finite_union_cones[of I K])
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9572
    using assms
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9573
    apply blast
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9574
    using False
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9575
    apply blast
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9576
    unfolding K_def
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9577
    apply rule
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9578
    apply (subst convex_cone_hull)
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9579
    apply (subst convex_Times)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  9580
    using assms cone_cone_hull \<open>\<forall>i\<in>I. K i \<noteq> {}\<close> K_def
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9581
    apply auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9582
    done
47444
d21c95af2df7 removed "setsum_set", now subsumed by generic setsum
krauss
parents: 47108
diff changeset
  9583
  finally have "K0 = setsum K I" by auto
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9584
  then have *: "rel_interior K0 = setsum (\<lambda>i. (rel_interior (K i))) I"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9585
    using rel_interior_sum_gen[of I K] convK by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9586
  {
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9587
    fix x
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9588
    assume "x \<in> ?lhs"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9589
    then have "(1::real, x) \<in> rel_interior K0"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9590
      using K0_def C0_def rel_interior_convex_cone_aux[of C0 "1::real" x] convex_convex_hull
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9591
      by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9592
    then obtain k where k: "(1::real, x) = setsum k I \<and> (\<forall>i\<in>I. k i \<in> rel_interior (K i))"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  9593
      using \<open>finite I\<close> * set_setsum_alt[of I "\<lambda>i. rel_interior (K i)"] by auto
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9594
    {
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9595
      fix i
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9596
      assume "i \<in> I"
55787
41a73a41f6c8 more symbols;
wenzelm
parents: 55697
diff changeset
  9597
      then have "convex (S i) \<and> k i \<in> rel_interior (cone hull {1} \<times> S i)"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9598
        using k K_def assms by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9599
      then have "\<exists>ci si. k i = (ci, ci *\<^sub>R si) \<and> 0 < ci \<and> si \<in> rel_interior (S i)"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9600
        using rel_interior_convex_cone[of "S i"] by auto
40887
ee8d0548c148 Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents: 40719
diff changeset
  9601
    }
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9602
    then obtain c s where
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9603
      cs: "\<forall>i\<in>I. k i = (c i, c i *\<^sub>R s i) \<and> 0 < c i \<and> s i \<in> rel_interior (S i)"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9604
      by metis
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9605
    then have "x = (\<Sum>i\<in>I. c i *\<^sub>R s i) \<and> setsum c I = 1"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9606
      using k by (simp add: setsum_prod)
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9607
    then have "x \<in> ?rhs"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9608
      using k
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9609
      apply auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9610
      apply (rule_tac x = c in exI)
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9611
      apply (rule_tac x = s in exI)
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9612
      using cs
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9613
      apply auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9614
      done
40887
ee8d0548c148 Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents: 40719
diff changeset
  9615
  }
ee8d0548c148 Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents: 40719
diff changeset
  9616
  moreover
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9617
  {
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9618
    fix x
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9619
    assume "x \<in> ?rhs"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9620
    then obtain c s where cs: "x = setsum (\<lambda>i. c i *\<^sub>R s i) I \<and>
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9621
        (\<forall>i\<in>I. c i > 0) \<and> setsum c I = 1 \<and> (\<forall>i\<in>I. s i \<in> rel_interior (S i))"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9622
      by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  9623
    define k where "k i = (c i, c i *\<^sub>R s i)" for i
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9624
    {
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9625
      fix i assume "i:I"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9626
      then have "k i \<in> rel_interior (K i)"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9627
        using k_def K_def assms cs rel_interior_convex_cone[of "S i"]
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9628
        by auto
40887
ee8d0548c148 Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents: 40719
diff changeset
  9629
    }
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9630
    then have "(1::real, x) \<in> rel_interior K0"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9631
      using K0_def * set_setsum_alt[of I "(\<lambda>i. rel_interior (K i))"] assms k_def cs
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9632
      apply auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9633
      apply (rule_tac x = k in exI)
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9634
      apply (simp add: setsum_prod)
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9635
      done
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9636
    then have "x \<in> ?lhs"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9637
      using K0_def C0_def rel_interior_convex_cone_aux[of C0 1 x]
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9638
      by (auto simp add: convex_convex_hull)
40887
ee8d0548c148 Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents: 40719
diff changeset
  9639
  }
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9640
  ultimately show ?thesis by blast
40887
ee8d0548c148 Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents: 40719
diff changeset
  9641
qed
ee8d0548c148 Prove rel_interior_convex_hull_union (by Grechuck Bogdan).
hoelzl
parents: 40719
diff changeset
  9642
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9643
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9644
lemma convex_le_Inf_differential:
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9645
  fixes f :: "real \<Rightarrow> real"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9646
  assumes "convex_on I f"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9647
    and "x \<in> interior I"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9648
    and "y \<in> I"
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9649
  shows "f y \<ge> f x + Inf ((\<lambda>t. (f x - f t) / (x - t)) ` ({x<..} \<inter> I)) * (y - x)"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9650
  (is "_ \<ge> _ + Inf (?F x) * (y - x)")
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9651
proof (cases rule: linorder_cases)
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9652
  assume "x < y"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9653
  moreover
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9654
  have "open (interior I)" by auto
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  9655
  from openE[OF this \<open>x \<in> interior I\<close>]
55697
abec82f4e3e9 tuned proofs;
wenzelm
parents: 54780
diff changeset
  9656
  obtain e where e: "0 < e" "ball x e \<subseteq> interior I" .
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  9657
  moreover define t where "t = min (x + e / 2) ((x + y) / 2)"
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9658
  ultimately have "x < t" "t < y" "t \<in> ball x e"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9659
    by (auto simp: dist_real_def field_simps split: split_min)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  9660
  with \<open>x \<in> interior I\<close> e interior_subset[of I] have "t \<in> I" "x \<in> I" by auto
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9661
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9662
  have "open (interior I)" by auto
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  9663
  from openE[OF this \<open>x \<in> interior I\<close>]
55697
abec82f4e3e9 tuned proofs;
wenzelm
parents: 54780
diff changeset
  9664
  obtain e where "0 < e" "ball x e \<subseteq> interior I" .
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  9665
  moreover define K where "K = x - e / 2"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  9666
  with \<open>0 < e\<close> have "K \<in> ball x e" "K < x"
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9667
    by (auto simp: dist_real_def)
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9668
  ultimately have "K \<in> I" "K < x" "x \<in> I"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  9669
    using interior_subset[of I] \<open>x \<in> interior I\<close> by auto
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9670
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9671
  have "Inf (?F x) \<le> (f x - f y) / (x - y)"
54258
adfc759263ab use bdd_above and bdd_below for conditionally complete lattices
hoelzl
parents: 54230
diff changeset
  9672
  proof (intro bdd_belowI cInf_lower2)
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9673
    show "(f x - f t) / (x - t) \<in> ?F x"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  9674
      using \<open>t \<in> I\<close> \<open>x < t\<close> by auto
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9675
    show "(f x - f t) / (x - t) \<le> (f x - f y) / (x - y)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  9676
      using \<open>convex_on I f\<close> \<open>x \<in> I\<close> \<open>y \<in> I\<close> \<open>x < t\<close> \<open>t < y\<close>
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9677
      by (rule convex_on_diff)
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9678
  next
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9679
    fix y
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9680
    assume "y \<in> ?F x"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  9681
    with order_trans[OF convex_on_diff[OF \<open>convex_on I f\<close> \<open>K \<in> I\<close> _ \<open>K < x\<close> _]]
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9682
    show "(f K - f x) / (K - x) \<le> y" by auto
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9683
  qed
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9684
  then show ?thesis
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  9685
    using \<open>x < y\<close> by (simp add: field_simps)
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9686
next
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9687
  assume "y < x"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9688
  moreover
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9689
  have "open (interior I)" by auto
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  9690
  from openE[OF this \<open>x \<in> interior I\<close>]
55697
abec82f4e3e9 tuned proofs;
wenzelm
parents: 54780
diff changeset
  9691
  obtain e where e: "0 < e" "ball x e \<subseteq> interior I" .
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
  9692
  moreover define t where "t = x + e / 2"
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9693
  ultimately have "x < t" "t \<in> ball x e"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9694
    by (auto simp: dist_real_def field_simps)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  9695
  with \<open>x \<in> interior I\<close> e interior_subset[of I] have "t \<in> I" "x \<in> I" by auto
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9696
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9697
  have "(f x - f y) / (x - y) \<le> Inf (?F x)"
51475
ebf9d4fd00ba introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
hoelzl
parents: 50979
diff changeset
  9698
  proof (rule cInf_greatest)
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9699
    have "(f x - f y) / (x - y) = (f y - f x) / (y - x)"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  9700
      using \<open>y < x\<close> by (auto simp: field_simps)
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9701
    also
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9702
    fix z
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9703
    assume "z \<in> ?F x"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  9704
    with order_trans[OF convex_on_diff[OF \<open>convex_on I f\<close> \<open>y \<in> I\<close> _ \<open>y < x\<close>]]
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9705
    have "(f y - f x) / (y - x) \<le> z"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9706
      by auto
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9707
    finally show "(f x - f y) / (x - y) \<le> z" .
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9708
  next
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9709
    have "open (interior I)" by auto
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  9710
    from openE[OF this \<open>x \<in> interior I\<close>]
55697
abec82f4e3e9 tuned proofs;
wenzelm
parents: 54780
diff changeset
  9711
    obtain e where e: "0 < e" "ball x e \<subseteq> interior I" .
54465
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9712
    then have "x + e / 2 \<in> ball x e"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9713
      by (auto simp: dist_real_def)
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9714
    with e interior_subset[of I] have "x + e / 2 \<in> {x<..} \<inter> I"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9715
      by auto
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9716
    then show "?F x \<noteq> {}"
2f7867850cc3 tuned proofs;
wenzelm
parents: 54263
diff changeset
  9717
      by blast
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9718
  qed
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9719
  then show ?thesis
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  9720
    using \<open>y < x\<close> by (simp add: field_simps)
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
  9721
qed simp
60762
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
  9722
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  9723
subsection\<open>Explicit formulas for interior and relative interior of convex hull\<close>
60762
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
  9724
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
  9725
lemma interior_atLeastAtMost [simp]:
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
  9726
  fixes a::real shows "interior {a..b} = {a<..<b}"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
  9727
  using interior_cbox [of a b] by auto
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
  9728
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
  9729
lemma interior_atLeastLessThan [simp]:
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
  9730
  fixes a::real shows "interior {a..<b} = {a<..<b}"
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  9731
  by (metis atLeastLessThan_def greaterThanLessThan_def interior_atLeastAtMost interior_Int interior_interior interior_real_semiline)
60762
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
  9732
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
  9733
lemma interior_lessThanAtMost [simp]:
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
  9734
  fixes a::real shows "interior {a<..b} = {a<..<b}"
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  9735
  by (metis atLeastAtMost_def greaterThanAtMost_def interior_atLeastAtMost interior_Int
60762
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
  9736
            interior_interior interior_real_semiline)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
  9737
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
  9738
lemma at_within_closed_interval:
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
  9739
    fixes x::real
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
  9740
    shows "a < x \<Longrightarrow> x < b \<Longrightarrow> (at x within {a..b}) = at x"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
  9741
  by (metis at_within_interior greaterThanLessThan_iff interior_atLeastAtMost)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
  9742
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9743
lemma affine_independent_convex_affine_hull:
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9744
  fixes s :: "'a::euclidean_space set"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9745
  assumes "~affine_dependent s" "t \<subseteq> s"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9746
    shows "convex hull t = affine hull t \<inter> convex hull s"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9747
proof -
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9748
  have fin: "finite s" "finite t" using assms aff_independent_finite finite_subset by auto
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9749
    { fix u v x
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
  9750
      assume uv: "setsum u t = 1" "\<forall>x\<in>s. 0 \<le> v x" "setsum v s = 1"
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9751
                 "(\<Sum>x\<in>s. v x *\<^sub>R x) = (\<Sum>v\<in>t. u v *\<^sub>R v)" "x \<in> t"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61762
diff changeset
  9752
      then have s: "s = (s - t) \<union> t" \<comment>\<open>split into separate cases\<close>
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9753
        using assms by auto
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9754
      have [simp]: "(\<Sum>x\<in>t. v x *\<^sub>R x) + (\<Sum>x\<in>s - t. v x *\<^sub>R x) = (\<Sum>x\<in>t. u x *\<^sub>R x)"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9755
                   "setsum v t + setsum v (s - t) = 1"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9756
        using uv fin s
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
  9757
        by (auto simp: setsum.union_disjoint [symmetric] Un_commute)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
  9758
      have "(\<Sum>x\<in>s. if x \<in> t then v x - u x else v x) = 0"
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9759
           "(\<Sum>x\<in>s. (if x \<in> t then v x - u x else v x) *\<^sub>R x) = 0"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9760
        using uv fin
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9761
        by (subst s, subst setsum.union_disjoint, auto simp: algebra_simps setsum_subtractf)+
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9762
    } note [simp] = this
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
  9763
  have "convex hull t \<subseteq> affine hull t"
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9764
    using convex_hull_subset_affine_hull by blast
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9765
  moreover have "convex hull t \<subseteq> convex hull s"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9766
    using assms hull_mono by blast
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9767
  moreover have "affine hull t \<inter> convex hull s \<subseteq> convex hull t"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
  9768
    using assms
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9769
    apply (simp add: convex_hull_finite affine_hull_finite fin affine_dependent_explicit)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9770
    apply (drule_tac x=s in spec)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9771
    apply (auto simp: fin)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9772
    apply (rule_tac x=u in exI)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9773
    apply (rename_tac v)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9774
    apply (drule_tac x="\<lambda>x. if x \<in> t then v x - u x else v x" in spec)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9775
    apply (force)+
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9776
    done
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9777
  ultimately show ?thesis
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9778
    by blast
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9779
qed
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9780
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
  9781
lemma affine_independent_span_eq:
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9782
  fixes s :: "'a::euclidean_space set"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9783
  assumes "~affine_dependent s" "card s = Suc (DIM ('a))"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9784
    shows "affine hull s = UNIV"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9785
proof (cases "s = {}")
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9786
  case True then show ?thesis
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9787
    using assms by simp
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9788
next
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9789
  case False
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9790
    then obtain a t where t: "a \<notin> t" "s = insert a t"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9791
      by blast
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
  9792
    then have fin: "finite t" using assms
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9793
      by (metis finite_insert aff_independent_finite)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9794
    show ?thesis
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9795
    using assms t fin
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9796
      apply (simp add: affine_dependent_iff_dependent affine_hull_insert_span_gen)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9797
      apply (rule subset_antisym)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9798
      apply force
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9799
      apply (rule Fun.vimage_subsetD)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9800
      apply (metis add.commute diff_add_cancel surj_def)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9801
      apply (rule card_ge_dim_independent)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9802
      apply (auto simp: card_image inj_on_def dim_subset_UNIV)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9803
      done
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9804
qed
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9805
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
  9806
lemma affine_independent_span_gt:
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9807
  fixes s :: "'a::euclidean_space set"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9808
  assumes ind: "~ affine_dependent s" and dim: "DIM ('a) < card s"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9809
    shows "affine hull s = UNIV"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9810
  apply (rule affine_independent_span_eq [OF ind])
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9811
  apply (rule antisym)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9812
  using assms
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9813
  apply auto
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9814
  apply (metis add_2_eq_Suc' not_less_eq_eq affine_dependent_biggerset aff_independent_finite)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9815
  done
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9816
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
  9817
lemma empty_interior_affine_hull:
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9818
  fixes s :: "'a::euclidean_space set"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9819
  assumes "finite s" and dim: "card s \<le> DIM ('a)"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9820
    shows "interior(affine hull s) = {}"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9821
  using assms
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9822
  apply (induct s rule: finite_induct)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9823
  apply (simp_all add:  affine_dependent_iff_dependent affine_hull_insert_span_gen interior_translation)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9824
  apply (rule empty_interior_lowdim)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9825
  apply (simp add: affine_dependent_iff_dependent affine_hull_insert_span_gen)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9826
  apply (metis Suc_le_lessD not_less order_trans card_image_le finite_imageI dim_le_card)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9827
  done
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9828
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
  9829
lemma empty_interior_convex_hull:
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9830
  fixes s :: "'a::euclidean_space set"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9831
  assumes "finite s" and dim: "card s \<le> DIM ('a)"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9832
    shows "interior(convex hull s) = {}"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
  9833
  by (metis Diff_empty Diff_eq_empty_iff convex_hull_subset_affine_hull
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9834
            interior_mono empty_interior_affine_hull [OF assms])
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9835
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9836
lemma explicit_subset_rel_interior_convex_hull:
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9837
  fixes s :: "'a::euclidean_space set"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
  9838
  shows "finite s
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9839
         \<Longrightarrow> {y. \<exists>u. (\<forall>x \<in> s. 0 < u x \<and> u x < 1) \<and> setsum u s = 1 \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y}
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9840
             \<subseteq> rel_interior (convex hull s)"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9841
  by (force simp add:  rel_interior_convex_hull_union [where S="\<lambda>x. {x}" and I=s, simplified])
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9842
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
  9843
lemma explicit_subset_rel_interior_convex_hull_minimal:
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9844
  fixes s :: "'a::euclidean_space set"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
  9845
  shows "finite s
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9846
         \<Longrightarrow> {y. \<exists>u. (\<forall>x \<in> s. 0 < u x) \<and> setsum u s = 1 \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y}
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9847
             \<subseteq> rel_interior (convex hull s)"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9848
  by (force simp add:  rel_interior_convex_hull_union [where S="\<lambda>x. {x}" and I=s, simplified])
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9849
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
  9850
lemma rel_interior_convex_hull_explicit:
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9851
  fixes s :: "'a::euclidean_space set"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9852
  assumes "~ affine_dependent s"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9853
  shows "rel_interior(convex hull s) =
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9854
         {y. \<exists>u. (\<forall>x \<in> s. 0 < u x) \<and> setsum u s = 1 \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y}"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
  9855
         (is "?lhs = ?rhs")
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9856
proof
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9857
  show "?rhs \<le> ?lhs"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9858
    by (simp add: aff_independent_finite explicit_subset_rel_interior_convex_hull_minimal assms)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9859
next
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9860
  show "?lhs \<le> ?rhs"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9861
  proof (cases "\<exists>a. s = {a}")
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9862
    case True then show "?lhs \<le> ?rhs"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9863
      by force
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9864
  next
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9865
    case False
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9866
    have fs: "finite s"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9867
      using assms by (simp add: aff_independent_finite)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9868
    { fix a b and d::real
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9869
      assume ab: "a \<in> s" "b \<in> s" "a \<noteq> b"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61762
diff changeset
  9870
      then have s: "s = (s - {a,b}) \<union> {a,b}" \<comment>\<open>split into separate cases\<close>
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9871
        by auto
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
  9872
      have "(\<Sum>x\<in>s. if x = a then - d else if x = b then d else 0) = 0"
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9873
           "(\<Sum>x\<in>s. (if x = a then - d else if x = b then d else 0) *\<^sub>R x) = d *\<^sub>R b - d *\<^sub>R a"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9874
        using ab fs
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9875
        by (subst s, subst setsum.union_disjoint, auto)+
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9876
    } note [simp] = this
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9877
    { fix y
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9878
      assume y: "y \<in> convex hull s" "y \<notin> ?rhs"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9879
      { fix u T a
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9880
        assume ua: "\<forall>x\<in>s. 0 \<le> u x" "setsum u s = 1" "\<not> 0 < u a" "a \<in> s"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9881
           and yT: "y = (\<Sum>x\<in>s. u x *\<^sub>R x)" "y \<in> T" "open T"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
  9882
           and sb: "T \<inter> affine hull s \<subseteq> {w. \<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> (\<Sum>x\<in>s. u x *\<^sub>R x) = w}"
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9883
        have ua0: "u a = 0"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9884
          using ua by auto
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9885
        obtain b where b: "b\<in>s" "a \<noteq> b"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9886
          using ua False by auto
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9887
        obtain e where e: "0 < e" "ball (\<Sum>x\<in>s. u x *\<^sub>R x) e \<subseteq> T"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9888
          using yT by (auto elim: openE)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9889
        with b obtain d where d: "0 < d" "norm(d *\<^sub>R (a-b)) < e"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9890
          by (auto intro: that [of "e / 2 / norm(a-b)"])
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9891
        have "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> affine hull s"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9892
          using yT y by (metis affine_hull_convex_hull hull_redundant_eq)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9893
        then have "(\<Sum>x\<in>s. u x *\<^sub>R x) - d *\<^sub>R (a - b) \<in> affine hull s"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9894
          using ua b by (auto simp: hull_inc intro: mem_affine_3_minus2)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9895
        then have "y - d *\<^sub>R (a - b) \<in> T \<inter> affine hull s"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9896
          using d e yT by auto
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
  9897
        then obtain v where "\<forall>x\<in>s. 0 \<le> v x"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
  9898
                            "setsum v s = 1"
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9899
                            "(\<Sum>x\<in>s. v x *\<^sub>R x) = (\<Sum>x\<in>s. u x *\<^sub>R x) - d *\<^sub>R (a - b)"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9900
          using subsetD [OF sb] yT
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9901
          by auto
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9902
        then have False
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
  9903
          using assms
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9904
          apply (simp add: affine_dependent_explicit_finite fs)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9905
          apply (drule_tac x="\<lambda>x. (v x - u x) - (if x = a then -d else if x = b then d else 0)" in spec)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9906
          using ua b d
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9907
          apply (auto simp: algebra_simps setsum_subtractf setsum.distrib)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9908
          done
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9909
      } note * = this
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9910
      have "y \<notin> rel_interior (convex hull s)"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9911
        using y
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9912
        apply (simp add: mem_rel_interior affine_hull_convex_hull)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9913
        apply (auto simp: convex_hull_finite [OF fs])
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9914
        apply (drule_tac x=u in spec)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9915
        apply (auto intro: *)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9916
        done
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9917
    } with rel_interior_subset show "?lhs \<le> ?rhs"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9918
      by blast
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9919
  qed
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9920
qed
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9921
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9922
lemma interior_convex_hull_explicit_minimal:
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9923
  fixes s :: "'a::euclidean_space set"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9924
  shows
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9925
   "~ affine_dependent s
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9926
        ==> interior(convex hull s) =
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9927
             (if card(s) \<le> DIM('a) then {}
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9928
              else {y. \<exists>u. (\<forall>x \<in> s. 0 < u x) \<and> setsum u s = 1 \<and> (\<Sum>x\<in>s. u x *\<^sub>R x) = y})"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9929
  apply (simp add: aff_independent_finite empty_interior_convex_hull, clarify)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9930
  apply (rule trans [of _ "rel_interior(convex hull s)"])
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9931
  apply (simp add: affine_hull_convex_hull affine_independent_span_gt rel_interior_interior)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9932
  by (simp add: rel_interior_convex_hull_explicit)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9933
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9934
lemma interior_convex_hull_explicit:
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9935
  fixes s :: "'a::euclidean_space set"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9936
  assumes "~ affine_dependent s"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9937
  shows
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9938
   "interior(convex hull s) =
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9939
             (if card(s) \<le> DIM('a) then {}
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9940
              else {y. \<exists>u. (\<forall>x \<in> s. 0 < u x \<and> u x < 1) \<and> setsum u s = 1 \<and> (\<Sum>x\<in>s. u x *\<^sub>R x) = y})"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9941
proof -
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9942
  { fix u :: "'a \<Rightarrow> real" and a
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9943
    assume "card Basis < card s" and u: "\<And>x. x\<in>s \<Longrightarrow> 0 < u x" "setsum u s = 1" and a: "a \<in> s"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9944
    then have cs: "Suc 0 < card s"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9945
      by (metis DIM_positive less_trans_Suc)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9946
    obtain b where b: "b \<in> s" "a \<noteq> b"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9947
    proof (cases "s \<le> {a}")
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9948
      case True
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9949
      then show thesis
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9950
        using cs subset_singletonD by fastforce
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9951
    next
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9952
      case False
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9953
      then show thesis
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9954
      by (blast intro: that)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
  9955
    qed
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9956
    have "u a + u b \<le> setsum u {a,b}"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9957
      using a b by simp
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9958
    also have "... \<le> setsum u s"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9959
      apply (rule Groups_Big.setsum_mono2)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9960
      using a b u
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9961
      apply (auto simp: less_imp_le aff_independent_finite assms)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
  9962
      done
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9963
    finally have "u a < 1"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
  9964
      using \<open>b \<in> s\<close> u by fastforce
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9965
  } note [simp] = this
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9966
  show ?thesis
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9967
    using assms
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9968
    apply (auto simp: interior_convex_hull_explicit_minimal)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9969
    apply (rule_tac x=u in exI)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9970
    apply (auto simp: not_le)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9971
    done
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9972
qed
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
  9973
62620
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  9974
lemma interior_closed_segment_ge2:
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  9975
  fixes a :: "'a::euclidean_space"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  9976
  assumes "2 \<le> DIM('a)"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  9977
    shows  "interior(closed_segment a b) = {}"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  9978
using assms unfolding segment_convex_hull
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  9979
proof -
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  9980
  have "card {a, b} \<le> DIM('a)"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  9981
    using assms
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  9982
    by (simp add: card_insert_if linear not_less_eq_eq numeral_2_eq_2)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  9983
  then show "interior (convex hull {a, b}) = {}"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  9984
    by (metis empty_interior_convex_hull finite.insertI finite.emptyI)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  9985
qed
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  9986
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  9987
lemma interior_open_segment:
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  9988
  fixes a :: "'a::euclidean_space"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  9989
  shows  "interior(open_segment a b) =
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  9990
                 (if 2 \<le> DIM('a) then {} else open_segment a b)"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  9991
proof (simp add: not_le, intro conjI impI)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  9992
  assume "2 \<le> DIM('a)"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  9993
  then show "interior (open_segment a b) = {}"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  9994
    apply (simp add: segment_convex_hull open_segment_def)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  9995
    apply (metis Diff_subset interior_mono segment_convex_hull subset_empty interior_closed_segment_ge2)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  9996
    done
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  9997
next
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  9998
  assume le2: "DIM('a) < 2"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
  9999
  show "interior (open_segment a b) = open_segment a b"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10000
  proof (cases "a = b")
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10001
    case True then show ?thesis by auto
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10002
  next
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10003
    case False
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10004
    with le2 have "affine hull (open_segment a b) = UNIV"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10005
      apply simp
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10006
      apply (rule affine_independent_span_gt)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10007
      apply (simp_all add: affine_dependent_def insert_Diff_if)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10008
      done
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10009
    then show "interior (open_segment a b) = open_segment a b"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10010
      using rel_interior_interior rel_interior_open_segment by blast
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10011
  qed
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10012
qed
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10013
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10014
lemma interior_closed_segment:
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10015
  fixes a :: "'a::euclidean_space"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10016
  shows "interior(closed_segment a b) =
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10017
                 (if 2 \<le> DIM('a) then {} else open_segment a b)"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10018
proof (cases "a = b")
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10019
  case True then show ?thesis by simp
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10020
next
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10021
  case False
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10022
  then have "closure (open_segment a b) = closed_segment a b"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10023
    by simp
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10024
  then show ?thesis
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10025
    by (metis (no_types) convex_interior_closure convex_open_segment interior_open_segment)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10026
qed
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10027
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10028
lemmas interior_segment = interior_closed_segment interior_open_segment
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10029
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10030
lemma closed_segment_eq [simp]:
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10031
  fixes a :: "'a::euclidean_space"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10032
  shows "closed_segment a b = closed_segment c d \<longleftrightarrow> {a,b} = {c,d}"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10033
proof
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10034
  assume abcd: "closed_segment a b = closed_segment c d"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10035
  show "{a,b} = {c,d}"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10036
  proof (cases "a=b \<or> c=d")
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10037
    case True with abcd show ?thesis by force
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10038
  next
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10039
    case False
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10040
    then have neq: "a \<noteq> b \<and> c \<noteq> d" by force
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10041
    have *: "closed_segment c d - {a, b} = rel_interior (closed_segment c d)"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10042
      using neq abcd by (metis (no_types) open_segment_def rel_interior_closed_segment)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10043
    have "b \<in> {c, d}"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10044
    proof -
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10045
      have "insert b (closed_segment c d) = closed_segment c d"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10046
        using abcd by blast
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10047
      then show ?thesis
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10048
        by (metis DiffD2 Diff_insert2 False * insertI1 insert_Diff_if open_segment_def rel_interior_closed_segment)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10049
    qed
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10050
    moreover have "a \<in> {c, d}"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10051
      by (metis Diff_iff False * abcd ends_in_segment(1) insertI1 open_segment_def rel_interior_closed_segment)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10052
    ultimately show "{a, b} = {c, d}"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10053
      using neq by fastforce
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10054
  qed
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10055
next
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10056
  assume "{a,b} = {c,d}"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10057
  then show "closed_segment a b = closed_segment c d"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10058
    by (simp add: segment_convex_hull)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10059
qed
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10060
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10061
lemma closed_open_segment_eq [simp]:
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10062
  fixes a :: "'a::euclidean_space"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10063
  shows "closed_segment a b \<noteq> open_segment c d"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10064
by (metis DiffE closed_segment_neq_empty closure_closed_segment closure_open_segment ends_in_segment(1) insertI1 open_segment_def)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10065
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10066
lemma open_closed_segment_eq [simp]:
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10067
  fixes a :: "'a::euclidean_space"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10068
  shows "open_segment a b \<noteq> closed_segment c d"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10069
using closed_open_segment_eq by blast
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10070
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10071
lemma open_segment_eq [simp]:
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10072
  fixes a :: "'a::euclidean_space"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10073
  shows "open_segment a b = open_segment c d \<longleftrightarrow> a = b \<and> c = d \<or> {a,b} = {c,d}"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10074
        (is "?lhs = ?rhs")
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10075
proof
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10076
  assume abcd: ?lhs
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10077
  show ?rhs
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10078
  proof (cases "a=b \<or> c=d")
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10079
    case True with abcd show ?thesis
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10080
      using finite_open_segment by fastforce
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10081
  next
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10082
    case False
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10083
    then have a2: "a \<noteq> b \<and> c \<noteq> d" by force
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10084
    with abcd show ?rhs
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10085
      unfolding open_segment_def
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10086
      by (metis (no_types) abcd closed_segment_eq closure_open_segment)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10087
  qed
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10088
next
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10089
  assume ?rhs
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10090
  then show ?lhs
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10091
    by (metis Diff_cancel convex_hull_singleton insert_absorb2 open_segment_def segment_convex_hull)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10092
qed
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10093
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60307
diff changeset
 10094
subsection\<open>Similar results for closure and (relative or absolute) frontier.\<close>
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10095
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10096
lemma closure_convex_hull [simp]:
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10097
  fixes s :: "'a::euclidean_space set"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10098
  shows "compact s ==> closure(convex hull s) = convex hull s"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10099
  by (simp add: compact_imp_closed compact_convex_hull)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10100
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10101
lemma rel_frontier_convex_hull_explicit:
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10102
  fixes s :: "'a::euclidean_space set"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10103
  assumes "~ affine_dependent s"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10104
  shows "rel_frontier(convex hull s) =
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10105
         {y. \<exists>u. (\<forall>x \<in> s. 0 \<le> u x) \<and> (\<exists>x \<in> s. u x = 0) \<and> setsum u s = 1 \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y}"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10106
proof -
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10107
  have fs: "finite s"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10108
    using assms by (simp add: aff_independent_finite)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10109
  show ?thesis
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10110
    apply (simp add: rel_frontier_def finite_imp_compact rel_interior_convex_hull_explicit assms fs)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10111
    apply (auto simp: convex_hull_finite fs)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10112
    apply (drule_tac x=u in spec)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10113
    apply (rule_tac x=u in exI)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10114
    apply force
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10115
    apply (rename_tac v)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10116
    apply (rule notE [OF assms])
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10117
    apply (simp add: affine_dependent_explicit)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10118
    apply (rule_tac x=s in exI)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10119
    apply (auto simp: fs)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10120
    apply (rule_tac x = "\<lambda>x. u x - v x" in exI)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10121
    apply (force simp: setsum_subtractf scaleR_diff_left)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10122
    done
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10123
qed
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10124
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10125
lemma frontier_convex_hull_explicit:
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10126
  fixes s :: "'a::euclidean_space set"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10127
  assumes "~ affine_dependent s"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10128
  shows "frontier(convex hull s) =
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
 10129
         {y. \<exists>u. (\<forall>x \<in> s. 0 \<le> u x) \<and> (DIM ('a) < card s \<longrightarrow> (\<exists>x \<in> s. u x = 0)) \<and>
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10130
             setsum u s = 1 \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y}"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10131
proof -
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10132
  have fs: "finite s"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10133
    using assms by (simp add: aff_independent_finite)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10134
  show ?thesis
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10135
  proof (cases "DIM ('a) < card s")
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10136
    case True
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10137
    with assms fs show ?thesis
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
 10138
      by (simp add: rel_frontier_def frontier_def rel_frontier_convex_hull_explicit [symmetric]
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10139
                    interior_convex_hull_explicit_minimal rel_interior_convex_hull_explicit)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10140
  next
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10141
    case False
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10142
    then have "card s \<le> DIM ('a)"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10143
      by linarith
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10144
    then show ?thesis
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10145
      using assms fs
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10146
      apply (simp add: frontier_def interior_convex_hull_explicit finite_imp_compact)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10147
      apply (simp add: convex_hull_finite)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10148
      done
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10149
  qed
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10150
qed
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10151
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10152
lemma rel_frontier_convex_hull_cases:
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10153
  fixes s :: "'a::euclidean_space set"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10154
  assumes "~ affine_dependent s"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10155
  shows "rel_frontier(convex hull s) = \<Union>{convex hull (s - {x}) |x. x \<in> s}"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10156
proof -
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10157
  have fs: "finite s"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10158
    using assms by (simp add: aff_independent_finite)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10159
  { fix u a
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10160
  have "\<forall>x\<in>s. 0 \<le> u x \<Longrightarrow> a \<in> s \<Longrightarrow> u a = 0 \<Longrightarrow> setsum u s = 1 \<Longrightarrow>
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10161
            \<exists>x v. x \<in> s \<and>
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10162
                  (\<forall>x\<in>s - {x}. 0 \<le> v x) \<and>
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10163
                      setsum v (s - {x}) = 1 \<and> (\<Sum>x\<in>s - {x}. v x *\<^sub>R x) = (\<Sum>x\<in>s. u x *\<^sub>R x)"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10164
    apply (rule_tac x=a in exI)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10165
    apply (rule_tac x=u in exI)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10166
    apply (simp add: Groups_Big.setsum_diff1 fs)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10167
    done }
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
 10168
  moreover
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10169
  { fix a u
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10170
    have "a \<in> s \<Longrightarrow> \<forall>x\<in>s - {a}. 0 \<le> u x \<Longrightarrow> setsum u (s - {a}) = 1 \<Longrightarrow>
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10171
            \<exists>v. (\<forall>x\<in>s. 0 \<le> v x) \<and>
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10172
                 (\<exists>x\<in>s. v x = 0) \<and> setsum v s = 1 \<and> (\<Sum>x\<in>s. v x *\<^sub>R x) = (\<Sum>x\<in>s - {a}. u x *\<^sub>R x)"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10173
    apply (rule_tac x="\<lambda>x. if x = a then 0 else u x" in exI)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10174
    apply (auto simp: setsum.If_cases Diff_eq if_smult fs)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10175
    done }
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10176
  ultimately show ?thesis
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10177
    using assms
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10178
    apply (simp add: rel_frontier_convex_hull_explicit)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10179
    apply (simp add: convex_hull_finite fs Union_SetCompr_eq, auto)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10180
    done
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10181
qed
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10182
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10183
lemma frontier_convex_hull_eq_rel_frontier:
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10184
  fixes s :: "'a::euclidean_space set"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10185
  assumes "~ affine_dependent s"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
 10186
  shows "frontier(convex hull s) =
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10187
           (if card s \<le> DIM ('a) then convex hull s else rel_frontier(convex hull s))"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
 10188
  using assms
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
 10189
  unfolding rel_frontier_def frontier_def
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
 10190
  by (simp add: affine_independent_span_gt rel_interior_interior
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10191
                finite_imp_compact empty_interior_convex_hull aff_independent_finite)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10192
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10193
lemma frontier_convex_hull_cases:
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10194
  fixes s :: "'a::euclidean_space set"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10195
  assumes "~ affine_dependent s"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
 10196
  shows "frontier(convex hull s) =
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10197
           (if card s \<le> DIM ('a) then convex hull s else \<Union>{convex hull (s - {x}) |x. x \<in> s})"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10198
by (simp add: assms frontier_convex_hull_eq_rel_frontier rel_frontier_convex_hull_cases)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10199
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10200
lemma in_frontier_convex_hull:
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10201
  fixes s :: "'a::euclidean_space set"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10202
  assumes "finite s" "card s \<le> Suc (DIM ('a))" "x \<in> s"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10203
  shows   "x \<in> frontier(convex hull s)"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10204
proof (cases "affine_dependent s")
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
 10205
  case True
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10206
  with assms show ?thesis
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10207
    apply (auto simp: affine_dependent_def frontier_def finite_imp_compact hull_inc)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10208
    by (metis card.insert_remove convex_hull_subset_affine_hull empty_interior_affine_hull finite_Diff hull_redundant insert_Diff insert_Diff_single insert_not_empty interior_mono not_less_eq_eq subset_empty)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10209
next
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
 10210
  case False
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
 10211
  { assume "card s = Suc (card Basis)"
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10212
    then have cs: "Suc 0 < card s"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10213
      by (simp add: DIM_positive)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10214
    with subset_singletonD have "\<exists>y \<in> s. y \<noteq> x"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10215
      by (cases "s \<le> {x}") fastforce+
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10216
  } note [dest!] = this
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10217
  show ?thesis using assms
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10218
    unfolding frontier_convex_hull_cases [OF False] Union_SetCompr_eq
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10219
    by (auto simp: le_Suc_eq hull_inc)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10220
qed
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10221
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10222
lemma not_in_interior_convex_hull:
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10223
  fixes s :: "'a::euclidean_space set"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10224
  assumes "finite s" "card s \<le> Suc (DIM ('a))" "x \<in> s"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10225
  shows   "x \<notin> interior(convex hull s)"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10226
using in_frontier_convex_hull [OF assms]
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10227
by (metis Diff_iff frontier_def)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10228
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10229
lemma interior_convex_hull_eq_empty:
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10230
  fixes s :: "'a::euclidean_space set"
60762
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10231
  assumes "card s = Suc (DIM ('a))"
60307
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10232
  shows   "interior(convex hull s) = {} \<longleftrightarrow> affine_dependent s"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10233
proof -
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10234
  { fix a b
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10235
    assume ab: "a \<in> interior (convex hull s)" "b \<in> s" "b \<in> affine hull (s - {b})"
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10236
    then have "interior(affine hull s) = {}" using assms
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10237
      by (metis DIM_positive One_nat_def Suc_mono card.remove card_infinite empty_interior_affine_hull eq_iff hull_redundant insert_Diff not_less zero_le_one)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10238
    then have False using ab
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10239
      by (metis convex_hull_subset_affine_hull equals0D interior_mono subset_eq)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10240
  } then
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10241
  show ?thesis
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10242
    using assms
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10243
    apply auto
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10244
    apply (metis UNIV_I affine_hull_convex_hull affine_hull_empty affine_independent_span_eq convex_convex_hull empty_iff rel_interior_interior rel_interior_same_affine_hull)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10245
    apply (auto simp: affine_dependent_def)
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10246
    done
75e1aa7a450e Convex hulls: theorems about interior, etc. And a few simple lemmas.
paulson <lp15@cam.ac.uk>
parents: 60303
diff changeset
 10247
qed
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 49962
diff changeset
 10248
60800
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
 10249
60762
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10250
subsection \<open>Coplanarity, and collinearity in terms of affine hull\<close>
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10251
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10252
definition coplanar  where
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10253
   "coplanar s \<equiv> \<exists>u v w. s \<subseteq> affine hull {u,v,w}"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10254
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10255
lemma collinear_affine_hull:
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10256
  "collinear s \<longleftrightarrow> (\<exists>u v. s \<subseteq> affine hull {u,v})"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10257
proof (cases "s={}")
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10258
  case True then show ?thesis
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10259
    by simp
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10260
next
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10261
  case False
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10262
  then obtain x where x: "x \<in> s" by auto
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10263
  { fix u
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10264
    assume *: "\<And>x y. \<lbrakk>x\<in>s; y\<in>s\<rbrakk> \<Longrightarrow> \<exists>c. x - y = c *\<^sub>R u"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10265
    have "\<exists>u v. s \<subseteq> {a *\<^sub>R u + b *\<^sub>R v |a b. a + b = 1}"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10266
      apply (rule_tac x=x in exI)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10267
      apply (rule_tac x="x+u" in exI, clarify)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10268
      apply (erule exE [OF * [OF x]])
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10269
      apply (rename_tac c)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10270
      apply (rule_tac x="1+c" in exI)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10271
      apply (rule_tac x="-c" in exI)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10272
      apply (simp add: algebra_simps)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10273
      done
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10274
  } moreover
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10275
  { fix u v x y
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10276
    assume *: "s \<subseteq> {a *\<^sub>R u + b *\<^sub>R v |a b. a + b = 1}"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10277
    have "x\<in>s \<Longrightarrow> y\<in>s \<Longrightarrow> \<exists>c. x - y = c *\<^sub>R (v-u)"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10278
      apply (drule subsetD [OF *])+
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10279
      apply simp
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10280
      apply clarify
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10281
      apply (rename_tac r1 r2)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10282
      apply (rule_tac x="r1-r2" in exI)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10283
      apply (simp add: algebra_simps)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10284
      apply (metis scaleR_left.add)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10285
      done
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10286
  } ultimately
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10287
  show ?thesis
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10288
  unfolding collinear_def affine_hull_2
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10289
    by blast
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10290
qed
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10291
62620
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10292
lemma collinear_closed_segment [simp]: "collinear (closed_segment a b)"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10293
by (metis affine_hull_convex_hull collinear_affine_hull hull_subset segment_convex_hull)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10294
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10295
lemma collinear_open_segment [simp]: "collinear (open_segment a b)"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10296
  unfolding open_segment_def
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10297
  by (metis convex_hull_subset_affine_hull segment_convex_hull dual_order.trans
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10298
    convex_hull_subset_affine_hull Diff_subset collinear_affine_hull)
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10299
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10300
lemma subset_continuous_image_segment_1:
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10301
  fixes f :: "'a::euclidean_space \<Rightarrow> real"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10302
  assumes "continuous_on (closed_segment a b) f"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10303
  shows "closed_segment (f a) (f b) \<subseteq> image f (closed_segment a b)"
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10304
by (metis connected_segment convex_contains_segment ends_in_segment imageI
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10305
           is_interval_connected_1 is_interval_convex connected_continuous_image [OF assms])
d21dab28b3f9 New results about paths, segments, etc. The notion of simply_connected.
paulson <lp15@cam.ac.uk>
parents: 62618
diff changeset
 10306
60762
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10307
lemma collinear_imp_coplanar:
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10308
  "collinear s ==> coplanar s"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10309
by (metis collinear_affine_hull coplanar_def insert_absorb2)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10310
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10311
lemma collinear_small:
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10312
  assumes "finite s" "card s \<le> 2"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10313
    shows "collinear s"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10314
proof -
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10315
  have "card s = 0 \<or> card s = 1 \<or> card s = 2"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10316
    using assms by linarith
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10317
  then show ?thesis using assms
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10318
    using card_eq_SucD
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10319
    by auto (metis collinear_2 numeral_2_eq_2)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10320
qed
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10321
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10322
lemma coplanar_small:
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10323
  assumes "finite s" "card s \<le> 3"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10324
    shows "coplanar s"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10325
proof -
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10326
  have "card s \<le> 2 \<or> card s = Suc (Suc (Suc 0))"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10327
    using assms by linarith
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10328
  then show ?thesis using assms
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10329
    apply safe
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10330
    apply (simp add: collinear_small collinear_imp_coplanar)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10331
    apply (safe dest!: card_eq_SucD)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10332
    apply (auto simp: coplanar_def)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10333
    apply (metis hull_subset insert_subset)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10334
    done
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10335
qed
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10336
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10337
lemma coplanar_empty: "coplanar {}"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10338
  by (simp add: coplanar_small)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10339
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10340
lemma coplanar_sing: "coplanar {a}"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10341
  by (simp add: coplanar_small)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10342
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10343
lemma coplanar_2: "coplanar {a,b}"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10344
  by (auto simp: card_insert_if coplanar_small)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10345
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10346
lemma coplanar_3: "coplanar {a,b,c}"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10347
  by (auto simp: card_insert_if coplanar_small)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10348
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10349
lemma collinear_affine_hull_collinear: "collinear(affine hull s) \<longleftrightarrow> collinear s"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10350
  unfolding collinear_affine_hull
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10351
  by (metis affine_affine_hull subset_hull hull_hull hull_mono)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10352
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10353
lemma coplanar_affine_hull_coplanar: "coplanar(affine hull s) \<longleftrightarrow> coplanar s"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10354
  unfolding coplanar_def
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10355
  by (metis affine_affine_hull subset_hull hull_hull hull_mono)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10356
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10357
lemma coplanar_linear_image:
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10358
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10359
  assumes "coplanar s" "linear f" shows "coplanar(f ` s)"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10360
proof -
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10361
  { fix u v w
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10362
    assume "s \<subseteq> affine hull {u, v, w}"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10363
    then have "f ` s \<subseteq> f ` (affine hull {u, v, w})"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10364
      by (simp add: image_mono)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10365
    then have "f ` s \<subseteq> affine hull (f ` {u, v, w})"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10366
      by (metis assms(2) linear_conv_bounded_linear affine_hull_linear_image)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10367
  } then
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10368
  show ?thesis
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10369
    by auto (meson assms(1) coplanar_def)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10370
qed
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10371
60800
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
 10372
lemma coplanar_translation_imp: "coplanar s \<Longrightarrow> coplanar ((\<lambda>x. a + x) ` s)"
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
 10373
  unfolding coplanar_def
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
 10374
  apply clarify
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
 10375
  apply (rule_tac x="u+a" in exI)
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
 10376
  apply (rule_tac x="v+a" in exI)
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
 10377
  apply (rule_tac x="w+a" in exI)
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
 10378
  using affine_hull_translation [of a "{u,v,w}" for u v w]
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
 10379
  apply (force simp: add.commute)
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
 10380
  done
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
 10381
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
 10382
lemma coplanar_translation_eq: "coplanar((\<lambda>x. a + x) ` s) \<longleftrightarrow> coplanar s"
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
 10383
    by (metis (no_types) coplanar_translation_imp translation_galois)
60762
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10384
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10385
lemma coplanar_linear_image_eq:
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10386
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10387
  assumes "linear f" "inj f" shows "coplanar(f ` s) = coplanar s"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60800
diff changeset
 10388
proof
60800
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
 10389
  assume "coplanar s"
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
 10390
  then show "coplanar (f ` s)"
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
 10391
    unfolding coplanar_def
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
 10392
    using affine_hull_linear_image [of f "{u,v,w}" for u v w]  assms
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
 10393
    by (meson coplanar_def coplanar_linear_image)
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
 10394
next
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
 10395
  obtain g where g: "linear g" "g \<circ> f = id"
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
 10396
    using linear_injective_left_inverse [OF assms]
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
 10397
    by blast
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
 10398
  assume "coplanar (f ` s)"
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
 10399
  then obtain u v w where "f ` s \<subseteq> affine hull {u, v, w}"
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
 10400
    by (auto simp: coplanar_def)
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
 10401
  then have "g ` f ` s \<subseteq> g ` (affine hull {u, v, w})"
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
 10402
    by blast
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
 10403
  then have "s \<subseteq> g ` (affine hull {u, v, w})"
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
 10404
    using g by (simp add: Fun.image_comp)
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
 10405
  then show "coplanar s"
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
 10406
    unfolding coplanar_def
61222
05d28dc76e5c isabelle update_cartouches;
wenzelm
parents: 61104
diff changeset
 10407
    using affine_hull_linear_image [of g "{u,v,w}" for u v w]  \<open>linear g\<close> linear_conv_bounded_linear
60800
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
 10408
    by fastforce
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
 10409
qed
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
 10410
(*The HOL Light proof is simply
7d04351c795a New material for Cauchy's integral theorem
paulson <lp15@cam.ac.uk>
parents: 60762
diff changeset
 10411
    MATCH_ACCEPT_TAC(LINEAR_INVARIANT_RULE COPLANAR_LINEAR_IMAGE));;
60762
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10412
*)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10413
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10414
lemma coplanar_subset: "\<lbrakk>coplanar t; s \<subseteq> t\<rbrakk> \<Longrightarrow> coplanar s"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10415
  by (meson coplanar_def order_trans)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10416
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10417
lemma affine_hull_3_imp_collinear: "c \<in> affine hull {a,b} \<Longrightarrow> collinear {a,b,c}"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10418
  by (metis collinear_2 collinear_affine_hull_collinear hull_redundant insert_commute)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10419
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10420
lemma collinear_3_imp_in_affine_hull: "\<lbrakk>collinear {a,b,c}; a \<noteq> b\<rbrakk> \<Longrightarrow> c \<in> affine hull {a,b}"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10421
  unfolding collinear_def
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10422
  apply clarify
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10423
  apply (frule_tac x=b in bspec, blast, drule_tac x=a in bspec, blast, erule exE)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10424
  apply (drule_tac x=c in bspec, blast, drule_tac x=a in bspec, blast, erule exE)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10425
  apply (rename_tac y x)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10426
  apply (simp add: affine_hull_2)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10427
  apply (rule_tac x="1 - x/y" in exI)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10428
  apply (simp add: algebra_simps)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10429
  done
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10430
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10431
lemma collinear_3_affine_hull:
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10432
  assumes "a \<noteq> b"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10433
    shows "collinear {a,b,c} \<longleftrightarrow> c \<in> affine hull {a,b}"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10434
using affine_hull_3_imp_collinear assms collinear_3_imp_in_affine_hull by blast
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10435
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10436
lemma collinear_3_eq_affine_dependent:
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10437
  "collinear{a,b,c} \<longleftrightarrow> a = b \<or> a = c \<or> b = c \<or> affine_dependent {a,b,c}"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10438
apply (case_tac "a=b", simp)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10439
apply (case_tac "a=c")
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10440
apply (simp add: insert_commute)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10441
apply (case_tac "b=c")
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10442
apply (simp add: insert_commute)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10443
apply (auto simp: affine_dependent_def collinear_3_affine_hull insert_Diff_if)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10444
apply (metis collinear_3_affine_hull insert_commute)+
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10445
done
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10446
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10447
lemma affine_dependent_imp_collinear_3:
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10448
  "affine_dependent {a,b,c} \<Longrightarrow> collinear{a,b,c}"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10449
by (simp add: collinear_3_eq_affine_dependent)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10450
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10451
lemma collinear_3: "NO_MATCH 0 x \<Longrightarrow> collinear {x,y,z} \<longleftrightarrow> collinear {0, x-y, z-y}"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10452
  by (auto simp add: collinear_def)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10453
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10454
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10455
thm affine_hull_nonempty
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10456
corollary affine_hull_eq_empty [simp]: "affine hull S = {} \<longleftrightarrow> S = {}"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10457
  using affine_hull_nonempty by blast
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10458
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10459
lemma affine_hull_2_alt:
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10460
  fixes a b :: "'a::real_vector"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10461
  shows "affine hull {a,b} = range (\<lambda>u. a + u *\<^sub>R (b - a))"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10462
apply (simp add: affine_hull_2, safe)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10463
apply (rule_tac x=v in image_eqI)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10464
apply (simp add: algebra_simps)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10465
apply (metis scaleR_add_left scaleR_one, simp)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10466
apply (rule_tac x="1-u" in exI)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10467
apply (simp add: algebra_simps)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10468
done
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10469
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10470
lemma interior_convex_hull_3_minimal:
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10471
  fixes a :: "'a::euclidean_space"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10472
  shows "\<lbrakk>~ collinear{a,b,c}; DIM('a) = 2\<rbrakk>
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10473
         \<Longrightarrow> interior(convex hull {a,b,c}) =
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10474
                {v. \<exists>x y z. 0 < x \<and> 0 < y \<and> 0 < z \<and> x + y + z = 1 \<and>
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10475
                            x *\<^sub>R a + y *\<^sub>R b + z *\<^sub>R c = v}"
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10476
apply (simp add: collinear_3_eq_affine_dependent interior_convex_hull_explicit_minimal, safe)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10477
apply (rule_tac x="u a" in exI, simp)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10478
apply (rule_tac x="u b" in exI, simp)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10479
apply (rule_tac x="u c" in exI, simp)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10480
apply (rename_tac uu x y z)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10481
apply (rule_tac x="\<lambda>r. (if r=a then x else if r=b then y else if r=c then z else 0)" in exI)
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10482
apply simp
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10483
done
bf0c76ccee8d new material for multivariate analysis, etc.
paulson
parents: 60585
diff changeset
 10484
60974
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10485
subsection\<open>The infimum of the distance between two sets\<close>
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10486
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10487
definition setdist :: "'a::metric_space set \<Rightarrow> 'a set \<Rightarrow> real" where
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10488
  "setdist s t \<equiv>
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10489
       (if s = {} \<or> t = {} then 0
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10490
        else Inf {dist x y| x y. x \<in> s \<and> y \<in> t})"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10491
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10492
lemma setdist_empty1 [simp]: "setdist {} t = 0"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10493
  by (simp add: setdist_def)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10494
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10495
lemma setdist_empty2 [simp]: "setdist t {} = 0"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10496
  by (simp add: setdist_def)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10497
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10498
lemma setdist_pos_le: "0 \<le> setdist s t"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10499
  by (auto simp: setdist_def ex_in_conv [symmetric] intro: cInf_greatest)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10500
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10501
lemma le_setdistI:
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10502
  assumes "s \<noteq> {}" "t \<noteq> {}" "\<And>x y. \<lbrakk>x \<in> s; y \<in> t\<rbrakk> \<Longrightarrow> d \<le> dist x y"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10503
    shows "d \<le> setdist s t"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10504
  using assms
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10505
  by (auto simp: setdist_def Set.ex_in_conv [symmetric] intro: cInf_greatest)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10506
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10507
lemma setdist_le_dist: "\<lbrakk>x \<in> s; y \<in> t\<rbrakk> \<Longrightarrow> setdist s t \<le> dist x y"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10508
  unfolding setdist_def
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10509
  by (auto intro!: bdd_belowI [where m=0] cInf_lower)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10510
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
 10511
lemma le_setdist_iff:
60974
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10512
        "d \<le> setdist s t \<longleftrightarrow>
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10513
        (\<forall>x \<in> s. \<forall>y \<in> t. d \<le> dist x y) \<and> (s = {} \<or> t = {} \<longrightarrow> d \<le> 0)"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10514
  apply (cases "s = {} \<or> t = {}")
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10515
  apply (force simp add: setdist_def)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10516
  apply (intro iffI conjI)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10517
  using setdist_le_dist apply fastforce
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10518
  apply (auto simp: intro: le_setdistI)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10519
  done
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10520
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
 10521
lemma setdist_ltE:
60974
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10522
  assumes "setdist s t < b" "s \<noteq> {}" "t \<noteq> {}"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10523
    obtains x y where "x \<in> s" "y \<in> t" "dist x y < b"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10524
using assms
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10525
by (auto simp: not_le [symmetric] le_setdist_iff)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10526
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10527
lemma setdist_refl: "setdist s s = 0"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10528
  apply (cases "s = {}")
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10529
  apply (force simp add: setdist_def)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10530
  apply (rule antisym [OF _ setdist_pos_le])
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10531
  apply (metis all_not_in_conv dist_self setdist_le_dist)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10532
  done
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10533
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10534
lemma setdist_sym: "setdist s t = setdist t s"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10535
  by (force simp: setdist_def dist_commute intro!: arg_cong [where f=Inf])
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10536
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10537
lemma setdist_triangle: "setdist s t \<le> setdist s {a} + setdist {a} t"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10538
proof (cases "s = {} \<or> t = {}")
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10539
  case True then show ?thesis
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10540
    using setdist_pos_le by fastforce
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10541
next
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10542
  case False
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
 10543
  have "\<And>x. x \<in> s \<Longrightarrow> setdist s t - dist x a \<le> setdist {a} t"
60974
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10544
    apply (rule le_setdistI, blast)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10545
    using False apply (fastforce intro: le_setdistI)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10546
    apply (simp add: algebra_simps)
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62398
diff changeset
 10547
    apply (metis dist_commute dist_triangle3 order_trans [OF setdist_le_dist])
60974
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10548
    done
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10549
  then have "setdist s t - setdist {a} t \<le> setdist s {a}"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10550
    using False by (fastforce intro: le_setdistI)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10551
  then show ?thesis
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10552
    by (simp add: algebra_simps)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10553
qed
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10554
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10555
lemma setdist_singletons [simp]: "setdist {x} {y} = dist x y"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10556
  by (simp add: setdist_def)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10557
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61880
diff changeset
 10558
lemma setdist_Lipschitz: "\<bar>setdist {x} s - setdist {y} s\<bar> \<le> dist x y"
60974
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10559
  apply (subst setdist_singletons [symmetric])
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10560
  by (metis abs_diff_le_iff diff_le_eq setdist_triangle setdist_sym)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10561
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10562
lemma continuous_at_setdist: "continuous (at x) (\<lambda>y. (setdist {y} s))"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10563
  by (force simp: continuous_at_eps_delta dist_real_def intro: le_less_trans [OF setdist_Lipschitz])
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10564
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10565
lemma continuous_on_setdist: "continuous_on t (\<lambda>y. (setdist {y} s))"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10566
  by (metis continuous_at_setdist continuous_at_imp_continuous_on)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10567
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10568
lemma uniformly_continuous_on_setdist: "uniformly_continuous_on t (\<lambda>y. (setdist {y} s))"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10569
  by (force simp: uniformly_continuous_on_def dist_real_def intro: le_less_trans [OF setdist_Lipschitz])
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10570
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10571
lemma setdist_subset_right: "\<lbrakk>t \<noteq> {}; t \<subseteq> u\<rbrakk> \<Longrightarrow> setdist s u \<le> setdist s t"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10572
  apply (cases "s = {} \<or> u = {}", force)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10573
  apply (auto simp: setdist_def intro!: bdd_belowI [where m=0] cInf_superset_mono)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10574
  done
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10575
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10576
lemma setdist_subset_left: "\<lbrakk>s \<noteq> {}; s \<subseteq> t\<rbrakk> \<Longrightarrow> setdist t u \<le> setdist s u"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10577
  by (metis setdist_subset_right setdist_sym)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10578
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10579
lemma setdist_closure_1 [simp]: "setdist (closure s) t = setdist s t"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10580
proof (cases "s = {} \<or> t = {}")
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10581
  case True then show ?thesis by force
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10582
next
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10583
  case False
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10584
  { fix y
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10585
    assume "y \<in> t"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10586
    have "continuous_on (closure s) (\<lambda>a. dist a y)"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10587
      by (auto simp: continuous_intros dist_norm)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10588
    then have *: "\<And>x. x \<in> closure s \<Longrightarrow> setdist s t \<le> dist x y"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10589
      apply (rule continuous_ge_on_closure)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10590
      apply assumption
61222
05d28dc76e5c isabelle update_cartouches;
wenzelm
parents: 61104
diff changeset
 10591
      apply (blast intro: setdist_le_dist \<open>y \<in> t\<close> )
60974
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10592
      done
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10593
  } note * = this
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10594
  show ?thesis
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10595
    apply (rule antisym)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10596
     using False closure_subset apply (blast intro: setdist_subset_left)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10597
    using False *
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10598
    apply (force simp add: closure_eq_empty intro!: le_setdistI)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10599
    done
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10600
qed
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10601
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10602
lemma setdist_closure_2 [simp]: "setdist t (closure s) = setdist t s"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10603
by (metis setdist_closure_1 setdist_sym)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10604
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10605
lemma setdist_compact_closed:
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10606
  fixes s :: "'a::euclidean_space set"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10607
  assumes s: "compact s" and t: "closed t"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10608
      and "s \<noteq> {}" "t \<noteq> {}"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10609
    shows "\<exists>x \<in> s. \<exists>y \<in> t. dist x y = setdist s t"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10610
proof -
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10611
  have "{x - y |x y. x \<in> s \<and> y \<in> t} \<noteq> {}"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10612
    using assms by blast
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10613
  then
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10614
  have "\<exists>x \<in> s. \<exists>y \<in> t. dist x y \<le> setdist s t"
62381
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
 10615
    apply (rule distance_attains_inf [where a=0, OF compact_closed_differences [OF s t]])
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
 10616
    apply (simp add: dist_norm le_setdist_iff)
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
 10617
    apply blast
60974
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10618
    done
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10619
  then show ?thesis
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10620
    by (blast intro!: antisym [OF _ setdist_le_dist] )
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10621
qed
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10622
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10623
lemma setdist_closed_compact:
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10624
  fixes s :: "'a::euclidean_space set"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10625
  assumes s: "closed s" and t: "compact t"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10626
      and "s \<noteq> {}" "t \<noteq> {}"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10627
    shows "\<exists>x \<in> s. \<exists>y \<in> t. dist x y = setdist s t"
61222
05d28dc76e5c isabelle update_cartouches;
wenzelm
parents: 61104
diff changeset
 10628
  using setdist_compact_closed [OF t s \<open>t \<noteq> {}\<close> \<open>s \<noteq> {}\<close>]
60974
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10629
  by (metis dist_commute setdist_sym)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10630
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10631
lemma setdist_eq_0I: "\<lbrakk>x \<in> s; x \<in> t\<rbrakk> \<Longrightarrow> setdist s t = 0"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10632
  by (metis antisym dist_self setdist_le_dist setdist_pos_le)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10633
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10634
lemma setdist_eq_0_compact_closed:
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10635
  fixes s :: "'a::euclidean_space set"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10636
  assumes s: "compact s" and t: "closed t"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10637
    shows "setdist s t = 0 \<longleftrightarrow> s = {} \<or> t = {} \<or> s \<inter> t \<noteq> {}"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10638
  apply (cases "s = {} \<or> t = {}", force)
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
 10639
  using setdist_compact_closed [OF s t]
60974
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10640
  apply (force intro: setdist_eq_0I )
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10641
  done
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10642
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10643
corollary setdist_gt_0_compact_closed:
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10644
  fixes s :: "'a::euclidean_space set"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10645
  assumes s: "compact s" and t: "closed t"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10646
    shows "setdist s t > 0 \<longleftrightarrow> (s \<noteq> {} \<and> t \<noteq> {} \<and> s \<inter> t = {})"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10647
  using setdist_pos_le [of s t] setdist_eq_0_compact_closed [OF assms]
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10648
  by linarith
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10649
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10650
lemma setdist_eq_0_closed_compact:
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10651
  fixes s :: "'a::euclidean_space set"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10652
  assumes s: "closed s" and t: "compact t"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10653
    shows "setdist s t = 0 \<longleftrightarrow> s = {} \<or> t = {} \<or> s \<inter> t \<noteq> {}"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10654
  using setdist_eq_0_compact_closed [OF t s]
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10655
  by (metis Int_commute setdist_sym)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10656
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10657
lemma setdist_eq_0_bounded:
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10658
  fixes s :: "'a::euclidean_space set"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10659
  assumes "bounded s \<or> bounded t"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10660
    shows "setdist s t = 0 \<longleftrightarrow> s = {} \<or> t = {} \<or> closure s \<inter> closure t \<noteq> {}"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10661
  apply (cases "s = {} \<or> t = {}", force)
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
 10662
  using setdist_eq_0_compact_closed [of "closure s" "closure t"]
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
 10663
        setdist_eq_0_closed_compact [of "closure s" "closure t"] assms
60974
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10664
  apply (force simp add:  bounded_closure compact_eq_bounded_closed)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10665
  done
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10666
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
 10667
lemma setdist_unique:
60974
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10668
  "\<lbrakk>a \<in> s; b \<in> t; \<And>x y. x \<in> s \<and> y \<in> t ==> dist a b \<le> dist x y\<rbrakk>
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10669
   \<Longrightarrow> setdist s t = dist a b"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10670
  by (force simp add: setdist_le_dist le_setdist_iff intro: antisym)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10671
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
 10672
lemma setdist_closest_point:
60974
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10673
    "\<lbrakk>closed s; s \<noteq> {}\<rbrakk> \<Longrightarrow> setdist {a} s = dist a (closest_point s a)"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10674
  apply (rule setdist_unique)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10675
  using closest_point_le
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10676
  apply (auto simp: closest_point_in_set)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10677
  done
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10678
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
 10679
lemma setdist_eq_0_sing_1 [simp]:
60974
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10680
  fixes s :: "'a::euclidean_space set"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10681
  shows "setdist {x} s = 0 \<longleftrightarrow> s = {} \<or> x \<in> closure s"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10682
by (auto simp: setdist_eq_0_bounded)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10683
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
 10684
lemma setdist_eq_0_sing_2 [simp]:
60974
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10685
  fixes s :: "'a::euclidean_space set"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10686
  shows "setdist s {x} = 0 \<longleftrightarrow> s = {} \<or> x \<in> closure s"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10687
by (auto simp: setdist_eq_0_bounded)
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10688
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10689
lemma setdist_sing_in_set:
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10690
  fixes s :: "'a::euclidean_space set"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10691
  shows "x \<in> s \<Longrightarrow> setdist {x} s = 0"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10692
using closure_subset by force
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10693
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10694
lemma setdist_le_sing: "x \<in> s ==> setdist s t \<le> setdist {x} t"
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10695
  using setdist_subset_left by auto
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60809
diff changeset
 10696
63016
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10697
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10698
subsection\<open>Basic lemmas about hyperplanes and halfspaces\<close>
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10699
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10700
lemma hyperplane_eq_Ex:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10701
  assumes "a \<noteq> 0" obtains x where "a \<bullet> x = b"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10702
  by (rule_tac x = "(b / (a \<bullet> a)) *\<^sub>R a" in that) (simp add: assms)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10703
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10704
lemma hyperplane_eq_empty:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10705
     "{x. a \<bullet> x = b} = {} \<longleftrightarrow> a = 0 \<and> b \<noteq> 0"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10706
  using hyperplane_eq_Ex apply auto[1]
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10707
  using inner_zero_right by blast
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10708
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10709
lemma hyperplane_eq_UNIV:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10710
   "{x. a \<bullet> x = b} = UNIV \<longleftrightarrow> a = 0 \<and> b = 0"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10711
proof -
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10712
  have "UNIV \<subseteq> {x. a \<bullet> x = b} \<Longrightarrow> a = 0 \<and> b = 0"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10713
    apply (drule_tac c = "((b+1) / (a \<bullet> a)) *\<^sub>R a" in subsetD)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10714
    apply simp_all
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10715
    by (metis add_cancel_right_right divide_1 zero_neq_one)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10716
  then show ?thesis by force
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10717
qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10718
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10719
lemma halfspace_eq_empty_lt:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10720
   "{x. a \<bullet> x < b} = {} \<longleftrightarrow> a = 0 \<and> b \<le> 0"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10721
proof -
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10722
  have "{x. a \<bullet> x < b} \<subseteq> {} \<Longrightarrow> a = 0 \<and> b \<le> 0"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10723
    apply (rule ccontr)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10724
    apply (drule_tac c = "((b-1) / (a \<bullet> a)) *\<^sub>R a" in subsetD)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10725
    apply force+
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10726
    done
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10727
  then show ?thesis by force
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10728
qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10729
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10730
lemma halfspace_eq_empty_gt:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10731
   "{x. a \<bullet> x > b} = {} \<longleftrightarrow> a = 0 \<and> b \<ge> 0"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10732
using halfspace_eq_empty_lt [of "-a" "-b"]
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10733
by simp
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10734
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10735
lemma halfspace_eq_empty_le:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10736
   "{x. a \<bullet> x \<le> b} = {} \<longleftrightarrow> a = 0 \<and> b < 0"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10737
proof -
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10738
  have "{x. a \<bullet> x \<le> b} \<subseteq> {} \<Longrightarrow> a = 0 \<and> b < 0"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10739
    apply (rule ccontr)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10740
    apply (drule_tac c = "((b-1) / (a \<bullet> a)) *\<^sub>R a" in subsetD)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10741
    apply force+
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10742
    done
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10743
  then show ?thesis by force
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10744
qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10745
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10746
lemma halfspace_eq_empty_ge:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10747
   "{x. a \<bullet> x \<ge> b} = {} \<longleftrightarrow> a = 0 \<and> b > 0"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10748
using halfspace_eq_empty_le [of "-a" "-b"]
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10749
by simp
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10750
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10751
subsubsection\<open>Representing affine hull as a finite intersection of hyperplanes\<close>
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10752
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10753
proposition affine_hull_convex_Int_nonempty_interior:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10754
  fixes S :: "'a::real_normed_vector set"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10755
  assumes "convex S" "S \<inter> interior T \<noteq> {}"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10756
    shows "affine hull (S \<inter> T) = affine hull S"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10757
proof
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10758
  show "affine hull (S \<inter> T) \<subseteq> affine hull S"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10759
    by (simp add: hull_mono)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10760
next
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10761
  obtain a where "a \<in> S" "a \<in> T" and at: "a \<in> interior T"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10762
    using assms interior_subset by blast
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10763
  then obtain e where "e > 0" and e: "cball a e \<subseteq> T"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10764
    using mem_interior_cball by blast
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10765
  have *: "x \<in> op + a ` span ((\<lambda>x. x - a) ` (S \<inter> T))" if "x \<in> S" for x
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10766
  proof (cases "x = a")
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10767
    case True with that span_0 eq_add_iff image_def mem_Collect_eq show ?thesis
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10768
      by blast
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10769
  next
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10770
    case False
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
 10771
    define k where "k = min (1/2) (e / norm (x-a))"
63016
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10772
    have k: "0 < k" "k < 1"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10773
      using \<open>e > 0\<close> False by (auto simp: k_def)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10774
    then have xa: "(x-a) = inverse k *\<^sub>R k *\<^sub>R (x-a)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10775
      by simp
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10776
    have "e / norm (x - a) \<ge> k"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10777
      using k_def by linarith
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10778
    then have "a + k *\<^sub>R (x - a) \<in> cball a e"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10779
      using \<open>0 < k\<close> False by (simp add: dist_norm field_simps)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10780
    then have T: "a + k *\<^sub>R (x - a) \<in> T"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10781
      using e by blast
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10782
    have S: "a + k *\<^sub>R (x - a) \<in> S"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10783
      using k \<open>a \<in> S\<close> convexD [OF \<open>convex S\<close> \<open>a \<in> S\<close> \<open>x \<in> S\<close>, of "1-k" k]
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10784
      by (simp add: algebra_simps)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10785
    have "inverse k *\<^sub>R k *\<^sub>R (x-a) \<in> span ((\<lambda>x. x - a) ` (S \<inter> T))"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10786
      apply (rule span_mul)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10787
      apply (rule span_superset)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10788
      apply (rule image_eqI [where x = "a + k *\<^sub>R (x - a)"])
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10789
      apply (auto simp: S T)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10790
      done
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10791
    with xa image_iff show ?thesis  by fastforce
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10792
  qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10793
  show "affine hull S \<subseteq> affine hull (S \<inter> T)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10794
    apply (simp add: subset_hull)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10795
    apply (simp add: \<open>a \<in> S\<close> \<open>a \<in> T\<close> hull_inc affine_hull_span_gen [of a])
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10796
    apply (force simp: *)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10797
    done
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10798
qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10799
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10800
corollary affine_hull_convex_Int_open:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10801
  fixes S :: "'a::real_normed_vector set"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10802
  assumes "convex S" "open T" "S \<inter> T \<noteq> {}"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10803
    shows "affine hull (S \<inter> T) = affine hull S"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10804
using affine_hull_convex_Int_nonempty_interior assms interior_eq by blast
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10805
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10806
corollary affine_hull_affine_Int_nonempty_interior:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10807
  fixes S :: "'a::real_normed_vector set"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10808
  assumes "affine S" "S \<inter> interior T \<noteq> {}"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10809
    shows "affine hull (S \<inter> T) = affine hull S"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10810
by (simp add: affine_hull_convex_Int_nonempty_interior affine_imp_convex assms)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10811
63018
ae2ec7d86ad4 tidying some proofs; getting rid of "nonempty_witness"
paulson <lp15@cam.ac.uk>
parents: 63016
diff changeset
 10812
corollary affine_hull_affine_Int_open:
63016
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10813
  fixes S :: "'a::real_normed_vector set"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10814
  assumes "affine S" "open T" "S \<inter> T \<noteq> {}"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10815
    shows "affine hull (S \<inter> T) = affine hull S"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10816
by (simp add: affine_hull_convex_Int_open affine_imp_convex assms)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10817
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10818
corollary affine_hull_convex_Int_openin:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10819
  fixes S :: "'a::real_normed_vector set"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10820
  assumes "convex S" "openin (subtopology euclidean (affine hull S)) T" "S \<inter> T \<noteq> {}"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10821
    shows "affine hull (S \<inter> T) = affine hull S"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10822
using assms unfolding openin_open
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10823
by (metis affine_hull_convex_Int_open hull_subset inf.orderE inf_assoc)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10824
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10825
corollary affine_hull_open_in:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10826
  fixes S :: "'a::real_normed_vector set"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10827
  assumes "openin (subtopology euclidean (affine hull T)) S" "S \<noteq> {}"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10828
    shows "affine hull S = affine hull T"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10829
using assms unfolding openin_open
63018
ae2ec7d86ad4 tidying some proofs; getting rid of "nonempty_witness"
paulson <lp15@cam.ac.uk>
parents: 63016
diff changeset
 10830
by (metis affine_affine_hull affine_hull_affine_Int_open hull_hull)
63016
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10831
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10832
corollary affine_hull_open:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10833
  fixes S :: "'a::real_normed_vector set"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10834
  assumes "open S" "S \<noteq> {}"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10835
    shows "affine hull S = UNIV"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10836
by (metis affine_hull_convex_Int_nonempty_interior assms convex_UNIV hull_UNIV inf_top.left_neutral interior_open)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10837
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10838
lemma aff_dim_convex_Int_nonempty_interior:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10839
  fixes S :: "'a::euclidean_space set"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10840
  shows "\<lbrakk>convex S; S \<inter> interior T \<noteq> {}\<rbrakk> \<Longrightarrow> aff_dim(S \<inter> T) = aff_dim S"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10841
using aff_dim_affine_hull2 affine_hull_convex_Int_nonempty_interior by blast
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10842
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10843
lemma aff_dim_convex_Int_open:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10844
  fixes S :: "'a::euclidean_space set"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10845
  shows "\<lbrakk>convex S; open T; S \<inter> T \<noteq> {}\<rbrakk> \<Longrightarrow>  aff_dim(S \<inter> T) = aff_dim S"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10846
using aff_dim_convex_Int_nonempty_interior interior_eq by blast
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10847
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10848
lemma affine_hull_halfspace_lt:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10849
  fixes a :: "'a::euclidean_space"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10850
  shows "affine hull {x. a \<bullet> x < r} = (if a = 0 \<and> r \<le> 0 then {} else UNIV)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10851
using halfspace_eq_empty_lt [of a r]
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10852
by (simp add: open_halfspace_lt affine_hull_open)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10853
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10854
lemma affine_hull_halfspace_le:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10855
  fixes a :: "'a::euclidean_space"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10856
  shows "affine hull {x. a \<bullet> x \<le> r} = (if a = 0 \<and> r < 0 then {} else UNIV)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10857
proof (cases "a = 0")
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10858
  case True then show ?thesis by simp
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10859
next
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10860
  case False
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10861
  then have "affine hull closure {x. a \<bullet> x < r} = UNIV"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10862
    using affine_hull_halfspace_lt closure_same_affine_hull by fastforce
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10863
  moreover have "{x. a \<bullet> x < r} \<subseteq> {x. a \<bullet> x \<le> r}"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10864
    by (simp add: Collect_mono)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10865
  ultimately show ?thesis using False antisym_conv hull_mono top_greatest
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10866
    by (metis affine_hull_halfspace_lt)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10867
qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10868
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10869
lemma affine_hull_halfspace_gt:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10870
  fixes a :: "'a::euclidean_space"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10871
  shows "affine hull {x. a \<bullet> x > r} = (if a = 0 \<and> r \<ge> 0 then {} else UNIV)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10872
using halfspace_eq_empty_gt [of r a]
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10873
by (simp add: open_halfspace_gt affine_hull_open)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10874
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10875
lemma affine_hull_halfspace_ge:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10876
  fixes a :: "'a::euclidean_space"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10877
  shows "affine hull {x. a \<bullet> x \<ge> r} = (if a = 0 \<and> r > 0 then {} else UNIV)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10878
using affine_hull_halfspace_le [of "-a" "-r"] by simp
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10879
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10880
lemma aff_dim_halfspace_lt:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10881
  fixes a :: "'a::euclidean_space"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10882
  shows "aff_dim {x. a \<bullet> x < r} =
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10883
        (if a = 0 \<and> r \<le> 0 then -1 else DIM('a))"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10884
by simp (metis aff_dim_open halfspace_eq_empty_lt open_halfspace_lt)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10885
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10886
lemma aff_dim_halfspace_le:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10887
  fixes a :: "'a::euclidean_space"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10888
  shows "aff_dim {x. a \<bullet> x \<le> r} =
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10889
        (if a = 0 \<and> r < 0 then -1 else DIM('a))"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10890
proof -
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10891
  have "int (DIM('a)) = aff_dim (UNIV::'a set)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10892
    by (simp add: aff_dim_univ)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10893
  then have "aff_dim (affine hull {x. a \<bullet> x \<le> r}) = DIM('a)" if "(a = 0 \<longrightarrow> r \<ge> 0)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10894
    using that by (simp add: affine_hull_halfspace_le not_less)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10895
  then show ?thesis
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10896
    by (force simp: aff_dim_affine_hull)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10897
qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10898
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10899
lemma aff_dim_halfspace_gt:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10900
  fixes a :: "'a::euclidean_space"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10901
  shows "aff_dim {x. a \<bullet> x > r} =
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10902
        (if a = 0 \<and> r \<ge> 0 then -1 else DIM('a))"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10903
by simp (metis aff_dim_open halfspace_eq_empty_gt open_halfspace_gt)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10904
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10905
lemma aff_dim_halfspace_ge:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10906
  fixes a :: "'a::euclidean_space"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10907
  shows "aff_dim {x. a \<bullet> x \<ge> r} =
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10908
        (if a = 0 \<and> r > 0 then -1 else DIM('a))"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10909
using aff_dim_halfspace_le [of "-a" "-r"] by simp
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10910
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10911
subsection\<open>Properties of special hyperplanes\<close>
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10912
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10913
lemma subspace_hyperplane: "subspace {x. a \<bullet> x = 0}"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10914
  by (simp add: subspace_def inner_right_distrib)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10915
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10916
lemma special_hyperplane_span:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10917
  fixes S :: "'n::euclidean_space set"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10918
  assumes "k \<in> Basis"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10919
  shows "{x. k \<bullet> x = 0} = span (Basis - {k})"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10920
proof -
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10921
  have *: "x \<in> span (Basis - {k})" if "k \<bullet> x = 0" for x
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10922
  proof -
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10923
    have "x = (\<Sum>b\<in>Basis. (x \<bullet> b) *\<^sub>R b)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10924
      by (simp add: euclidean_representation)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10925
    also have "... = (\<Sum>b \<in> Basis - {k}. (x \<bullet> b) *\<^sub>R b)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10926
      by (auto simp: setsum.remove [of _ k] inner_commute assms that)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10927
    finally have "x = (\<Sum>b\<in>Basis - {k}. (x \<bullet> b) *\<^sub>R b)" .
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10928
    then show ?thesis
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10929
      by (simp add: Linear_Algebra.span_finite) metis
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10930
  qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10931
  show ?thesis
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10932
    apply (rule span_subspace [symmetric])
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10933
    using assms
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10934
    apply (auto simp: inner_not_same_Basis intro: * subspace_hyperplane)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10935
    done
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10936
qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10937
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10938
lemma dim_special_hyperplane:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10939
  fixes k :: "'n::euclidean_space"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10940
  shows "k \<in> Basis \<Longrightarrow> dim {x. k \<bullet> x = 0} = DIM('n) - 1"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10941
apply (simp add: special_hyperplane_span)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10942
apply (rule Linear_Algebra.dim_unique [OF subset_refl])
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10943
apply (auto simp: Diff_subset independent_substdbasis)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10944
apply (metis member_remove remove_def span_clauses(1))
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10945
done
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10946
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10947
proposition dim_hyperplane:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10948
  fixes a :: "'a::euclidean_space"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10949
  assumes "a \<noteq> 0"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10950
    shows "dim {x. a \<bullet> x = 0} = DIM('a) - 1"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10951
proof -
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10952
  have span0: "span {x. a \<bullet> x = 0} = {x. a \<bullet> x = 0}"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10953
    by (rule span_unique) (auto simp: subspace_hyperplane)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10954
  then obtain B where "independent B"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10955
              and Bsub: "B \<subseteq> {x. a \<bullet> x = 0}"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10956
              and subspB: "{x. a \<bullet> x = 0} \<subseteq> span B"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10957
              and card0: "(card B = dim {x. a \<bullet> x = 0})"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10958
              and ortho: "pairwise orthogonal B"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10959
    using orthogonal_basis_exists by metis
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10960
  with assms have "a \<notin> span B"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10961
    by (metis (mono_tags, lifting) span_eq inner_eq_zero_iff mem_Collect_eq span0 span_subspace)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10962
  then have ind: "independent (insert a B)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10963
    by (simp add: \<open>independent B\<close> independent_insert)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10964
  have "finite B"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10965
    using \<open>independent B\<close> independent_bound by blast
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10966
  have "UNIV \<subseteq> span (insert a B)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10967
  proof fix y::'a
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10968
    obtain r z where z: "y = r *\<^sub>R a + z" "a \<bullet> z = 0"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10969
      apply (rule_tac r="(a \<bullet> y) / (a \<bullet> a)" and z = "y - ((a \<bullet> y) / (a \<bullet> a)) *\<^sub>R a" in that)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10970
      using assms
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10971
      by (auto simp: algebra_simps)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10972
    show "y \<in> span (insert a B)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10973
      by (metis (mono_tags, lifting) z Bsub Convex_Euclidean_Space.span_eq
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10974
         add_diff_cancel_left' mem_Collect_eq span0 span_breakdown_eq span_subspace subspB)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10975
  qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10976
  then have dima: "DIM('a) = dim(insert a B)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10977
    by (metis antisym dim_UNIV dim_subset_UNIV subset_le_dim)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10978
  then show ?thesis
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10979
    by (metis (mono_tags, lifting) Bsub Diff_insert_absorb \<open>a \<notin> span B\<close> ind card0 card_Diff_singleton dim_span indep_card_eq_dim_span insertI1 subsetCE subspB)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10980
qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10981
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10982
lemma lowdim_eq_hyperplane:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10983
  fixes S :: "'a::euclidean_space set"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10984
  assumes "dim S = DIM('a) - 1"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10985
  obtains a where "a \<noteq> 0" and "span S = {x. a \<bullet> x = 0}"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10986
proof -
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10987
  have [simp]: "dim S < DIM('a)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10988
    by (simp add: DIM_positive assms)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10989
  then obtain b where b: "b \<noteq> 0" "span S \<subseteq> {a. b \<bullet> a = 0}"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10990
    using lowdim_subset_hyperplane [of S] by fastforce
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10991
  show ?thesis
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10992
    using b that real_vector_class.subspace_span [of S]
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10993
    by (simp add: assms dim_hyperplane subspace_dim_equal subspace_hyperplane)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10994
qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10995
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10996
lemma dim_eq_hyperplane:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10997
  fixes S :: "'n::euclidean_space set"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10998
  shows "dim S = DIM('n) - 1 \<longleftrightarrow> (\<exists>a. a \<noteq> 0 \<and> span S = {x. a \<bullet> x = 0})"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 10999
by (metis One_nat_def dim_hyperplane dim_span lowdim_eq_hyperplane)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11000
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11001
proposition aff_dim_eq_hyperplane:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11002
  fixes S :: "'a::euclidean_space set"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11003
  shows "aff_dim S = DIM('a) - 1 \<longleftrightarrow> (\<exists>a b. a \<noteq> 0 \<and> affine hull S = {x. a \<bullet> x = b})"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11004
proof (cases "S = {}")
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11005
  case True then show ?thesis
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11006
    by (auto simp: dest: hyperplane_eq_Ex)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11007
next
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11008
  case False
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11009
  then obtain c where "c \<in> S" by blast
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11010
  show ?thesis
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11011
  proof (cases "c = 0")
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11012
    case True show ?thesis
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11013
    apply (simp add: aff_dim_eq_dim [of c] affine_hull_span_gen [of c] \<open>c \<in> S\<close> hull_inc dim_eq_hyperplane
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11014
                del: One_nat_def)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11015
    apply (rule ex_cong)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11016
    using span_0
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11017
    apply (force simp: \<open>c = 0\<close>)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11018
    done
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11019
  next
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11020
    case False
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11021
    have xc_im: "x \<in> op + c ` {y. a \<bullet> y = 0}" if "a \<bullet> x = a \<bullet> c" for a x
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11022
    proof -
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11023
      have "\<exists>y. a \<bullet> y = 0 \<and> c + y = x"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11024
        by (metis that add.commute diff_add_cancel inner_commute inner_diff_left right_minus_eq)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11025
      then show "x \<in> op + c ` {y. a \<bullet> y = 0}"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11026
        by blast
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11027
    qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11028
    have 2: "span ((\<lambda>x. x - c) ` S) = {x. a \<bullet> x = 0}"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11029
         if "op + c ` span ((\<lambda>x. x - c) ` S) = {x. a \<bullet> x = b}" for a b
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11030
    proof -
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11031
      have "b = a \<bullet> c"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11032
        using span_0 that by fastforce
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11033
      with that have "op + c ` span ((\<lambda>x. x - c) ` S) = {x. a \<bullet> x = a \<bullet> c}"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11034
        by simp
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11035
      then have "span ((\<lambda>x. x - c) ` S) = (\<lambda>x. x - c) ` {x. a \<bullet> x = a \<bullet> c}"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11036
        by (metis (no_types) image_cong translation_galois uminus_add_conv_diff)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11037
      also have "... = {x. a \<bullet> x = 0}"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11038
        by (force simp: inner_distrib inner_diff_right
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11039
             intro: image_eqI [where x="x+c" for x])
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11040
      finally show ?thesis .
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11041
    qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11042
    show ?thesis
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11043
      apply (simp add: aff_dim_eq_dim [of c] affine_hull_span_gen [of c] \<open>c \<in> S\<close> hull_inc dim_eq_hyperplane
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11044
                  del: One_nat_def, safe)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11045
      apply (fastforce simp add: inner_distrib intro: xc_im)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11046
      apply (force simp: intro!: 2)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11047
      done
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11048
  qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11049
qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11050
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11051
corollary aff_dim_hyperplane:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11052
  fixes a :: "'a::euclidean_space"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11053
  shows "a \<noteq> 0 \<Longrightarrow> aff_dim {x. a \<bullet> x = r} = DIM('a) - 1"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11054
by (metis aff_dim_eq_hyperplane affine_hull_eq affine_hyperplane)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11055
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11056
subsection\<open>Some stepping theorems\<close>
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11057
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11058
lemma dim_empty [simp]: "dim ({} :: 'a::euclidean_space set) = 0"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11059
  by (force intro!: dim_unique)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11060
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11061
lemma dim_insert:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11062
  fixes x :: "'a::euclidean_space"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11063
  shows "dim (insert x S) = (if x \<in> span S then dim S else dim S + 1)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11064
proof -
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11065
  show ?thesis
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11066
  proof (cases "x \<in> span S")
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11067
    case True then show ?thesis
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11068
      by (metis dim_span span_redundant)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11069
  next
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11070
    case False
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11071
    obtain B where B: "B \<subseteq> span S" "independent B" "span S \<subseteq> span B" "card B = dim (span S)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11072
      using basis_exists [of "span S"] by blast
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11073
    have 1: "insert x B \<subseteq> span (insert x S)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11074
      by (meson \<open>B \<subseteq> span S\<close> dual_order.trans insertI1 insert_subsetI span_mono span_superset subset_insertI)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11075
    have 2: "span (insert x S) \<subseteq> span (insert x B)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11076
      by (metis \<open>B \<subseteq> span S\<close> \<open>span S \<subseteq> span B\<close> span_breakdown_eq span_subspace subsetI subspace_span)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11077
    have 3: "independent (insert x B)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11078
      by (metis B independent_insert span_subspace subspace_span False)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11079
    have "dim (span (insert x S)) = Suc (dim S)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11080
      apply (rule dim_unique [OF 1 2 3])
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11081
      by (metis B False card_insert_disjoint dim_span independent_imp_finite subsetCE)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11082
    then show ?thesis
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11083
      by (simp add: False)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11084
  qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11085
qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11086
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11087
lemma dim_singleton [simp]:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11088
  fixes x :: "'a::euclidean_space"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11089
  shows "dim{x} = (if x = 0 then 0 else 1)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11090
by (simp add: dim_insert)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11091
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11092
lemma dim_eq_0 [simp]:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11093
  fixes S :: "'a::euclidean_space set"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11094
  shows "dim S = 0 \<longleftrightarrow> S \<subseteq> {0}"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11095
apply safe
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11096
apply (metis DIM_positive DIM_real card_ge_dim_independent contra_subsetD dim_empty dim_insert dim_singleton empty_subsetI independent_empty less_not_refl zero_le)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11097
by (metis dim_singleton dim_subset le_0_eq)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11098
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11099
lemma aff_dim_insert:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11100
  fixes a :: "'a::euclidean_space"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11101
  shows "aff_dim (insert a S) = (if a \<in> affine hull S then aff_dim S else aff_dim S + 1)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11102
proof (cases "S = {}")
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11103
  case True then show ?thesis
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11104
    by simp
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11105
next
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11106
  case False
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11107
  then obtain x s' where S: "S = insert x s'" "x \<notin> s'"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11108
    by (meson Set.set_insert all_not_in_conv)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11109
  show ?thesis using S
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11110
    apply (simp add: hull_redundant cong: aff_dim_affine_hull2)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11111
    apply (simp add: affine_hull_insert_span_gen hull_inc)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11112
    apply (simp add: insert_commute [of a] hull_inc aff_dim_eq_dim [of x] dim_insert span_0)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11113
    apply (metis (no_types, lifting) add_minus_cancel image_iff uminus_add_conv_diff)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11114
    done
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11115
qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11116
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11117
lemma subspace_bounded_eq_trivial:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11118
  fixes S :: "'a::real_normed_vector set"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11119
  assumes "subspace S"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11120
    shows "bounded S \<longleftrightarrow> S = {0}"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11121
proof -
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11122
  have "False" if "bounded S" "x \<in> S" "x \<noteq> 0" for x
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11123
  proof -
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11124
    obtain B where B: "\<And>y. y \<in> S \<Longrightarrow> norm y < B" "B > 0"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11125
      using \<open>bounded S\<close> by (force simp: bounded_pos_less)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11126
    have "(B / norm x) *\<^sub>R x \<in> S"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11127
      using assms subspace_mul \<open>x \<in> S\<close> by auto
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11128
    moreover have "norm ((B / norm x) *\<^sub>R x) = B"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11129
      using that B by (simp add: algebra_simps)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11130
    ultimately show False using B by force
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11131
  qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11132
  then have "bounded S \<Longrightarrow> S = {0}"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11133
    using assms subspace_0 by fastforce
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11134
  then show ?thesis
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11135
    by blast
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11136
qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11137
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11138
lemma affine_bounded_eq_trivial:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11139
  fixes S :: "'a::real_normed_vector set"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11140
  assumes "affine S"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11141
    shows "bounded S \<longleftrightarrow> S = {} \<or> (\<exists>a. S = {a})"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11142
proof (cases "S = {}")
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11143
  case True then show ?thesis
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11144
    by simp
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11145
next
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11146
  case False
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11147
  then obtain b where "b \<in> S" by blast
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11148
  with False assms show ?thesis
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11149
    apply safe
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11150
    using affine_diffs_subspace [OF assms \<open>b \<in> S\<close>]
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11151
    apply (metis (no_types, lifting) subspace_bounded_eq_trivial ab_left_minus bounded_translation
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11152
                image_empty image_insert translation_invert)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11153
    apply force
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11154
    done
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11155
qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11156
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11157
lemma affine_bounded_eq_lowdim:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11158
  fixes S :: "'a::euclidean_space set"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11159
  assumes "affine S"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11160
    shows "bounded S \<longleftrightarrow> aff_dim S \<le> 0"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11161
apply safe
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11162
using affine_bounded_eq_trivial assms apply fastforce
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11163
by (metis aff_dim_sing aff_dim_subset affine_dim_equal affine_sing all_not_in_conv assms bounded_empty bounded_insert dual_order.antisym empty_subsetI insert_subset)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11164
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11165
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11166
lemma bounded_hyperplane_eq_trivial_0:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11167
  fixes a :: "'a::euclidean_space"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11168
  assumes "a \<noteq> 0"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11169
  shows "bounded {x. a \<bullet> x = 0} \<longleftrightarrow> DIM('a) = 1"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11170
proof
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11171
  assume "bounded {x. a \<bullet> x = 0}"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11172
  then have "aff_dim {x. a \<bullet> x = 0} \<le> 0"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11173
    by (simp add: affine_bounded_eq_lowdim affine_hyperplane)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11174
  with assms show "DIM('a) = 1"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11175
    by (simp add: le_Suc_eq aff_dim_hyperplane)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11176
next
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11177
  assume "DIM('a) = 1"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11178
  then show "bounded {x. a \<bullet> x = 0}"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11179
    by (simp add: aff_dim_hyperplane affine_bounded_eq_lowdim affine_hyperplane assms)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11180
qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11181
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11182
lemma bounded_hyperplane_eq_trivial:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11183
  fixes a :: "'a::euclidean_space"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11184
  shows "bounded {x. a \<bullet> x = r} \<longleftrightarrow> (if a = 0 then r \<noteq> 0 else DIM('a) = 1)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11185
proof (simp add: bounded_hyperplane_eq_trivial_0, clarify)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11186
  assume "r \<noteq> 0" "a \<noteq> 0"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11187
  have "aff_dim {x. y \<bullet> x = 0} = aff_dim {x. a \<bullet> x = r}" if "y \<noteq> 0" for y::'a
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11188
    by (metis that \<open>a \<noteq> 0\<close> aff_dim_hyperplane)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11189
  then show "bounded {x. a \<bullet> x = r} = (DIM('a) = Suc 0)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11190
    by (metis One_nat_def \<open>a \<noteq> 0\<close> affine_bounded_eq_lowdim affine_hyperplane bounded_hyperplane_eq_trivial_0)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11191
qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11192
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11193
subsection\<open>General case without assuming closure and getting non-strict separation\<close>
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11194
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11195
proposition separating_hyperplane_closed_point_inset:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11196
  fixes S :: "'a::euclidean_space set"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11197
  assumes "convex S" "closed S" "S \<noteq> {}" "z \<notin> S"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11198
  obtains a b where "a \<in> S" "(a - z) \<bullet> z < b" "\<And>x. x \<in> S \<Longrightarrow> b < (a - z) \<bullet> x"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11199
proof -
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11200
  obtain y where "y \<in> S" and y: "\<And>u. u \<in> S \<Longrightarrow> dist z y \<le> dist z u"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11201
    using distance_attains_inf [of S z] assms by auto
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11202
  then have *: "(y - z) \<bullet> z < (y - z) \<bullet> z + (norm (y - z))\<^sup>2 / 2"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11203
    using \<open>y \<in> S\<close> \<open>z \<notin> S\<close> by auto
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11204
  show ?thesis
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11205
  proof (rule that [OF \<open>y \<in> S\<close> *])
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11206
    fix x
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11207
    assume "x \<in> S"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11208
    have yz: "0 < (y - z) \<bullet> (y - z)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11209
      using \<open>y \<in> S\<close> \<open>z \<notin> S\<close> by auto
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11210
    { assume 0: "0 < ((z - y) \<bullet> (x - y))"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11211
      with any_closest_point_dot [OF \<open>convex S\<close> \<open>closed S\<close>]
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11212
      have False
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11213
        using y \<open>x \<in> S\<close> \<open>y \<in> S\<close> not_less by blast
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11214
    }
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11215
    then have "0 \<le> ((y - z) \<bullet> (x - y))"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11216
      by (force simp: not_less inner_diff_left)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11217
    with yz have "0 < 2 * ((y - z) \<bullet> (x - y)) + (y - z) \<bullet> (y - z)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11218
      by (simp add: algebra_simps)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11219
    then show "(y - z) \<bullet> z + (norm (y - z))\<^sup>2 / 2 < (y - z) \<bullet> x"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11220
      by (simp add: field_simps inner_diff_left inner_diff_right dot_square_norm [symmetric])
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11221
  qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11222
qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11223
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11224
lemma separating_hyperplane_closed_0_inset:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11225
  fixes S :: "'a::euclidean_space set"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11226
  assumes "convex S" "closed S" "S \<noteq> {}" "0 \<notin> S"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11227
  obtains a b where "a \<in> S" "a \<noteq> 0" "0 < b" "\<And>x. x \<in> S \<Longrightarrow> a \<bullet> x > b"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11228
using separating_hyperplane_closed_point_inset [OF assms]
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11229
by simp (metis \<open>0 \<notin> S\<close>)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11230
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11231
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11232
proposition separating_hyperplane_set_0_inspan:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11233
  fixes S :: "'a::euclidean_space set"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11234
  assumes "convex S" "S \<noteq> {}" "0 \<notin> S"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11235
  obtains a where "a \<in> span S" "a \<noteq> 0" "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> a \<bullet> x"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11236
proof -
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
 11237
  define k where [abs_def]: "k c = {x. 0 \<le> c \<bullet> x}" for c :: 'a
63016
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11238
  have *: "span S \<inter> frontier (cball 0 1) \<inter> \<Inter>f' \<noteq> {}"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11239
          if f': "finite f'" "f' \<subseteq> k ` S" for f'
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11240
  proof -
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11241
    obtain C where "C \<subseteq> S" "finite C" and C: "f' = k ` C"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11242
      using finite_subset_image [OF f'] by blast
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11243
    obtain a where "a \<in> S" "a \<noteq> 0"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11244
      using \<open>S \<noteq> {}\<close> \<open>0 \<notin> S\<close> ex_in_conv by blast
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11245
    then have "norm (a /\<^sub>R (norm a)) = 1"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11246
      by simp
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11247
    moreover have "a /\<^sub>R (norm a) \<in> span S"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11248
      by (simp add: \<open>a \<in> S\<close> span_mul span_superset)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11249
    ultimately have ass: "a /\<^sub>R (norm a) \<in> span S \<inter> sphere 0 1"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11250
      by simp
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11251
    show ?thesis
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11252
    proof (cases "C = {}")
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11253
      case True with C ass show ?thesis
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11254
        by auto
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11255
    next
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11256
      case False
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11257
      have "closed (convex hull C)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11258
        using \<open>finite C\<close> compact_eq_bounded_closed finite_imp_compact_convex_hull by auto
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11259
      moreover have "convex hull C \<noteq> {}"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11260
        by (simp add: False)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11261
      moreover have "0 \<notin> convex hull C"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11262
        by (metis \<open>C \<subseteq> S\<close> \<open>convex S\<close> \<open>0 \<notin> S\<close> convex_hull_subset hull_same insert_absorb insert_subset)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11263
      ultimately obtain a b
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11264
            where "a \<in> convex hull C" "a \<noteq> 0" "0 < b"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11265
                  and ab: "\<And>x. x \<in> convex hull C \<Longrightarrow> a \<bullet> x > b"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11266
        using separating_hyperplane_closed_0_inset by blast
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11267
      then have "a \<in> S"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11268
        by (metis \<open>C \<subseteq> S\<close> assms(1) subsetCE subset_hull)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11269
      moreover have "norm (a /\<^sub>R (norm a)) = 1"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11270
        using \<open>a \<noteq> 0\<close> by simp
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11271
      moreover have "a /\<^sub>R (norm a) \<in> span S"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11272
        by (simp add: \<open>a \<in> S\<close> span_mul span_superset)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11273
      ultimately have ass: "a /\<^sub>R (norm a) \<in> span S \<inter> sphere 0 1"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11274
        by simp
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11275
      have aa: "a /\<^sub>R (norm a) \<in> (\<Inter>c\<in>C. {x. 0 \<le> c \<bullet> x})"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11276
        apply (clarsimp simp add: divide_simps)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11277
        using ab \<open>0 < b\<close>
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11278
        by (metis hull_inc inner_commute less_eq_real_def less_trans)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11279
      show ?thesis
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11280
        apply (simp add: C k_def)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11281
        using ass aa Int_iff empty_iff by blast
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11282
    qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11283
  qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11284
  have "(span S \<inter> frontier(cball 0 1)) \<inter> (\<Inter> (k ` S)) \<noteq> {}"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11285
    apply (rule compact_imp_fip)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11286
    apply (blast intro: compact_cball)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11287
    using closed_halfspace_ge k_def apply blast
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11288
    apply (metis *)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11289
    done
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11290
  then show ?thesis
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11291
    unfolding set_eq_iff k_def
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11292
    by simp (metis inner_commute norm_eq_zero that zero_neq_one)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11293
qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11294
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11295
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11296
lemma separating_hyperplane_set_point_inaff:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11297
  fixes S :: "'a::euclidean_space set"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11298
  assumes "convex S" "S \<noteq> {}" and zno: "z \<notin> S"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11299
  obtains a b where "(z + a) \<in> affine hull (insert z S)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11300
                and "a \<noteq> 0" and "a \<bullet> z \<le> b"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11301
                and "\<And>x. x \<in> S \<Longrightarrow> a \<bullet> x \<ge> b"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11302
proof -
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11303
from separating_hyperplane_set_0_inspan [of "image (\<lambda>x. -z + x) S"]
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11304
  have "convex (op + (- z) ` S)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11305
    by (simp add: \<open>convex S\<close>)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11306
  moreover have "op + (- z) ` S \<noteq> {}"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11307
    by (simp add: \<open>S \<noteq> {}\<close>)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11308
  moreover have "0 \<notin> op + (- z) ` S"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11309
    using zno by auto
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11310
  ultimately obtain a where "a \<in> span (op + (- z) ` S)" "a \<noteq> 0"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11311
                  and a:  "\<And>x. x \<in> (op + (- z) ` S) \<Longrightarrow> 0 \<le> a \<bullet> x"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11312
    using separating_hyperplane_set_0_inspan [of "image (\<lambda>x. -z + x) S"]
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11313
    by blast
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11314
  then have szx: "\<And>x. x \<in> S \<Longrightarrow> a \<bullet> z \<le> a \<bullet> x"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11315
    by (metis (no_types, lifting) imageI inner_minus_right inner_right_distrib minus_add neg_le_0_iff_le neg_le_iff_le real_add_le_0_iff)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11316
  show ?thesis
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11317
    apply (rule_tac a=a and b = "a  \<bullet> z" in that, simp_all)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11318
    using \<open>a \<in> span (op + (- z) ` S)\<close> affine_hull_insert_span_gen apply blast
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11319
    apply (simp_all add: \<open>a \<noteq> 0\<close> szx)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11320
    done
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11321
qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11322
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11323
proposition supporting_hyperplane_relative_boundary:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11324
  fixes S :: "'a::euclidean_space set"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11325
  assumes "convex S" "x \<in> S" and xno: "x \<notin> rel_interior S"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11326
  obtains a where "a \<noteq> 0"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11327
              and "\<And>y. y \<in> S \<Longrightarrow> a \<bullet> x \<le> a \<bullet> y"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11328
              and "\<And>y. y \<in> rel_interior S \<Longrightarrow> a \<bullet> x < a \<bullet> y"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11329
proof -
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11330
  obtain a b where aff: "(x + a) \<in> affine hull (insert x (rel_interior S))"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11331
                  and "a \<noteq> 0" and "a \<bullet> x \<le> b"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11332
                  and ageb: "\<And>u. u \<in> (rel_interior S) \<Longrightarrow> a \<bullet> u \<ge> b"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11333
    using separating_hyperplane_set_point_inaff [of "rel_interior S" x] assms
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11334
    by (auto simp: rel_interior_eq_empty convex_rel_interior)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11335
  have le_ay: "a \<bullet> x \<le> a \<bullet> y" if "y \<in> S" for y
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11336
  proof -
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11337
    have con: "continuous_on (closure (rel_interior S)) (op \<bullet> a)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11338
      by (rule continuous_intros continuous_on_subset | blast)+
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11339
    have y: "y \<in> closure (rel_interior S)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11340
      using \<open>convex S\<close> closure_def convex_closure_rel_interior \<open>y \<in> S\<close>
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11341
      by fastforce
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11342
    show ?thesis
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11343
      using continuous_ge_on_closure [OF con y] ageb \<open>a \<bullet> x \<le> b\<close>
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11344
      by fastforce
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11345
  qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11346
  have 3: "a \<bullet> x < a \<bullet> y" if "y \<in> rel_interior S" for y
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11347
  proof -
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11348
    obtain e where "0 < e" "y \<in> S" and e: "cball y e \<inter> affine hull S \<subseteq> S"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11349
      using \<open>y \<in> rel_interior S\<close> by (force simp: rel_interior_cball)
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
 11350
    define y' where "y' = y - (e / norm a) *\<^sub>R ((x + a) - x)"
63016
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11351
    have "y' \<in> cball y e"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11352
      unfolding y'_def using \<open>0 < e\<close> by force
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11353
    moreover have "y' \<in> affine hull S"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11354
      unfolding y'_def
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11355
      by (metis \<open>x \<in> S\<close> \<open>y \<in> S\<close> \<open>convex S\<close> aff affine_affine_hull hull_redundant
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11356
                rel_interior_same_affine_hull hull_inc mem_affine_3_minus2)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11357
    ultimately have "y' \<in> S"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11358
      using e by auto
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11359
    have "a \<bullet> x \<le> a \<bullet> y"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11360
      using le_ay \<open>a \<noteq> 0\<close> \<open>y \<in> S\<close> by blast
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11361
    moreover have "a \<bullet> x \<noteq> a \<bullet> y"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11362
      using le_ay [OF \<open>y' \<in> S\<close>] \<open>a \<noteq> 0\<close>
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11363
      apply (simp add: y'_def inner_diff dot_square_norm power2_eq_square)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11364
      by (metis \<open>0 < e\<close> add_le_same_cancel1 inner_commute inner_real_def inner_zero_left le_diff_eq norm_le_zero_iff real_mult_le_cancel_iff2)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11365
    ultimately show ?thesis by force
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11366
  qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11367
  show ?thesis
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11368
    by (rule that [OF \<open>a \<noteq> 0\<close> le_ay 3])
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11369
qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11370
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11371
lemma supporting_hyperplane_relative_frontier:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11372
  fixes S :: "'a::euclidean_space set"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11373
  assumes "convex S" "x \<in> closure S" "x \<notin> rel_interior S"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11374
  obtains a where "a \<noteq> 0"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11375
              and "\<And>y. y \<in> closure S \<Longrightarrow> a \<bullet> x \<le> a \<bullet> y"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11376
              and "\<And>y. y \<in> rel_interior S \<Longrightarrow> a \<bullet> x < a \<bullet> y"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11377
using supporting_hyperplane_relative_boundary [of "closure S" x]
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11378
by (metis assms convex_closure convex_rel_interior_closure)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11379
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11380
subsection\<open> Some results on decomposing convex hulls, e.g. simplicial subdivision\<close>
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11381
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11382
lemma 
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11383
  fixes s :: "'a::euclidean_space set"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11384
  assumes "~ (affine_dependent(s \<union> t))" 
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11385
    shows convex_hull_Int_subset: "convex hull s \<inter> convex hull t \<subseteq> convex hull (s \<inter> t)" (is ?C)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11386
      and affine_hull_Int_subset: "affine hull s \<inter> affine hull t \<subseteq> affine hull (s \<inter> t)" (is ?A)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11387
proof -
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11388
  have [simp]: "finite s" "finite t"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11389
    using aff_independent_finite assms by blast+
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11390
    have "setsum u (s \<inter> t) = 1 \<and> 
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11391
          (\<Sum>v\<in>s \<inter> t. u v *\<^sub>R v) = (\<Sum>v\<in>s. u v *\<^sub>R v)" 
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11392
      if [simp]:  "setsum u s = 1" 
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11393
                 "setsum v t = 1"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11394
         and eq: "(\<Sum>x\<in>t. v x *\<^sub>R x) = (\<Sum>x\<in>s. u x *\<^sub>R x)" for u v
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11395
    proof -
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63018
diff changeset
 11396
    define f where "f x = (if x \<in> s then u x else 0) - (if x \<in> t then v x else 0)" for x
63016
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11397
    have "setsum f (s \<union> t) = 0"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11398
      apply (simp add: f_def setsum_Un setsum_subtractf)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11399
      apply (simp add: setsum.inter_restrict [symmetric] Int_commute)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11400
      done
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11401
    moreover have "(\<Sum>x\<in>(s \<union> t). f x *\<^sub>R x) = 0"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11402
      apply (simp add: f_def setsum_Un scaleR_left_diff_distrib setsum_subtractf)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11403
      apply (simp add: if_smult setsum.inter_restrict [symmetric] Int_commute eq cong del: if_cong)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11404
      done
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11405
    ultimately have "\<And>v. v \<in> s \<union> t \<Longrightarrow> f v = 0"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11406
      using aff_independent_finite assms unfolding affine_dependent_explicit
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11407
      by blast
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11408
    then have u [simp]: "\<And>x. x \<in> s \<Longrightarrow> u x = (if x \<in> t then v x else 0)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11409
      by (simp add: f_def) presburger
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11410
    have "setsum u (s \<inter> t) = setsum u s"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11411
      by (simp add: setsum.inter_restrict) 
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11412
    then have "setsum u (s \<inter> t) = 1"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11413
      using that by linarith
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11414
    moreover have "(\<Sum>v\<in>s \<inter> t. u v *\<^sub>R v) = (\<Sum>v\<in>s. u v *\<^sub>R v)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11415
      by (auto simp: if_smult setsum.inter_restrict intro: setsum.cong)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11416
    ultimately show ?thesis
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11417
      by force
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11418
    qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11419
    then show ?A ?C
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11420
      by (auto simp: convex_hull_finite affine_hull_finite)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11421
qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11422
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11423
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11424
proposition affine_hull_Int:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11425
  fixes s :: "'a::euclidean_space set"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11426
  assumes "~ (affine_dependent(s \<union> t))" 
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11427
    shows "affine hull (s \<inter> t) = affine hull s \<inter> affine hull t"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11428
apply (rule subset_antisym)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11429
apply (simp add: hull_mono)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11430
by (simp add: affine_hull_Int_subset assms)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11431
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11432
proposition convex_hull_Int:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11433
  fixes s :: "'a::euclidean_space set"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11434
  assumes "~ (affine_dependent(s \<union> t))" 
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11435
    shows "convex hull (s \<inter> t) = convex hull s \<inter> convex hull t"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11436
apply (rule subset_antisym)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11437
apply (simp add: hull_mono)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11438
by (simp add: convex_hull_Int_subset assms)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11439
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11440
proposition 
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11441
  fixes s :: "'a::euclidean_space set set"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11442
  assumes "~ (affine_dependent (\<Union>s))" 
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11443
    shows affine_hull_Inter: "affine hull (\<Inter>s) = (\<Inter>t\<in>s. affine hull t)" (is "?A")
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11444
      and convex_hull_Inter: "convex hull (\<Inter>s) = (\<Inter>t\<in>s. convex hull t)" (is "?C")
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11445
proof -
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11446
  have "finite s"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11447
    using aff_independent_finite assms finite_UnionD by blast  
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11448
  then have "?A \<and> ?C" using assms
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11449
  proof (induction s rule: finite_induct)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11450
    case empty then show ?case by auto
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11451
  next
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11452
    case (insert t F)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11453
    then show ?case
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11454
    proof (cases "F={}")
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11455
      case True then show ?thesis by simp
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11456
    next
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11457
      case False
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11458
      with "insert.prems" have [simp]: "\<not> affine_dependent (t \<union> \<Inter>F)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11459
        by (auto intro: affine_dependent_subset)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11460
      have [simp]: "\<not> affine_dependent (\<Union>F)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11461
        using affine_independent_subset insert.prems by fastforce
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11462
      show ?thesis
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11463
        by (simp add: affine_hull_Int convex_hull_Int insert.IH)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11464
    qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11465
  qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11466
  then show "?A" "?C"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11467
    by auto
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11468
qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11469
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11470
lemma affine_hull_finite_intersection_hyperplanes:
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11471
  fixes s :: "'a::euclidean_space set"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11472
  obtains f where
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11473
     "finite f"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11474
     "of_nat (card f) + aff_dim s = DIM('a)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11475
     "affine hull s = \<Inter>f"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11476
     "\<And>h. h \<in> f \<Longrightarrow> \<exists>a b. a \<noteq> 0 \<and> h = {x. a \<bullet> x = b}"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11477
proof -
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11478
  obtain b where "b \<subseteq> s"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11479
             and indb: "\<not> affine_dependent b"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11480
             and eq: "affine hull s = affine hull b"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11481
    using affine_basis_exists by blast
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11482
  obtain c where indc: "\<not> affine_dependent c" and "b \<subseteq> c"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11483
             and affc: "affine hull c = UNIV"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11484
    by (metis extend_to_affine_basis affine_UNIV hull_same indb subset_UNIV)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11485
  then have "finite c"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11486
    by (simp add: aff_independent_finite)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11487
  then have fbc: "finite b" "card b \<le> card c"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11488
    using \<open>b \<subseteq> c\<close> infinite_super by (auto simp: card_mono)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11489
  have imeq: "(\<lambda>x. affine hull x) ` ((\<lambda>a. c - {a}) ` (c - b)) = ((\<lambda>a. affine hull (c - {a})) ` (c - b))"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11490
    by blast
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11491
  have card1: "card ((\<lambda>a. affine hull (c - {a})) ` (c - b)) = card (c - b)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11492
    apply (rule card_image [OF inj_onI])
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11493
    by (metis Diff_eq_empty_iff Diff_iff indc affine_dependent_def hull_subset insert_iff)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11494
  have card2: "(card (c - b)) + aff_dim s = DIM('a)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11495
  proof -
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11496
    have aff: "aff_dim (UNIV::'a set) = aff_dim c"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11497
      by (metis aff_dim_affine_hull affc)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11498
    have "aff_dim b = aff_dim s"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11499
      by (metis (no_types) aff_dim_affine_hull eq)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11500
    then have "int (card b) = 1 + aff_dim s"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11501
      by (simp add: aff_dim_affine_independent indb)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11502
    then show ?thesis
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11503
      using fbc aff
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11504
      by (simp add: \<open>\<not> affine_dependent c\<close> \<open>b \<subseteq> c\<close> aff_dim_affine_independent aff_dim_univ card_Diff_subset of_nat_diff)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11505
  qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11506
  show ?thesis
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11507
  proof (cases "c = b")
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11508
    case True show ?thesis
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11509
      apply (rule_tac f="{}" in that)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11510
      using True affc
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11511
      apply (simp_all add: eq [symmetric])
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11512
      by (metis aff_dim_univ aff_dim_affine_hull)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11513
  next
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11514
    case False
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11515
    have ind: "\<not> affine_dependent (\<Union>a\<in>c - b. c - {a})"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11516
      by (rule affine_independent_subset [OF indc]) auto
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11517
    have affeq: "affine hull s = (\<Inter>x\<in>(\<lambda>a. c - {a}) ` (c - b). affine hull x)"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11518
      using \<open>b \<subseteq> c\<close> False 
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11519
      apply (subst affine_hull_Inter [OF ind, symmetric])
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11520
      apply (simp add: eq double_diff)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11521
      done
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11522
    have *: "1 + aff_dim (c - {t}) = int (DIM('a))" 
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11523
            if t: "t \<in> c" for t
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11524
    proof -
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11525
      have "insert t c = c"
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11526
        using t by blast
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11527
      then show ?thesis
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11528
        by (metis (full_types) add.commute aff_dim_affine_hull aff_dim_insert aff_dim_univ affc affine_dependent_def indc insert_Diff_single t)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11529
    qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11530
    show ?thesis
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11531
      apply (rule_tac f = "(\<lambda>x. affine hull x) ` ((\<lambda>a. c - {a}) ` (c - b))" in that)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11532
         using \<open>finite c\<close> apply blast
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11533
        apply (simp add: imeq card1 card2)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11534
      apply (simp add: affeq, clarify)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11535
      apply (metis DIM_positive One_nat_def Suc_leI add_diff_cancel_left' of_nat_1 aff_dim_eq_hyperplane of_nat_diff *)
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11536
      done
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11537
  qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11538
qed
3590590699b1 numerous theorems about affine hulls, hyperplanes, etc.
paulson <lp15@cam.ac.uk>
parents: 63007
diff changeset
 11539
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
 11540
end