src/HOL/Analysis/Affine.thy
author wenzelm
Sat, 26 Feb 2022 21:40:53 +0100
changeset 75155 0b6c43a87fa6
parent 74224 e04ec2b9ed97
child 76836 30182f9e1818
permissions -rw-r--r--
support Isabelle fonts via patch of vscode resources;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
71243
nipkow
parents: 71242
diff changeset
     1
section "Affine Sets"
nipkow
parents: 71242
diff changeset
     2
71242
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
     3
theory Affine
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
     4
imports Linear_Algebra
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
     5
begin
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
     6
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
     7
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)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
     8
  by (fact if_distrib)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
     9
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    10
lemma sum_delta_notmem:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    11
  assumes "x \<notin> s"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    12
  shows "sum (\<lambda>y. if (y = x) then P x else Q y) s = sum Q s"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    13
    and "sum (\<lambda>y. if (x = y) then P x else Q y) s = sum Q s"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    14
    and "sum (\<lambda>y. if (y = x) then P y else Q y) s = sum Q s"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    15
    and "sum (\<lambda>y. if (x = y) then P y else Q y) s = sum Q s"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    16
  apply (rule_tac [!] sum.cong)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    17
  using assms
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    18
  apply auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    19
  done
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    20
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    21
lemmas independent_finite = independent_imp_finite
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    22
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    23
lemma span_substd_basis:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    24
  assumes d: "d \<subseteq> Basis"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    25
  shows "span d = {x. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    26
  (is "_ = ?B")
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    27
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    28
  have "d \<subseteq> ?B"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    29
    using d by (auto simp: inner_Basis)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    30
  moreover have s: "subspace ?B"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    31
    using subspace_substandard[of "\<lambda>i. i \<notin> d"] .
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    32
  ultimately have "span d \<subseteq> ?B"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    33
    using span_mono[of d "?B"] span_eq_iff[of "?B"] by blast
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    34
  moreover have *: "card d \<le> dim (span d)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    35
    using independent_card_le_dim[of d "span d"] independent_substdbasis[OF assms]
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    36
      span_superset[of d]
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    37
    by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    38
  moreover from * have "dim ?B \<le> dim (span d)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    39
    using dim_substandard[OF assms] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    40
  ultimately show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    41
    using s subspace_dim_equal[of "span d" "?B"] subspace_span[of d] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    42
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    43
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    44
lemma basis_to_substdbasis_subspace_isomorphism:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    45
  fixes B :: "'a::euclidean_space set"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    46
  assumes "independent B"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    47
  shows "\<exists>f d::'a set. card d = card B \<and> linear f \<and> f ` B = d \<and>
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    48
    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"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    49
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    50
  have B: "card B = dim B"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    51
    using dim_unique[of B B "card B"] assms span_superset[of B] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    52
  have "dim B \<le> card (Basis :: 'a set)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    53
    using dim_subset_UNIV[of B] by simp
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    54
  from ex_card[OF this] obtain d :: "'a set" where d: "d \<subseteq> Basis" and t: "card d = dim B"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    55
    by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    56
  let ?t = "{x::'a::euclidean_space. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    57
  have "\<exists>f. linear f \<and> f ` B = d \<and> f ` span B = ?t \<and> inj_on f (span B)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    58
  proof (intro basis_to_basis_subspace_isomorphism subspace_span subspace_substandard span_superset)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    59
    show "d \<subseteq> {x. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    60
      using d inner_not_same_Basis by blast
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    61
  qed (auto simp: span_substd_basis independent_substdbasis dim_substandard d t B assms)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    62
  with t \<open>card B = dim B\<close> d show ?thesis by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    63
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    64
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    65
subsection \<open>Affine set and affine hull\<close>
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    66
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    67
definition\<^marker>\<open>tag important\<close> affine :: "'a::real_vector set \<Rightarrow> bool"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    68
  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)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    69
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    70
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)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    71
  unfolding affine_def by (metis eq_diff_eq')
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    72
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    73
lemma affine_empty [iff]: "affine {}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    74
  unfolding affine_def by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    75
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    76
lemma affine_sing [iff]: "affine {x}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    77
  unfolding affine_alt by (auto simp: scaleR_left_distrib [symmetric])
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    78
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    79
lemma affine_UNIV [iff]: "affine UNIV"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    80
  unfolding affine_def by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    81
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    82
lemma affine_Inter [intro]: "(\<And>s. s\<in>f \<Longrightarrow> affine s) \<Longrightarrow> affine (\<Inter>f)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    83
  unfolding affine_def by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    84
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    85
lemma affine_Int[intro]: "affine s \<Longrightarrow> affine t \<Longrightarrow> affine (s \<inter> t)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    86
  unfolding affine_def by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    87
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    88
lemma affine_scaling: "affine s \<Longrightarrow> affine (image (\<lambda>x. c *\<^sub>R x) s)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    89
  apply (clarsimp simp add: affine_def)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    90
  apply (rule_tac x="u *\<^sub>R x + v *\<^sub>R y" in image_eqI)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    91
  apply (auto simp: algebra_simps)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    92
  done
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    93
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    94
lemma affine_affine_hull [simp]: "affine(affine hull s)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    95
  unfolding hull_def
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    96
  using affine_Inter[of "{t. affine t \<and> s \<subseteq> t}"] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    97
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    98
lemma affine_hull_eq[simp]: "(affine hull s = s) \<longleftrightarrow> affine s"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
    99
  by (metis affine_affine_hull hull_same)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   100
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   101
lemma affine_hyperplane: "affine {x. a \<bullet> x = b}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   102
  by (simp add: affine_def algebra_simps) (metis distrib_right mult.left_neutral)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   103
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   104
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   105
subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Some explicit formulations\<close>
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   106
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   107
text "Formalized by Lars Schewe."
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   108
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   109
lemma affine:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   110
  fixes V::"'a::real_vector set"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   111
  shows "affine V \<longleftrightarrow>
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   112
         (\<forall>S u. finite S \<and> S \<noteq> {} \<and> S \<subseteq> V \<and> sum u S = 1 \<longrightarrow> (\<Sum>x\<in>S. u x *\<^sub>R x) \<in> V)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   113
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   114
  have "u *\<^sub>R x + v *\<^sub>R y \<in> V" if "x \<in> V" "y \<in> V" "u + v = (1::real)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   115
    and *: "\<And>S u. \<lbrakk>finite S; S \<noteq> {}; S \<subseteq> V; sum u S = 1\<rbrakk> \<Longrightarrow> (\<Sum>x\<in>S. u x *\<^sub>R x) \<in> V" for x y u v
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   116
  proof (cases "x = y")
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   117
    case True
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   118
    then show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   119
      using that by (metis scaleR_add_left scaleR_one)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   120
  next
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   121
    case False
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   122
    then show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   123
      using that *[of "{x,y}" "\<lambda>w. if w = x then u else v"] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   124
  qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   125
  moreover have "(\<Sum>x\<in>S. u x *\<^sub>R x) \<in> V"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   126
                if *: "\<And>x y u v. \<lbrakk>x\<in>V; y\<in>V; u + v = 1\<rbrakk> \<Longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> V"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   127
                  and "finite S" "S \<noteq> {}" "S \<subseteq> V" "sum u S = 1" for S u
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   128
  proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   129
    define n where "n = card S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   130
    consider "card S = 0" | "card S = 1" | "card S = 2" | "card S > 2" by linarith
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   131
    then show "(\<Sum>x\<in>S. u x *\<^sub>R x) \<in> V"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   132
    proof cases
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   133
      assume "card S = 1"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   134
      then obtain a where "S={a}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   135
        by (auto simp: card_Suc_eq)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   136
      then show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   137
        using that by simp
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   138
    next
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   139
      assume "card S = 2"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   140
      then obtain a b where "S = {a, b}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   141
        by (metis Suc_1 card_1_singletonE card_Suc_eq)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   142
      then show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   143
        using *[of a b] that
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   144
        by (auto simp: sum_clauses(2))
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   145
    next
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   146
      assume "card S > 2"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   147
      then show ?thesis using that n_def
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   148
      proof (induct n arbitrary: u S)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   149
        case 0
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   150
        then show ?case by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   151
      next
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   152
        case (Suc n u S)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   153
        have "sum u S = card S" if "\<not> (\<exists>x\<in>S. u x \<noteq> 1)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   154
          using that unfolding card_eq_sum by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   155
        with Suc.prems obtain x where "x \<in> S" and x: "u x \<noteq> 1" by force
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   156
        have c: "card (S - {x}) = card S - 1"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   157
          by (simp add: Suc.prems(3) \<open>x \<in> S\<close>)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   158
        have "sum u (S - {x}) = 1 - u x"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   159
          by (simp add: Suc.prems sum_diff1 \<open>x \<in> S\<close>)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   160
        with x have eq1: "inverse (1 - u x) * sum u (S - {x}) = 1"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   161
          by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   162
        have inV: "(\<Sum>y\<in>S - {x}. (inverse (1 - u x) * u y) *\<^sub>R y) \<in> V"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   163
        proof (cases "card (S - {x}) > 2")
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   164
          case True
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   165
          then have S: "S - {x} \<noteq> {}" "card (S - {x}) = n"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   166
            using Suc.prems c by force+
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   167
          show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   168
          proof (rule Suc.hyps)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   169
            show "(\<Sum>a\<in>S - {x}. inverse (1 - u x) * u a) = 1"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   170
              by (auto simp: eq1 sum_distrib_left[symmetric])
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   171
          qed (use S Suc.prems True in auto)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   172
        next
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   173
          case False
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   174
          then have "card (S - {x}) = Suc (Suc 0)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   175
            using Suc.prems c by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   176
          then obtain a b where ab: "(S - {x}) = {a, b}" "a\<noteq>b"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   177
            unfolding card_Suc_eq by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   178
          then show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   179
            using eq1 \<open>S \<subseteq> V\<close>
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   180
            by (auto simp: sum_distrib_left distrib_left intro!: Suc.prems(2)[of a b])
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   181
        qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   182
        have "u x + (1 - u x) = 1 \<Longrightarrow>
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   183
          u x *\<^sub>R x + (1 - u x) *\<^sub>R ((\<Sum>y\<in>S - {x}. u y *\<^sub>R y) /\<^sub>R (1 - u x)) \<in> V"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   184
          by (rule Suc.prems) (use \<open>x \<in> S\<close> Suc.prems inV in \<open>auto simp: scaleR_right.sum\<close>)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   185
        moreover have "(\<Sum>a\<in>S. u a *\<^sub>R a) = u x *\<^sub>R x + (\<Sum>a\<in>S - {x}. u a *\<^sub>R a)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   186
          by (meson Suc.prems(3) sum.remove \<open>x \<in> S\<close>)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   187
        ultimately show "(\<Sum>x\<in>S. u x *\<^sub>R x) \<in> V"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   188
          by (simp add: x)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   189
      qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   190
    qed (use \<open>S\<noteq>{}\<close> \<open>finite S\<close> in auto)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   191
  qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   192
  ultimately show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   193
    unfolding affine_def by meson
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   194
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   195
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   196
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   197
lemma affine_hull_explicit:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   198
  "affine hull p = {y. \<exists>S u. finite S \<and> S \<noteq> {} \<and> S \<subseteq> p \<and> sum u S = 1 \<and> sum (\<lambda>v. u v *\<^sub>R v) S = y}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   199
  (is "_ = ?rhs")
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   200
proof (rule hull_unique)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   201
  show "p \<subseteq> ?rhs"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   202
  proof (intro subsetI CollectI exI conjI)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   203
    show "\<And>x. sum (\<lambda>z. 1) {x} = 1"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   204
      by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   205
  qed auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   206
  show "?rhs \<subseteq> T" if "p \<subseteq> T" "affine T" for T
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   207
    using that unfolding affine by blast
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   208
  show "affine ?rhs"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   209
    unfolding affine_def
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   210
  proof clarify
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   211
    fix u v :: real and sx ux sy uy
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   212
    assume uv: "u + v = 1"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   213
      and x: "finite sx" "sx \<noteq> {}" "sx \<subseteq> p" "sum ux sx = (1::real)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   214
      and y: "finite sy" "sy \<noteq> {}" "sy \<subseteq> p" "sum uy sy = (1::real)" 
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   215
    have **: "(sx \<union> sy) \<inter> sx = sx" "(sx \<union> sy) \<inter> sy = sy"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   216
      by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   217
    show "\<exists>S w. finite S \<and> S \<noteq> {} \<and> S \<subseteq> p \<and>
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   218
        sum w S = 1 \<and> (\<Sum>v\<in>S. w v *\<^sub>R v) = u *\<^sub>R (\<Sum>v\<in>sx. ux v *\<^sub>R v) + v *\<^sub>R (\<Sum>v\<in>sy. uy v *\<^sub>R v)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   219
    proof (intro exI conjI)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   220
      show "finite (sx \<union> sy)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   221
        using x y by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   222
      show "sum (\<lambda>i. (if i\<in>sx then u * ux i else 0) + (if i\<in>sy then v * uy i else 0)) (sx \<union> sy) = 1"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   223
        using x y uv
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   224
        by (simp add: sum_Un sum.distrib sum.inter_restrict[symmetric] sum_distrib_left [symmetric] **)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   225
      have "(\<Sum>i\<in>sx \<union> sy. ((if i \<in> sx then u * ux i else 0) + (if i \<in> sy then v * uy i else 0)) *\<^sub>R i)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   226
          = (\<Sum>i\<in>sx. (u * ux i) *\<^sub>R i) + (\<Sum>i\<in>sy. (v * uy i) *\<^sub>R i)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   227
        using x y
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   228
        unfolding scaleR_left_distrib scaleR_zero_left if_smult
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   229
        by (simp add: sum_Un sum.distrib sum.inter_restrict[symmetric]  **)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   230
      also have "\<dots> = u *\<^sub>R (\<Sum>v\<in>sx. ux v *\<^sub>R v) + v *\<^sub>R (\<Sum>v\<in>sy. uy v *\<^sub>R v)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   231
        unfolding scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] by blast
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   232
      finally show "(\<Sum>i\<in>sx \<union> sy. ((if i \<in> sx then u * ux i else 0) + (if i \<in> sy then v * uy i else 0)) *\<^sub>R i) 
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   233
                  = u *\<^sub>R (\<Sum>v\<in>sx. ux v *\<^sub>R v) + v *\<^sub>R (\<Sum>v\<in>sy. uy v *\<^sub>R v)" .
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   234
    qed (use x y in auto)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   235
  qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   236
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   237
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   238
lemma affine_hull_finite:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   239
  assumes "finite S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   240
  shows "affine hull S = {y. \<exists>u. sum u S = 1 \<and> sum (\<lambda>v. u v *\<^sub>R v) S = y}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   241
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   242
  have *: "\<exists>h. sum h S = 1 \<and> (\<Sum>v\<in>S. h v *\<^sub>R v) = x" 
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   243
    if "F \<subseteq> S" "finite F" "F \<noteq> {}" and sum: "sum u F = 1" and x: "(\<Sum>v\<in>F. u v *\<^sub>R v) = x" for x F u
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   244
  proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   245
    have "S \<inter> F = F"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   246
      using that by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   247
    show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   248
    proof (intro exI conjI)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   249
      show "(\<Sum>x\<in>S. if x \<in> F then u x else 0) = 1"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   250
        by (metis (mono_tags, lifting) \<open>S \<inter> F = F\<close> assms sum.inter_restrict sum)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   251
      show "(\<Sum>v\<in>S. (if v \<in> F then u v else 0) *\<^sub>R v) = x"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   252
        by (simp add: if_smult cong: if_cong) (metis (no_types) \<open>S \<inter> F = F\<close> assms sum.inter_restrict x)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   253
    qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   254
  qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   255
  show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   256
    unfolding affine_hull_explicit using assms
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   257
    by (fastforce dest: *)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   258
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   259
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   260
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   261
subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Stepping theorems and hence small special cases\<close>
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   262
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   263
lemma affine_hull_empty[simp]: "affine hull {} = {}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   264
  by simp
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   265
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   266
lemma affine_hull_finite_step:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   267
  fixes y :: "'a::real_vector"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   268
  shows "finite S \<Longrightarrow>
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   269
      (\<exists>u. sum u (insert a S) = w \<and> sum (\<lambda>x. u x *\<^sub>R x) (insert a S) = y) \<longleftrightarrow>
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   270
      (\<exists>v u. sum u S = w - v \<and> sum (\<lambda>x. u x *\<^sub>R x) S = y - v *\<^sub>R a)" (is "_ \<Longrightarrow> ?lhs = ?rhs")
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   271
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   272
  assume fin: "finite S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   273
  show "?lhs = ?rhs"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   274
  proof
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   275
    assume ?lhs
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   276
    then obtain u where u: "sum u (insert a S) = w \<and> (\<Sum>x\<in>insert a S. u x *\<^sub>R x) = y"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   277
      by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   278
    show ?rhs
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   279
    proof (cases "a \<in> S")
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   280
      case True
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   281
      then show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   282
        using u by (simp add: insert_absorb) (metis diff_zero real_vector.scale_zero_left)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   283
    next
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   284
      case False
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   285
      show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   286
        by (rule exI [where x="u a"]) (use u fin False in auto)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   287
    qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   288
  next
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   289
    assume ?rhs
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   290
    then obtain v u where vu: "sum u S = w - v"  "(\<Sum>x\<in>S. u x *\<^sub>R x) = y - v *\<^sub>R a"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   291
      by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   292
    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)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   293
      by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   294
    show ?lhs
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   295
    proof (cases "a \<in> S")
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   296
      case True
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   297
      show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   298
        by (rule exI [where x="\<lambda>x. (if x=a then v else 0) + u x"])
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   299
           (simp add: True scaleR_left_distrib sum.distrib sum_clauses fin vu * cong: if_cong)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   300
    next
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   301
      case False
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   302
      then show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   303
        apply (rule_tac x="\<lambda>x. if x=a then v else u x" in exI) 
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   304
        apply (simp add: vu sum_clauses(2)[OF fin] *)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   305
        by (simp add: sum_delta_notmem(3) vu)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   306
    qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   307
  qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   308
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   309
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   310
lemma affine_hull_2:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   311
  fixes a b :: "'a::real_vector"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   312
  shows "affine hull {a,b} = {u *\<^sub>R a + v *\<^sub>R b| u v. (u + v = 1)}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   313
  (is "?lhs = ?rhs")
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   314
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   315
  have *:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   316
    "\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::real)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   317
    "\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::'a)" by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   318
  have "?lhs = {y. \<exists>u. sum u {a, b} = 1 \<and> (\<Sum>v\<in>{a, b}. u v *\<^sub>R v) = y}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   319
    using affine_hull_finite[of "{a,b}"] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   320
  also have "\<dots> = {y. \<exists>v u. u b = 1 - v \<and> u b *\<^sub>R b = y - v *\<^sub>R a}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   321
    by (simp add: affine_hull_finite_step[of "{b}" a])
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   322
  also have "\<dots> = ?rhs" unfolding * by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   323
  finally show ?thesis by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   324
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   325
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   326
lemma affine_hull_3:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   327
  fixes a b c :: "'a::real_vector"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   328
  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}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   329
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   330
  have *:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   331
    "\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::real)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   332
    "\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::'a)" by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   333
  show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   334
    apply (simp add: affine_hull_finite affine_hull_finite_step)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   335
    unfolding *
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   336
    apply safe
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   337
     apply (metis add.assoc)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   338
    apply (rule_tac x=u in exI, force)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   339
    done
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   340
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   341
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   342
lemma mem_affine:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   343
  assumes "affine S" "x \<in> S" "y \<in> S" "u + v = 1"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   344
  shows "u *\<^sub>R x + v *\<^sub>R y \<in> S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   345
  using assms affine_def[of S] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   346
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   347
lemma mem_affine_3:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   348
  assumes "affine S" "x \<in> S" "y \<in> S" "z \<in> S" "u + v + w = 1"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   349
  shows "u *\<^sub>R x + v *\<^sub>R y + w *\<^sub>R z \<in> S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   350
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   351
  have "u *\<^sub>R x + v *\<^sub>R y + w *\<^sub>R z \<in> affine hull {x, y, z}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   352
    using affine_hull_3[of x y z] assms by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   353
  moreover
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   354
  have "affine hull {x, y, z} \<subseteq> affine hull S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   355
    using hull_mono[of "{x, y, z}" "S"] assms by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   356
  moreover
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   357
  have "affine hull S = S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   358
    using assms affine_hull_eq[of S] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   359
  ultimately show ?thesis by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   360
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   361
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   362
lemma mem_affine_3_minus:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   363
  assumes "affine S" "x \<in> S" "y \<in> S" "z \<in> S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   364
  shows "x + v *\<^sub>R (y-z) \<in> S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   365
  using mem_affine_3[of S x y z 1 v "-v"] assms
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   366
  by (simp add: algebra_simps)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   367
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   368
corollary%unimportant mem_affine_3_minus2:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   369
    "\<lbrakk>affine S; x \<in> S; y \<in> S; z \<in> S\<rbrakk> \<Longrightarrow> x - v *\<^sub>R (y-z) \<in> S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   370
  by (metis add_uminus_conv_diff mem_affine_3_minus real_vector.scale_minus_left)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   371
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   372
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   373
subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Some relations between affine hull and subspaces\<close>
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   374
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   375
lemma affine_hull_insert_subset_span:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   376
  "affine hull (insert a S) \<subseteq> {a + v| v . v \<in> span {x - a | x . x \<in> S}}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   377
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   378
  have "\<exists>v T u. x = a + v \<and> (finite T \<and> T \<subseteq> {x - a |x. x \<in> S} \<and> (\<Sum>v\<in>T. u v *\<^sub>R v) = v)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   379
    if "finite F" "F \<noteq> {}" "F \<subseteq> insert a S" "sum u F = 1" "(\<Sum>v\<in>F. u v *\<^sub>R v) = x"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   380
    for x F u
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   381
  proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   382
    have *: "(\<lambda>x. x - a) ` (F - {a}) \<subseteq> {x - a |x. x \<in> S}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   383
      using that by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   384
    show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   385
    proof (intro exI conjI)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   386
      show "finite ((\<lambda>x. x - a) ` (F - {a}))"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   387
        by (simp add: that(1))
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   388
      show "(\<Sum>v\<in>(\<lambda>x. x - a) ` (F - {a}). u(v+a) *\<^sub>R v) = x-a"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   389
        by (simp add: sum.reindex[unfolded inj_on_def] algebra_simps
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   390
            sum_subtractf scaleR_left.sum[symmetric] sum_diff1 that)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   391
    qed (use \<open>F \<subseteq> insert a S\<close> in auto)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   392
  qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   393
  then show ?thesis
71840
8ed78bb0b915 Tuned some proofs in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 71243
diff changeset
   394
    unfolding affine_hull_explicit span_explicit by fast
71242
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   395
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   396
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   397
lemma affine_hull_insert_span:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   398
  assumes "a \<notin> S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   399
  shows "affine hull (insert a S) = {a + v | v . v \<in> span {x - a | x.  x \<in> S}}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   400
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   401
  have *: "\<exists>G u. finite G \<and> G \<noteq> {} \<and> G \<subseteq> insert a S \<and> sum u G = 1 \<and> (\<Sum>v\<in>G. u v *\<^sub>R v) = y"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   402
    if "v \<in> span {x - a |x. x \<in> S}" "y = a + v" for y v
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   403
  proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   404
    from that
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   405
    obtain T u where u: "finite T" "T \<subseteq> {x - a |x. x \<in> S}" "a + (\<Sum>v\<in>T. u v *\<^sub>R v) = y"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   406
      unfolding span_explicit by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   407
    define F where "F = (\<lambda>x. x + a) ` T"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   408
    have F: "finite F" "F \<subseteq> S" "(\<Sum>v\<in>F. u (v - a) *\<^sub>R (v - a)) = y - a"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   409
      unfolding F_def using u by (auto simp: sum.reindex[unfolded inj_on_def])
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   410
    have *: "F \<inter> {a} = {}" "F \<inter> - {a} = F"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   411
      using F assms by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   412
    show "\<exists>G u. finite G \<and> G \<noteq> {} \<and> G \<subseteq> insert a S \<and> sum u G = 1 \<and> (\<Sum>v\<in>G. u v *\<^sub>R v) = y"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   413
      apply (rule_tac x = "insert a F" in exI)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   414
      apply (rule_tac x = "\<lambda>x. if x=a then 1 - sum (\<lambda>x. u (x - a)) F else u (x - a)" in exI)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   415
      using assms F
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   416
      apply (auto simp:  sum_clauses sum.If_cases if_smult sum_subtractf scaleR_left.sum algebra_simps *)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   417
      done
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   418
  qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   419
  show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   420
    by (intro subset_antisym affine_hull_insert_subset_span) (auto simp: affine_hull_explicit dest!: *)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   421
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   422
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   423
lemma affine_hull_span:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   424
  assumes "a \<in> S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   425
  shows "affine hull S = {a + v | v. v \<in> span {x - a | x. x \<in> S - {a}}}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   426
  using affine_hull_insert_span[of a "S - {a}", unfolded insert_Diff[OF assms]] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   427
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   428
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   429
subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Parallel affine sets\<close>
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   430
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   431
definition affine_parallel :: "'a::real_vector set \<Rightarrow> 'a::real_vector set \<Rightarrow> bool"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   432
  where "affine_parallel S T \<longleftrightarrow> (\<exists>a. T = (\<lambda>x. a + x) ` S)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   433
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   434
lemma affine_parallel_expl_aux:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   435
  fixes S T :: "'a::real_vector set"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   436
  assumes "\<And>x. x \<in> S \<longleftrightarrow> a + x \<in> T"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   437
  shows "T = (\<lambda>x. a + x) ` S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   438
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   439
  have "x \<in> ((\<lambda>x. a + x) ` S)" if "x \<in> T" for x
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   440
    using that
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   441
    by (simp add: image_iff) (metis add.commute diff_add_cancel assms)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   442
  moreover have "T \<ge> (\<lambda>x. a + x) ` S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   443
    using assms by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   444
  ultimately show ?thesis by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   445
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   446
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   447
lemma affine_parallel_expl: "affine_parallel S T \<longleftrightarrow> (\<exists>a. \<forall>x. x \<in> S \<longleftrightarrow> a + x \<in> T)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   448
  by (auto simp add: affine_parallel_def)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   449
    (use affine_parallel_expl_aux [of S _ T] in blast)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   450
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   451
lemma affine_parallel_reflex: "affine_parallel S S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   452
  unfolding affine_parallel_def
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   453
  using image_add_0 by blast
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   454
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   455
lemma affine_parallel_commut:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   456
  assumes "affine_parallel A B"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   457
  shows "affine_parallel B A"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   458
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   459
  from assms obtain a where B: "B = (\<lambda>x. a + x) ` A"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   460
    unfolding affine_parallel_def by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   461
  have [simp]: "(\<lambda>x. x - a) = plus (- a)" by (simp add: fun_eq_iff)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   462
  from B show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   463
    using translation_galois [of B a A]
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   464
    unfolding affine_parallel_def by blast
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   465
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   466
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   467
lemma affine_parallel_assoc:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   468
  assumes "affine_parallel A B"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   469
    and "affine_parallel B C"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   470
  shows "affine_parallel A C"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   471
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   472
  from assms obtain ab where "B = (\<lambda>x. ab + x) ` A"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   473
    unfolding affine_parallel_def by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   474
  moreover
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   475
  from assms obtain bc where "C = (\<lambda>x. bc + x) ` B"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   476
    unfolding affine_parallel_def by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   477
  ultimately show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   478
    using translation_assoc[of bc ab A] unfolding affine_parallel_def by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   479
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   480
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   481
lemma affine_translation_aux:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   482
  fixes a :: "'a::real_vector"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   483
  assumes "affine ((\<lambda>x. a + x) ` S)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   484
  shows "affine S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   485
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   486
  {
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   487
    fix x y u v
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   488
    assume xy: "x \<in> S" "y \<in> S" "(u :: real) + v = 1"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   489
    then have "(a + x) \<in> ((\<lambda>x. a + x) ` S)" "(a + y) \<in> ((\<lambda>x. a + x) ` S)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   490
      by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   491
    then have h1: "u *\<^sub>R  (a + x) + v *\<^sub>R (a + y) \<in> (\<lambda>x. a + x) ` S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   492
      using xy assms unfolding affine_def by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   493
    have "u *\<^sub>R (a + x) + v *\<^sub>R (a + y) = (u + v) *\<^sub>R a + (u *\<^sub>R x + v *\<^sub>R y)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   494
      by (simp add: algebra_simps)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   495
    also have "\<dots> = a + (u *\<^sub>R x + v *\<^sub>R y)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   496
      using \<open>u + v = 1\<close> by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   497
    ultimately have "a + (u *\<^sub>R x + v *\<^sub>R y) \<in> (\<lambda>x. a + x) ` S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   498
      using h1 by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   499
    then have "u *\<^sub>R x + v *\<^sub>R y \<in> S" by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   500
  }
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   501
  then show ?thesis unfolding affine_def by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   502
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   503
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   504
lemma affine_translation:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   505
  "affine S \<longleftrightarrow> affine ((+) a ` S)" for a :: "'a::real_vector"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   506
proof
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   507
  show "affine ((+) a ` S)" if "affine S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   508
    using that translation_assoc [of "- a" a S]
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   509
    by (auto intro: affine_translation_aux [of "- a" "((+) a ` S)"])
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   510
  show "affine S" if "affine ((+) a ` S)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   511
    using that by (rule affine_translation_aux)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   512
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   513
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   514
lemma parallel_is_affine:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   515
  fixes S T :: "'a::real_vector set"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   516
  assumes "affine S" "affine_parallel S T"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   517
  shows "affine T"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   518
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   519
  from assms obtain a where "T = (\<lambda>x. a + x) ` S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   520
    unfolding affine_parallel_def by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   521
  then show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   522
    using affine_translation assms by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   523
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   524
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   525
lemma subspace_imp_affine: "subspace s \<Longrightarrow> affine s"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   526
  unfolding subspace_def affine_def by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   527
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   528
lemma affine_hull_subset_span: "(affine hull s) \<subseteq> (span s)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   529
  by (metis hull_minimal span_superset subspace_imp_affine subspace_span)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   530
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   531
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   532
subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Subspace parallel to an affine set\<close>
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   533
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   534
lemma subspace_affine: "subspace S \<longleftrightarrow> affine S \<and> 0 \<in> S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   535
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   536
  have h0: "subspace S \<Longrightarrow> affine S \<and> 0 \<in> S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   537
    using subspace_imp_affine[of S] subspace_0 by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   538
  {
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   539
    assume assm: "affine S \<and> 0 \<in> S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   540
    {
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   541
      fix c :: real
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   542
      fix x
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   543
      assume x: "x \<in> S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   544
      have "c *\<^sub>R x = (1-c) *\<^sub>R 0 + c *\<^sub>R x" by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   545
      moreover
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   546
      have "(1 - c) *\<^sub>R 0 + c *\<^sub>R x \<in> S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   547
        using affine_alt[of S] assm x by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   548
      ultimately have "c *\<^sub>R x \<in> S" by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   549
    }
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   550
    then have h1: "\<forall>c. \<forall>x \<in> S. c *\<^sub>R x \<in> S" by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   551
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   552
    {
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   553
      fix x y
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   554
      assume xy: "x \<in> S" "y \<in> S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   555
      define u where "u = (1 :: real)/2"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   556
      have "(1/2) *\<^sub>R (x+y) = (1/2) *\<^sub>R (x+y)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   557
        by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   558
      moreover
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   559
      have "(1/2) *\<^sub>R (x+y)=(1/2) *\<^sub>R x + (1-(1/2)) *\<^sub>R y"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   560
        by (simp add: algebra_simps)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   561
      moreover
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   562
      have "(1 - u) *\<^sub>R x + u *\<^sub>R y \<in> S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   563
        using affine_alt[of S] assm xy by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   564
      ultimately
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   565
      have "(1/2) *\<^sub>R (x+y) \<in> S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   566
        using u_def by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   567
      moreover
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   568
      have "x + y = 2 *\<^sub>R ((1/2) *\<^sub>R (x+y))"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   569
        by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   570
      ultimately
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   571
      have "x + y \<in> S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   572
        using h1[rule_format, of "(1/2) *\<^sub>R (x+y)" "2"] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   573
    }
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   574
    then have "\<forall>x \<in> S. \<forall>y \<in> S. x + y \<in> S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   575
      by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   576
    then have "subspace S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   577
      using h1 assm unfolding subspace_def by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   578
  }
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   579
  then show ?thesis using h0 by metis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   580
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   581
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   582
lemma affine_diffs_subspace:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   583
  assumes "affine S" "a \<in> S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   584
  shows "subspace ((\<lambda>x. (-a)+x) ` S)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   585
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   586
  have [simp]: "(\<lambda>x. x - a) = plus (- a)" by (simp add: fun_eq_iff)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   587
  have "affine ((\<lambda>x. (-a)+x) ` S)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   588
    using affine_translation assms by blast
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   589
  moreover have "0 \<in> ((\<lambda>x. (-a)+x) ` S)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   590
    using assms exI[of "(\<lambda>x. x\<in>S \<and> -a+x = 0)" a] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   591
  ultimately show ?thesis using subspace_affine by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   592
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   593
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   594
lemma affine_diffs_subspace_subtract:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   595
  "subspace ((\<lambda>x. x - a) ` S)" if "affine S" "a \<in> S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   596
  using that affine_diffs_subspace [of _ a] by simp
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   597
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   598
lemma parallel_subspace_explicit:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   599
  assumes "affine S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   600
    and "a \<in> S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   601
  assumes "L \<equiv> {y. \<exists>x \<in> S. (-a) + x = y}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   602
  shows "subspace L \<and> affine_parallel S L"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   603
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   604
  from assms have "L = plus (- a) ` S" by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   605
  then have par: "affine_parallel S L"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   606
    unfolding affine_parallel_def ..
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   607
  then have "affine L" using assms parallel_is_affine by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   608
  moreover have "0 \<in> L"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   609
    using assms by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   610
  ultimately show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   611
    using subspace_affine par by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   612
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   613
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   614
lemma parallel_subspace_aux:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   615
  assumes "subspace A"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   616
    and "subspace B"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   617
    and "affine_parallel A B"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   618
  shows "A \<supseteq> B"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   619
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   620
  from assms obtain a where a: "\<forall>x. x \<in> A \<longleftrightarrow> a + x \<in> B"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   621
    using affine_parallel_expl[of A B] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   622
  then have "-a \<in> A"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   623
    using assms subspace_0[of B] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   624
  then have "a \<in> A"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   625
    using assms subspace_neg[of A "-a"] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   626
  then show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   627
    using assms a unfolding subspace_def by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   628
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   629
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   630
lemma parallel_subspace:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   631
  assumes "subspace A"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   632
    and "subspace B"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   633
    and "affine_parallel A B"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   634
  shows "A = B"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   635
proof
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   636
  show "A \<supseteq> B"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   637
    using assms parallel_subspace_aux by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   638
  show "A \<subseteq> B"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   639
    using assms parallel_subspace_aux[of B A] affine_parallel_commut by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   640
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   641
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   642
lemma affine_parallel_subspace:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   643
  assumes "affine S" "S \<noteq> {}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   644
  shows "\<exists>!L. subspace L \<and> affine_parallel S L"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   645
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   646
  have ex: "\<exists>L. subspace L \<and> affine_parallel S L"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   647
    using assms parallel_subspace_explicit by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   648
  {
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   649
    fix L1 L2
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   650
    assume ass: "subspace L1 \<and> affine_parallel S L1" "subspace L2 \<and> affine_parallel S L2"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   651
    then have "affine_parallel L1 L2"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   652
      using affine_parallel_commut[of S L1] affine_parallel_assoc[of L1 S L2] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   653
    then have "L1 = L2"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   654
      using ass parallel_subspace by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   655
  }
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   656
  then show ?thesis using ex by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   657
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   658
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   659
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   660
subsection \<open>Affine Dependence\<close>
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   661
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   662
text "Formalized by Lars Schewe."
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   663
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   664
definition\<^marker>\<open>tag important\<close> affine_dependent :: "'a::real_vector set \<Rightarrow> bool"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   665
  where "affine_dependent s \<longleftrightarrow> (\<exists>x\<in>s. x \<in> affine hull (s - {x}))"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   666
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   667
lemma affine_dependent_imp_dependent: "affine_dependent s \<Longrightarrow> dependent s"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   668
  unfolding affine_dependent_def dependent_def
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   669
  using affine_hull_subset_span by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   670
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   671
lemma affine_dependent_subset:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   672
   "\<lbrakk>affine_dependent s; s \<subseteq> t\<rbrakk> \<Longrightarrow> affine_dependent t"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   673
apply (simp add: affine_dependent_def Bex_def)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   674
apply (blast dest: hull_mono [OF Diff_mono [OF _ subset_refl]])
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   675
done
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   676
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   677
lemma affine_independent_subset:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   678
  shows "\<lbrakk>\<not> affine_dependent t; s \<subseteq> t\<rbrakk> \<Longrightarrow> \<not> affine_dependent s"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   679
by (metis affine_dependent_subset)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   680
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   681
lemma affine_independent_Diff:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   682
   "\<not> affine_dependent s \<Longrightarrow> \<not> affine_dependent(s - t)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   683
by (meson Diff_subset affine_dependent_subset)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   684
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   685
proposition affine_dependent_explicit:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   686
  "affine_dependent p \<longleftrightarrow>
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   687
    (\<exists>S u. finite S \<and> S \<subseteq> p \<and> sum u S = 0 \<and> (\<exists>v\<in>S. u v \<noteq> 0) \<and> sum (\<lambda>v. u v *\<^sub>R v) S = 0)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   688
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   689
  have "\<exists>S u. finite S \<and> S \<subseteq> p \<and> sum u S = 0 \<and> (\<exists>v\<in>S. u v \<noteq> 0) \<and> (\<Sum>w\<in>S. u w *\<^sub>R w) = 0"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   690
    if "(\<Sum>w\<in>S. u w *\<^sub>R w) = x" "x \<in> p" "finite S" "S \<noteq> {}" "S \<subseteq> p - {x}" "sum u S = 1" for x S u
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   691
  proof (intro exI conjI)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   692
    have "x \<notin> S" 
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   693
      using that by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   694
    then show "(\<Sum>v \<in> insert x S. if v = x then - 1 else u v) = 0"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   695
      using that by (simp add: sum_delta_notmem)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   696
    show "(\<Sum>w \<in> insert x S. (if w = x then - 1 else u w) *\<^sub>R w) = 0"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   697
      using that \<open>x \<notin> S\<close> by (simp add: if_smult sum_delta_notmem cong: if_cong)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   698
  qed (use that in auto)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   699
  moreover have "\<exists>x\<in>p. \<exists>S u. finite S \<and> S \<noteq> {} \<and> S \<subseteq> p - {x} \<and> sum u S = 1 \<and> (\<Sum>v\<in>S. u v *\<^sub>R v) = x"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   700
    if "(\<Sum>v\<in>S. u v *\<^sub>R v) = 0" "finite S" "S \<subseteq> p" "sum u S = 0" "v \<in> S" "u v \<noteq> 0" for S u v
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   701
  proof (intro bexI exI conjI)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   702
    have "S \<noteq> {v}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   703
      using that by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   704
    then show "S - {v} \<noteq> {}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   705
      using that by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   706
    show "(\<Sum>x \<in> S - {v}. - (1 / u v) * u x) = 1"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   707
      unfolding sum_distrib_left[symmetric] sum_diff1[OF \<open>finite S\<close>] by (simp add: that)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   708
    show "(\<Sum>x\<in>S - {v}. (- (1 / u v) * u x) *\<^sub>R x) = v"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   709
      unfolding sum_distrib_left [symmetric] scaleR_scaleR[symmetric]
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   710
                scaleR_right.sum [symmetric] sum_diff1[OF \<open>finite S\<close>] 
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   711
      using that by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   712
    show "S - {v} \<subseteq> p - {v}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   713
      using that by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   714
  qed (use that in auto)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   715
  ultimately show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   716
    unfolding affine_dependent_def affine_hull_explicit by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   717
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   718
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   719
lemma affine_dependent_explicit_finite:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   720
  fixes S :: "'a::real_vector set"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   721
  assumes "finite S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   722
  shows "affine_dependent S \<longleftrightarrow>
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   723
    (\<exists>u. sum u S = 0 \<and> (\<exists>v\<in>S. u v \<noteq> 0) \<and> sum (\<lambda>v. u v *\<^sub>R v) S = 0)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   724
  (is "?lhs = ?rhs")
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   725
proof
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   726
  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)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   727
    by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   728
  assume ?lhs
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   729
  then obtain t u v where
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   730
    "finite t" "t \<subseteq> S" "sum u t = 0" "v\<in>t" "u v \<noteq> 0"  "(\<Sum>v\<in>t. u v *\<^sub>R v) = 0"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   731
    unfolding affine_dependent_explicit by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   732
  then show ?rhs
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   733
    apply (rule_tac x="\<lambda>x. if x\<in>t then u x else 0" in exI)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   734
    apply (auto simp: * sum.inter_restrict[OF assms, symmetric] Int_absorb1[OF \<open>t\<subseteq>S\<close>])
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   735
    done
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   736
next
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   737
  assume ?rhs
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   738
  then obtain u v where "sum u S = 0"  "v\<in>S" "u v \<noteq> 0" "(\<Sum>v\<in>S. u v *\<^sub>R v) = 0"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   739
    by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   740
  then show ?lhs unfolding affine_dependent_explicit
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   741
    using assms by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   742
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   743
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   744
lemma dependent_imp_affine_dependent:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   745
  assumes "dependent {x - a| x . x \<in> s}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   746
    and "a \<notin> s"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   747
  shows "affine_dependent (insert a s)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   748
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   749
  from assms(1)[unfolded dependent_explicit] obtain S u v
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   750
    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"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   751
    by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   752
  define t where "t = (\<lambda>x. x + a) ` S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   753
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   754
  have inj: "inj_on (\<lambda>x. x + a) S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   755
    unfolding inj_on_def by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   756
  have "0 \<notin> S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   757
    using obt(2) assms(2) unfolding subset_eq by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   758
  have fin: "finite t" and "t \<subseteq> s"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   759
    unfolding t_def using obt(1,2) by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   760
  then have "finite (insert a t)" and "insert a t \<subseteq> insert a s"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   761
    by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   762
  moreover have *: "\<And>P Q. (\<Sum>x\<in>t. (if x = a then P x else Q x)) = (\<Sum>x\<in>t. Q x)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   763
    apply (rule sum.cong)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   764
    using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close>
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   765
    apply auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   766
    done
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   767
  have "(\<Sum>x\<in>insert a t. if x = a then - (\<Sum>x\<in>t. u (x - a)) else u (x - a)) = 0"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   768
    unfolding sum_clauses(2)[OF fin] * using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close> by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   769
  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"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   770
    using obt(3,4) \<open>0\<notin>S\<close>
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   771
    by (rule_tac x="v + a" in bexI) (auto simp: t_def)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   772
  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)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   773
    using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close> by (auto intro!: sum.cong)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   774
  have "(\<Sum>x\<in>t. u (x - a)) *\<^sub>R a = (\<Sum>v\<in>t. u (v - a) *\<^sub>R v)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   775
    unfolding scaleR_left.sum
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   776
    unfolding t_def and sum.reindex[OF inj] and o_def
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   777
    using obt(5)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   778
    by (auto simp: sum.distrib scaleR_right_distrib)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   779
  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"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   780
    unfolding sum_clauses(2)[OF fin]
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   781
    using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close>
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   782
    by (auto simp: *)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   783
  ultimately show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   784
    unfolding affine_dependent_explicit
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   785
    apply (rule_tac x="insert a t" in exI, auto)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   786
    done
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   787
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   788
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   789
lemma affine_dependent_biggerset:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   790
  fixes s :: "'a::euclidean_space set"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   791
  assumes "finite s" "card s \<ge> DIM('a) + 2"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   792
  shows "affine_dependent s"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   793
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   794
  have "s \<noteq> {}" using assms by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   795
  then obtain a where "a\<in>s" by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   796
  have *: "{x - a |x. x \<in> s - {a}} = (\<lambda>x. x - a) ` (s - {a})"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   797
    by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   798
  have "card {x - a |x. x \<in> s - {a}} = card (s - {a})"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   799
    unfolding * by (simp add: card_image inj_on_def)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   800
  also have "\<dots> > DIM('a)" using assms(2)
74224
e04ec2b9ed97 some fixes connected with card_Diff_singleton
paulson <lp15@cam.ac.uk>
parents: 73372
diff changeset
   801
    unfolding card_Diff_singleton[OF \<open>a\<in>s\<close>] by auto
71242
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   802
  finally show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   803
    apply (subst insert_Diff[OF \<open>a\<in>s\<close>, symmetric])
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   804
    apply (rule dependent_imp_affine_dependent)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   805
    apply (rule dependent_biggerset, auto)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   806
    done
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   807
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   808
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   809
lemma affine_dependent_biggerset_general:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   810
  assumes "finite (S :: 'a::euclidean_space set)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   811
    and "card S \<ge> dim S + 2"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   812
  shows "affine_dependent S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   813
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   814
  from assms(2) have "S \<noteq> {}" by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   815
  then obtain a where "a\<in>S" by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   816
  have *: "{x - a |x. x \<in> S - {a}} = (\<lambda>x. x - a) ` (S - {a})"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   817
    by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   818
  have **: "card {x - a |x. x \<in> S - {a}} = card (S - {a})"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   819
    by (metis (no_types, lifting) "*" card_image diff_add_cancel inj_on_def)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   820
  have "dim {x - a |x. x \<in> S - {a}} \<le> dim S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   821
    using \<open>a\<in>S\<close> by (auto simp: span_base span_diff intro: subset_le_dim)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   822
  also have "\<dots> < dim S + 1" by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   823
  also have "\<dots> \<le> card (S - {a})"
74224
e04ec2b9ed97 some fixes connected with card_Diff_singleton
paulson <lp15@cam.ac.uk>
parents: 73372
diff changeset
   824
    using assms card_Diff_singleton[OF \<open>a\<in>S\<close>] by auto
71242
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   825
  finally show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   826
    apply (subst insert_Diff[OF \<open>a\<in>S\<close>, symmetric])
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   827
    apply (rule dependent_imp_affine_dependent)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   828
    apply (rule dependent_biggerset_general)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   829
    unfolding **
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   830
    apply auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   831
    done
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   832
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   833
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   834
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   835
subsection\<^marker>\<open>tag unimportant\<close> \<open>Some Properties of Affine Dependent Sets\<close>
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   836
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   837
lemma affine_independent_0 [simp]: "\<not> affine_dependent {}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   838
  by (simp add: affine_dependent_def)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   839
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   840
lemma affine_independent_1 [simp]: "\<not> affine_dependent {a}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   841
  by (simp add: affine_dependent_def)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   842
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   843
lemma affine_independent_2 [simp]: "\<not> affine_dependent {a,b}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   844
  by (simp add: affine_dependent_def insert_Diff_if hull_same)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   845
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   846
lemma affine_hull_translation: "affine hull ((\<lambda>x. a + x) `  S) = (\<lambda>x. a + x) ` (affine hull S)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   847
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   848
  have "affine ((\<lambda>x. a + x) ` (affine hull S))"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   849
    using affine_translation affine_affine_hull by blast
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   850
  moreover have "(\<lambda>x. a + x) `  S \<subseteq> (\<lambda>x. a + x) ` (affine hull S)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   851
    using hull_subset[of S] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   852
  ultimately have h1: "affine hull ((\<lambda>x. a + x) `  S) \<subseteq> (\<lambda>x. a + x) ` (affine hull S)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   853
    by (metis hull_minimal)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   854
  have "affine((\<lambda>x. -a + x) ` (affine hull ((\<lambda>x. a + x) `  S)))"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   855
    using affine_translation affine_affine_hull by blast
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   856
  moreover have "(\<lambda>x. -a + x) ` (\<lambda>x. a + x) `  S \<subseteq> (\<lambda>x. -a + x) ` (affine hull ((\<lambda>x. a + x) `  S))"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   857
    using hull_subset[of "(\<lambda>x. a + x) `  S"] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   858
  moreover have "S = (\<lambda>x. -a + x) ` (\<lambda>x. a + x) `  S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   859
    using translation_assoc[of "-a" a] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   860
  ultimately have "(\<lambda>x. -a + x) ` (affine hull ((\<lambda>x. a + x) `  S)) >= (affine hull S)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   861
    by (metis hull_minimal)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   862
  then have "affine hull ((\<lambda>x. a + x) ` S) >= (\<lambda>x. a + x) ` (affine hull S)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   863
    by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   864
  then show ?thesis using h1 by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   865
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   866
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   867
lemma affine_dependent_translation:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   868
  assumes "affine_dependent S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   869
  shows "affine_dependent ((\<lambda>x. a + x) ` S)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   870
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   871
  obtain x where x: "x \<in> S \<and> x \<in> affine hull (S - {x})"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   872
    using assms affine_dependent_def by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   873
  have "(+) a ` (S - {x}) = (+) a ` S - {a + x}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   874
    by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   875
  then have "a + x \<in> affine hull ((\<lambda>x. a + x) ` S - {a + x})"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   876
    using affine_hull_translation[of a "S - {x}"] x by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   877
  moreover have "a + x \<in> (\<lambda>x. a + x) ` S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   878
    using x by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   879
  ultimately show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   880
    unfolding affine_dependent_def by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   881
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   882
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   883
lemma affine_dependent_translation_eq:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   884
  "affine_dependent S \<longleftrightarrow> affine_dependent ((\<lambda>x. a + x) ` S)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   885
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   886
  {
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   887
    assume "affine_dependent ((\<lambda>x. a + x) ` S)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   888
    then have "affine_dependent S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   889
      using affine_dependent_translation[of "((\<lambda>x. a + x) ` S)" "-a"] translation_assoc[of "-a" a]
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   890
      by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   891
  }
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   892
  then show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   893
    using affine_dependent_translation by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   894
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   895
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   896
lemma affine_hull_0_dependent:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   897
  assumes "0 \<in> affine hull S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   898
  shows "dependent S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   899
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   900
  obtain s u where s_u: "finite s \<and> s \<noteq> {} \<and> s \<subseteq> S \<and> sum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = 0"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   901
    using assms affine_hull_explicit[of S] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   902
  then have "\<exists>v\<in>s. u v \<noteq> 0" by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   903
  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)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   904
    using s_u by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   905
  then show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   906
    unfolding dependent_explicit[of S] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   907
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   908
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   909
lemma affine_dependent_imp_dependent2:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   910
  assumes "affine_dependent (insert 0 S)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   911
  shows "dependent S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   912
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   913
  obtain x where x: "x \<in> insert 0 S \<and> x \<in> affine hull (insert 0 S - {x})"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   914
    using affine_dependent_def[of "(insert 0 S)"] assms by blast
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   915
  then have "x \<in> span (insert 0 S - {x})"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   916
    using affine_hull_subset_span by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   917
  moreover have "span (insert 0 S - {x}) = span (S - {x})"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   918
    using insert_Diff_if[of "0" S "{x}"] span_insert_0[of "S-{x}"] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   919
  ultimately have "x \<in> span (S - {x})" by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   920
  then have "x \<noteq> 0 \<Longrightarrow> dependent S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   921
    using x dependent_def by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   922
  moreover
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   923
  {
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   924
    assume "x = 0"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   925
    then have "0 \<in> affine hull S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   926
      using x hull_mono[of "S - {0}" S] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   927
    then have "dependent S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   928
      using affine_hull_0_dependent by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   929
  }
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   930
  ultimately show ?thesis by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   931
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   932
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   933
lemma affine_dependent_iff_dependent:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   934
  assumes "a \<notin> S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   935
  shows "affine_dependent (insert a S) \<longleftrightarrow> dependent ((\<lambda>x. -a + x) ` S)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   936
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   937
  have "((+) (- a) ` S) = {x - a| x . x \<in> S}" by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   938
  then show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   939
    using affine_dependent_translation_eq[of "(insert a S)" "-a"]
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   940
      affine_dependent_imp_dependent2 assms
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   941
      dependent_imp_affine_dependent[of a S]
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   942
    by (auto simp del: uminus_add_conv_diff)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   943
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   944
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   945
lemma affine_dependent_iff_dependent2:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   946
  assumes "a \<in> S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   947
  shows "affine_dependent S \<longleftrightarrow> dependent ((\<lambda>x. -a + x) ` (S-{a}))"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   948
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   949
  have "insert a (S - {a}) = S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   950
    using assms by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   951
  then show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   952
    using assms affine_dependent_iff_dependent[of a "S-{a}"] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   953
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   954
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   955
lemma affine_hull_insert_span_gen:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   956
  "affine hull (insert a s) = (\<lambda>x. a + x) ` span ((\<lambda>x. - a + x) ` s)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   957
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   958
  have h1: "{x - a |x. x \<in> s} = ((\<lambda>x. -a+x) ` s)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   959
    by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   960
  {
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   961
    assume "a \<notin> s"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   962
    then have ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   963
      using affine_hull_insert_span[of a s] h1 by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   964
  }
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   965
  moreover
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   966
  {
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   967
    assume a1: "a \<in> s"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   968
    have "\<exists>x. x \<in> s \<and> -a+x=0"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   969
      apply (rule exI[of _ a])
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   970
      using a1
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   971
      apply auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   972
      done
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   973
    then have "insert 0 ((\<lambda>x. -a+x) ` (s - {a})) = (\<lambda>x. -a+x) ` s"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   974
      by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   975
    then have "span ((\<lambda>x. -a+x) ` (s - {a}))=span ((\<lambda>x. -a+x) ` s)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   976
      using span_insert_0[of "(+) (- a) ` (s - {a})"] by (auto simp del: uminus_add_conv_diff)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   977
    moreover have "{x - a |x. x \<in> (s - {a})} = ((\<lambda>x. -a+x) ` (s - {a}))"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   978
      by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   979
    moreover have "insert a (s - {a}) = insert a s"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   980
      by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   981
    ultimately have ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   982
      using affine_hull_insert_span[of "a" "s-{a}"] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   983
  }
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   984
  ultimately show ?thesis by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   985
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   986
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   987
lemma affine_hull_span2:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   988
  assumes "a \<in> s"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   989
  shows "affine hull s = (\<lambda>x. a+x) ` span ((\<lambda>x. -a+x) ` (s-{a}))"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   990
  using affine_hull_insert_span_gen[of a "s - {a}", unfolded insert_Diff[OF assms]]
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   991
  by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   992
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   993
lemma affine_hull_span_gen:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   994
  assumes "a \<in> affine hull s"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   995
  shows "affine hull s = (\<lambda>x. a+x) ` span ((\<lambda>x. -a+x) ` s)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   996
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   997
  have "affine hull (insert a s) = affine hull s"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   998
    using hull_redundant[of a affine s] assms by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
   999
  then show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1000
    using affine_hull_insert_span_gen[of a "s"] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1001
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1002
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1003
lemma affine_hull_span_0:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1004
  assumes "0 \<in> affine hull S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1005
  shows "affine hull S = span S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1006
  using affine_hull_span_gen[of "0" S] assms by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1007
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1008
lemma extend_to_affine_basis_nonempty:
72492
2dd41a8893aa type class reduction
paulson <lp15@cam.ac.uk>
parents: 71840
diff changeset
  1009
  fixes S V :: "'n::real_vector set"
71242
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1010
  assumes "\<not> affine_dependent S" "S \<subseteq> V" "S \<noteq> {}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1011
  shows "\<exists>T. \<not> affine_dependent T \<and> S \<subseteq> T \<and> T \<subseteq> V \<and> affine hull T = affine hull V"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1012
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1013
  obtain a where a: "a \<in> S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1014
    using assms by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1015
  then have h0: "independent  ((\<lambda>x. -a + x) ` (S-{a}))"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1016
    using affine_dependent_iff_dependent2 assms by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1017
  obtain B where B:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1018
    "(\<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"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1019
    using assms
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1020
    by (blast intro: maximal_independent_subset_extend[OF _ h0, of "(\<lambda>x. -a + x) ` V"])
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1021
  define T where "T = (\<lambda>x. a+x) ` insert 0 B"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1022
  then have "T = insert a ((\<lambda>x. a+x) ` B)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1023
    by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1024
  then have "affine hull T = (\<lambda>x. a+x) ` span B"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1025
    using affine_hull_insert_span_gen[of a "((\<lambda>x. a+x) ` B)"] translation_assoc[of "-a" a B]
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1026
    by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1027
  then have "V \<subseteq> affine hull T"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1028
    using B assms translation_inverse_subset[of a V "span B"]
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1029
    by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1030
  moreover have "T \<subseteq> V"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1031
    using T_def B a assms by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1032
  ultimately have "affine hull T = affine hull V"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1033
    by (metis Int_absorb1 Int_absorb2 hull_hull hull_mono)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1034
  moreover have "S \<subseteq> T"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1035
    using T_def B translation_inverse_subset[of a "S-{a}" B]
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1036
    by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1037
  moreover have "\<not> affine_dependent T"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1038
    using T_def affine_dependent_translation_eq[of "insert 0 B"]
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1039
      affine_dependent_imp_dependent2 B
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1040
    by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1041
  ultimately show ?thesis using \<open>T \<subseteq> V\<close> by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1042
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1043
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1044
lemma affine_basis_exists:
72492
2dd41a8893aa type class reduction
paulson <lp15@cam.ac.uk>
parents: 71840
diff changeset
  1045
  fixes V :: "'n::real_vector set"
71242
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1046
  shows "\<exists>B. B \<subseteq> V \<and> \<not> affine_dependent B \<and> affine hull V = affine hull B"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1047
proof (cases "V = {}")
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1048
  case True
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1049
  then show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1050
    using affine_independent_0 by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1051
next
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1052
  case False
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1053
  then obtain x where "x \<in> V" by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1054
  then show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1055
    using affine_dependent_def[of "{x}"] extend_to_affine_basis_nonempty[of "{x}" V]
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1056
    by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1057
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1058
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1059
proposition extend_to_affine_basis:
72492
2dd41a8893aa type class reduction
paulson <lp15@cam.ac.uk>
parents: 71840
diff changeset
  1060
  fixes S V :: "'n::real_vector set"
71242
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1061
  assumes "\<not> affine_dependent S" "S \<subseteq> V"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1062
  obtains T where "\<not> affine_dependent T" "S \<subseteq> T" "T \<subseteq> V" "affine hull T = affine hull V"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1063
proof (cases "S = {}")
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1064
  case True then show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1065
    using affine_basis_exists by (metis empty_subsetI that)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1066
next
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1067
  case False
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1068
  then show ?thesis by (metis assms extend_to_affine_basis_nonempty that)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1069
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1070
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1071
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1072
subsection \<open>Affine Dimension of a Set\<close>
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1073
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1074
definition\<^marker>\<open>tag important\<close> aff_dim :: "('a::euclidean_space) set \<Rightarrow> int"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1075
  where "aff_dim V =
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1076
  (SOME d :: int.
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1077
    \<exists>B. affine hull B = affine hull V \<and> \<not> affine_dependent B \<and> of_nat (card B) = d + 1)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1078
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1079
lemma aff_dim_basis_exists:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1080
  fixes V :: "('n::euclidean_space) set"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1081
  shows "\<exists>B. affine hull B = affine hull V \<and> \<not> affine_dependent B \<and> of_nat (card B) = aff_dim V + 1"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1082
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1083
  obtain B where "\<not> affine_dependent B \<and> affine hull B = affine hull V"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1084
    using affine_basis_exists[of V] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1085
  then show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1086
    unfolding aff_dim_def
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1087
      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"]
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1088
    apply auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1089
    apply (rule exI[of _ "int (card B) - (1 :: int)"])
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1090
    apply (rule exI[of _ "B"], auto)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1091
    done
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1092
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1093
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1094
lemma affine_hull_eq_empty [simp]: "affine hull S = {} \<longleftrightarrow> S = {}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1095
by (metis affine_empty subset_empty subset_hull)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1096
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1097
lemma empty_eq_affine_hull[simp]: "{} = affine hull S \<longleftrightarrow> S = {}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1098
by (metis affine_hull_eq_empty)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1099
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1100
lemma aff_dim_parallel_subspace_aux:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1101
  fixes B :: "'n::euclidean_space set"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1102
  assumes "\<not> affine_dependent B" "a \<in> B"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1103
  shows "finite B \<and> ((card B) - 1 = dim (span ((\<lambda>x. -a+x) ` (B-{a}))))"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1104
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1105
  have "independent ((\<lambda>x. -a + x) ` (B-{a}))"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1106
    using affine_dependent_iff_dependent2 assms by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1107
  then have fin: "dim (span ((\<lambda>x. -a+x) ` (B-{a}))) = card ((\<lambda>x. -a + x) ` (B-{a}))"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1108
    "finite ((\<lambda>x. -a + x) ` (B - {a}))"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1109
    using indep_card_eq_dim_span[of "(\<lambda>x. -a+x) ` (B-{a})"] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1110
  show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1111
  proof (cases "(\<lambda>x. -a + x) ` (B - {a}) = {}")
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1112
    case True
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1113
    have "B = insert a ((\<lambda>x. a + x) ` (\<lambda>x. -a + x) ` (B - {a}))"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1114
      using translation_assoc[of "a" "-a" "(B - {a})"] assms by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1115
    then have "B = {a}" using True by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1116
    then show ?thesis using assms fin by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1117
  next
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1118
    case False
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1119
    then have "card ((\<lambda>x. -a + x) ` (B - {a})) > 0"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1120
      using fin by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1121
    moreover have h1: "card ((\<lambda>x. -a + x) ` (B-{a})) = card (B-{a})"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1122
      by (rule card_image) (use translate_inj_on in blast)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1123
    ultimately have "card (B-{a}) > 0" by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1124
    then have *: "finite (B - {a})"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1125
      using card_gt_0_iff[of "(B - {a})"] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1126
    then have "card (B - {a}) = card B - 1"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1127
      using card_Diff_singleton assms by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1128
    with * show ?thesis using fin h1 by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1129
  qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1130
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1131
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1132
lemma aff_dim_parallel_subspace:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1133
  fixes V L :: "'n::euclidean_space set"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1134
  assumes "V \<noteq> {}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1135
    and "subspace L"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1136
    and "affine_parallel (affine hull V) L"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1137
  shows "aff_dim V = int (dim L)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1138
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1139
  obtain B where
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1140
    B: "affine hull B = affine hull V \<and> \<not> affine_dependent B \<and> int (card B) = aff_dim V + 1"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1141
    using aff_dim_basis_exists by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1142
  then have "B \<noteq> {}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1143
    using assms B
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1144
    by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1145
  then obtain a where a: "a \<in> B" by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1146
  define Lb where "Lb = span ((\<lambda>x. -a+x) ` (B-{a}))"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1147
  moreover have "affine_parallel (affine hull B) Lb"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1148
    using Lb_def B assms affine_hull_span2[of a B] a
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1149
      affine_parallel_commut[of "Lb" "(affine hull B)"]
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1150
    unfolding affine_parallel_def
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1151
    by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1152
  moreover have "subspace Lb"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1153
    using Lb_def subspace_span by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1154
  moreover have "affine hull B \<noteq> {}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1155
    using assms B by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1156
  ultimately have "L = Lb"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1157
    using assms affine_parallel_subspace[of "affine hull B"] affine_affine_hull[of B] B
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1158
    by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1159
  then have "dim L = dim Lb"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1160
    by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1161
  moreover have "card B - 1 = dim Lb" and "finite B"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1162
    using Lb_def aff_dim_parallel_subspace_aux a B by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1163
  ultimately show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1164
    using B \<open>B \<noteq> {}\<close> card_gt_0_iff[of B] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1165
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1166
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1167
lemma aff_independent_finite:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1168
  fixes B :: "'n::euclidean_space set"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1169
  assumes "\<not> affine_dependent B"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1170
  shows "finite B"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1171
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1172
  {
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1173
    assume "B \<noteq> {}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1174
    then obtain a where "a \<in> B" by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1175
    then have ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1176
      using aff_dim_parallel_subspace_aux assms by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1177
  }
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1178
  then show ?thesis by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1179
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1180
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1181
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1182
lemma aff_dim_empty:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1183
  fixes S :: "'n::euclidean_space set"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1184
  shows "S = {} \<longleftrightarrow> aff_dim S = -1"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1185
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1186
  obtain B where *: "affine hull B = affine hull S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1187
    and "\<not> affine_dependent B"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1188
    and "int (card B) = aff_dim S + 1"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1189
    using aff_dim_basis_exists by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1190
  moreover
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1191
  from * have "S = {} \<longleftrightarrow> B = {}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1192
    by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1193
  ultimately show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1194
    using aff_independent_finite[of B] card_gt_0_iff[of B] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1195
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1196
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1197
lemma aff_dim_empty_eq [simp]: "aff_dim ({}::'a::euclidean_space set) = -1"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1198
  by (simp add: aff_dim_empty [symmetric])
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1199
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1200
lemma aff_dim_affine_hull [simp]: "aff_dim (affine hull S) = aff_dim S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1201
  unfolding aff_dim_def using hull_hull[of _ S] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1202
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1203
lemma aff_dim_affine_hull2:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1204
  assumes "affine hull S = affine hull T"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1205
  shows "aff_dim S = aff_dim T"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1206
  unfolding aff_dim_def using assms by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1207
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1208
lemma aff_dim_unique:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1209
  fixes B V :: "'n::euclidean_space set"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1210
  assumes "affine hull B = affine hull V \<and> \<not> affine_dependent B"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1211
  shows "of_nat (card B) = aff_dim V + 1"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1212
proof (cases "B = {}")
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1213
  case True
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1214
  then have "V = {}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1215
    using assms
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1216
    by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1217
  then have "aff_dim V = (-1::int)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1218
    using aff_dim_empty by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1219
  then show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1220
    using \<open>B = {}\<close> by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1221
next
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1222
  case False
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1223
  then obtain a where a: "a \<in> B" by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1224
  define Lb where "Lb = span ((\<lambda>x. -a+x) ` (B-{a}))"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1225
  have "affine_parallel (affine hull B) Lb"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1226
    using Lb_def affine_hull_span2[of a B] a
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1227
      affine_parallel_commut[of "Lb" "(affine hull B)"]
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1228
    unfolding affine_parallel_def by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1229
  moreover have "subspace Lb"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1230
    using Lb_def subspace_span by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1231
  ultimately have "aff_dim B = int(dim Lb)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1232
    using aff_dim_parallel_subspace[of B Lb] \<open>B \<noteq> {}\<close> by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1233
  moreover have "(card B) - 1 = dim Lb" "finite B"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1234
    using Lb_def aff_dim_parallel_subspace_aux a assms by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1235
  ultimately have "of_nat (card B) = aff_dim B + 1"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1236
    using \<open>B \<noteq> {}\<close> card_gt_0_iff[of B] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1237
  then show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1238
    using aff_dim_affine_hull2 assms by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1239
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1240
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1241
lemma aff_dim_affine_independent:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1242
  fixes B :: "'n::euclidean_space set"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1243
  assumes "\<not> affine_dependent B"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1244
  shows "of_nat (card B) = aff_dim B + 1"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1245
  using aff_dim_unique[of B B] assms by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1246
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1247
lemma affine_independent_iff_card:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1248
    fixes s :: "'a::euclidean_space set"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1249
    shows "\<not> affine_dependent s \<longleftrightarrow> finite s \<and> aff_dim s = int(card s) - 1"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1250
  apply (rule iffI)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1251
  apply (simp add: aff_dim_affine_independent aff_independent_finite)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1252
  by (metis affine_basis_exists [of s] aff_dim_unique card_subset_eq diff_add_cancel of_nat_eq_iff)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1253
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1254
lemma aff_dim_sing [simp]:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1255
  fixes a :: "'n::euclidean_space"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1256
  shows "aff_dim {a} = 0"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1257
  using aff_dim_affine_independent[of "{a}"] affine_independent_1 by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1258
72492
2dd41a8893aa type class reduction
paulson <lp15@cam.ac.uk>
parents: 71840
diff changeset
  1259
lemma aff_dim_2 [simp]:
2dd41a8893aa type class reduction
paulson <lp15@cam.ac.uk>
parents: 71840
diff changeset
  1260
  fixes a :: "'n::euclidean_space"
2dd41a8893aa type class reduction
paulson <lp15@cam.ac.uk>
parents: 71840
diff changeset
  1261
  shows "aff_dim {a,b} = (if a = b then 0 else 1)"
71242
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1262
proof (clarsimp)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1263
  assume "a \<noteq> b"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1264
  then have "aff_dim{a,b} = card{a,b} - 1"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1265
    using affine_independent_2 [of a b] aff_dim_affine_independent by fastforce
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1266
  also have "\<dots> = 1"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1267
    using \<open>a \<noteq> b\<close> by simp
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1268
  finally show "aff_dim {a, b} = 1" .
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1269
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1270
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1271
lemma aff_dim_inner_basis_exists:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1272
  fixes V :: "('n::euclidean_space) set"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1273
  shows "\<exists>B. B \<subseteq> V \<and> affine hull B = affine hull V \<and>
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1274
    \<not> affine_dependent B \<and> of_nat (card B) = aff_dim V + 1"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1275
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1276
  obtain B where B: "\<not> affine_dependent B" "B \<subseteq> V" "affine hull B = affine hull V"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1277
    using affine_basis_exists[of V] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1278
  then have "of_nat(card B) = aff_dim V+1" using aff_dim_unique by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1279
  with B show ?thesis by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1280
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1281
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1282
lemma aff_dim_le_card:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1283
  fixes V :: "'n::euclidean_space set"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1284
  assumes "finite V"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1285
  shows "aff_dim V \<le> of_nat (card V) - 1"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1286
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1287
  obtain B where B: "B \<subseteq> V" "of_nat (card B) = aff_dim V + 1"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1288
    using aff_dim_inner_basis_exists[of V] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1289
  then have "card B \<le> card V"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1290
    using assms card_mono by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1291
  with B show ?thesis by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1292
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1293
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1294
lemma aff_dim_parallel_eq:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1295
  fixes S T :: "'n::euclidean_space set"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1296
  assumes "affine_parallel (affine hull S) (affine hull T)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1297
  shows "aff_dim S = aff_dim T"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1298
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1299
  {
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1300
    assume "T \<noteq> {}" "S \<noteq> {}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1301
    then obtain L where L: "subspace L \<and> affine_parallel (affine hull T) L"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1302
      using affine_parallel_subspace[of "affine hull T"]
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1303
        affine_affine_hull[of T]
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1304
      by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1305
    then have "aff_dim T = int (dim L)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1306
      using aff_dim_parallel_subspace \<open>T \<noteq> {}\<close> by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1307
    moreover have *: "subspace L \<and> affine_parallel (affine hull S) L"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1308
       using L affine_parallel_assoc[of "affine hull S" "affine hull T" L] assms by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1309
    moreover from * have "aff_dim S = int (dim L)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1310
      using aff_dim_parallel_subspace \<open>S \<noteq> {}\<close> by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1311
    ultimately have ?thesis by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1312
  }
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1313
  moreover
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1314
  {
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1315
    assume "S = {}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1316
    then have "S = {}" and "T = {}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1317
      using assms
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1318
      unfolding affine_parallel_def
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1319
      by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1320
    then have ?thesis using aff_dim_empty by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1321
  }
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1322
  moreover
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1323
  {
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1324
    assume "T = {}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1325
    then have "S = {}" and "T = {}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1326
      using assms
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1327
      unfolding affine_parallel_def
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1328
      by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1329
    then have ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1330
      using aff_dim_empty by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1331
  }
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1332
  ultimately show ?thesis by blast
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1333
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1334
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1335
lemma aff_dim_translation_eq:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1336
  "aff_dim ((+) a ` S) = aff_dim S" for a :: "'n::euclidean_space"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1337
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1338
  have "affine_parallel (affine hull S) (affine hull ((\<lambda>x. a + x) ` S))"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1339
    unfolding affine_parallel_def
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1340
    apply (rule exI[of _ "a"])
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1341
    using affine_hull_translation[of a S]
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1342
    apply auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1343
    done
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1344
  then show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1345
    using aff_dim_parallel_eq[of S "(\<lambda>x. a + x) ` S"] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1346
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1347
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1348
lemma aff_dim_translation_eq_subtract:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1349
  "aff_dim ((\<lambda>x. x - a) ` S) = aff_dim S" for a :: "'n::euclidean_space"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1350
  using aff_dim_translation_eq [of "- a"] by (simp cong: image_cong_simp)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1351
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1352
lemma aff_dim_affine:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1353
  fixes S L :: "'n::euclidean_space set"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1354
  assumes "S \<noteq> {}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1355
    and "affine S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1356
    and "subspace L"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1357
    and "affine_parallel S L"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1358
  shows "aff_dim S = int (dim L)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1359
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1360
  have *: "affine hull S = S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1361
    using assms affine_hull_eq[of S] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1362
  then have "affine_parallel (affine hull S) L"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1363
    using assms by (simp add: *)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1364
  then show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1365
    using assms aff_dim_parallel_subspace[of S L] by blast
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1366
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1367
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1368
lemma dim_affine_hull:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1369
  fixes S :: "'n::euclidean_space set"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1370
  shows "dim (affine hull S) = dim S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1371
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1372
  have "dim (affine hull S) \<ge> dim S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1373
    using dim_subset by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1374
  moreover have "dim (span S) \<ge> dim (affine hull S)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1375
    using dim_subset affine_hull_subset_span by blast
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1376
  moreover have "dim (span S) = dim S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1377
    using dim_span by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1378
  ultimately show ?thesis by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1379
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1380
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1381
lemma aff_dim_subspace:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1382
  fixes S :: "'n::euclidean_space set"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1383
  assumes "subspace S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1384
  shows "aff_dim S = int (dim S)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1385
proof (cases "S={}")
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1386
  case True with assms show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1387
    by (simp add: subspace_affine)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1388
next
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1389
  case False
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1390
  with aff_dim_affine[of S S] assms subspace_imp_affine[of S] affine_parallel_reflex[of S] subspace_affine
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1391
  show ?thesis by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1392
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1393
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1394
lemma aff_dim_zero:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1395
  fixes S :: "'n::euclidean_space set"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1396
  assumes "0 \<in> affine hull S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1397
  shows "aff_dim S = int (dim S)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1398
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1399
  have "subspace (affine hull S)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1400
    using subspace_affine[of "affine hull S"] affine_affine_hull assms
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1401
    by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1402
  then have "aff_dim (affine hull S) = int (dim (affine hull S))"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1403
    using assms aff_dim_subspace[of "affine hull S"] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1404
  then show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1405
    using aff_dim_affine_hull[of S] dim_affine_hull[of S]
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1406
    by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1407
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1408
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1409
lemma aff_dim_eq_dim:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1410
  "aff_dim S = int (dim ((+) (- a) ` S))" if "a \<in> affine hull S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1411
    for S :: "'n::euclidean_space set"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1412
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1413
  have "0 \<in> affine hull (+) (- a) ` S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1414
    unfolding affine_hull_translation
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1415
    using that by (simp add: ac_simps)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1416
  with aff_dim_zero show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1417
    by (metis aff_dim_translation_eq)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1418
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1419
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1420
lemma aff_dim_eq_dim_subtract:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1421
  "aff_dim S = int (dim ((\<lambda>x. x - a) ` S))" if "a \<in> affine hull S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1422
    for S :: "'n::euclidean_space set"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1423
  using aff_dim_eq_dim [of a] that by (simp cong: image_cong_simp)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1424
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1425
lemma aff_dim_UNIV [simp]: "aff_dim (UNIV :: 'n::euclidean_space set) = int(DIM('n))"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1426
  using aff_dim_subspace[of "(UNIV :: 'n::euclidean_space set)"]
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1427
    dim_UNIV[where 'a="'n::euclidean_space"]
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1428
  by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1429
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1430
lemma aff_dim_geq:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1431
  fixes V :: "'n::euclidean_space set"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1432
  shows "aff_dim V \<ge> -1"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1433
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1434
  obtain B where "affine hull B = affine hull V"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1435
    and "\<not> affine_dependent B"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1436
    and "int (card B) = aff_dim V + 1"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1437
    using aff_dim_basis_exists by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1438
  then show ?thesis by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1439
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1440
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1441
lemma aff_dim_negative_iff [simp]:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1442
  fixes S :: "'n::euclidean_space set"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1443
  shows "aff_dim S < 0 \<longleftrightarrow>S = {}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1444
by (metis aff_dim_empty aff_dim_geq diff_0 eq_iff zle_diff1_eq)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1445
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1446
lemma aff_lowdim_subset_hyperplane:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1447
  fixes S :: "'a::euclidean_space set"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1448
  assumes "aff_dim S < DIM('a)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1449
  obtains a b where "a \<noteq> 0" "S \<subseteq> {x. a \<bullet> x = b}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1450
proof (cases "S={}")
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1451
  case True
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1452
  moreover
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1453
  have "(SOME b. b \<in> Basis) \<noteq> 0"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1454
    by (metis norm_some_Basis norm_zero zero_neq_one)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1455
  ultimately show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1456
    using that by blast
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1457
next
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1458
  case False
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1459
  then obtain c S' where "c \<notin> S'" "S = insert c S'"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1460
    by (meson equals0I mk_disjoint_insert)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1461
  have "dim ((+) (-c) ` S) < DIM('a)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1462
    by (metis \<open>S = insert c S'\<close> aff_dim_eq_dim assms hull_inc insertI1 of_nat_less_imp_less)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1463
  then obtain a where "a \<noteq> 0" "span ((+) (-c) ` S) \<subseteq> {x. a \<bullet> x = 0}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1464
    using lowdim_subset_hyperplane by blast
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1465
  moreover
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1466
  have "a \<bullet> w = a \<bullet> c" if "span ((+) (- c) ` S) \<subseteq> {x. a \<bullet> x = 0}" "w \<in> S" for w
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1467
  proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1468
    have "w-c \<in> span ((+) (- c) ` S)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1469
      by (simp add: span_base \<open>w \<in> S\<close>)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1470
    with that have "w-c \<in> {x. a \<bullet> x = 0}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1471
      by blast
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1472
    then show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1473
      by (auto simp: algebra_simps)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1474
  qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1475
  ultimately have "S \<subseteq> {x. a \<bullet> x = a \<bullet> c}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1476
    by blast
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1477
  then show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1478
    by (rule that[OF \<open>a \<noteq> 0\<close>])
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1479
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1480
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1481
lemma affine_independent_card_dim_diffs:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1482
  fixes S :: "'a :: euclidean_space set"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1483
  assumes "\<not> affine_dependent S" "a \<in> S"
72567
aeac6424d3b5 cleanup
paulson <lp15@cam.ac.uk>
parents: 72492
diff changeset
  1484
    shows "card S = dim ((\<lambda>x. x - a) ` S) + 1"
71242
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1485
proof -
72567
aeac6424d3b5 cleanup
paulson <lp15@cam.ac.uk>
parents: 72492
diff changeset
  1486
  have non: "\<not> affine_dependent (insert a S)"
71242
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1487
    by (simp add: assms insert_absorb)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1488
  have "finite S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1489
    by (meson assms aff_independent_finite)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1490
  with \<open>a \<in> S\<close> have "card S \<noteq> 0" by auto
72567
aeac6424d3b5 cleanup
paulson <lp15@cam.ac.uk>
parents: 72492
diff changeset
  1491
  moreover have "dim ((\<lambda>x. x - a) ` S) = card S - 1"
aeac6424d3b5 cleanup
paulson <lp15@cam.ac.uk>
parents: 72492
diff changeset
  1492
    using aff_dim_eq_dim_subtract aff_dim_unique \<open>a \<in> S\<close> hull_inc insert_absorb non by fastforce
71242
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1493
  ultimately show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1494
    by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1495
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1496
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1497
lemma independent_card_le_aff_dim:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1498
  fixes B :: "'n::euclidean_space set"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1499
  assumes "B \<subseteq> V"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1500
  assumes "\<not> affine_dependent B"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1501
  shows "int (card B) \<le> aff_dim V + 1"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1502
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1503
  obtain T where T: "\<not> affine_dependent T \<and> B \<subseteq> T \<and> T \<subseteq> V \<and> affine hull T = affine hull V"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1504
    by (metis assms extend_to_affine_basis[of B V])
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1505
  then have "of_nat (card T) = aff_dim V + 1"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1506
    using aff_dim_unique by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1507
  then show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1508
    using T card_mono[of T B] aff_independent_finite[of T] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1509
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1510
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1511
lemma aff_dim_subset:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1512
  fixes S T :: "'n::euclidean_space set"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1513
  assumes "S \<subseteq> T"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1514
  shows "aff_dim S \<le> aff_dim T"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1515
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1516
  obtain B where B: "\<not> affine_dependent B" "B \<subseteq> S" "affine hull B = affine hull S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1517
    "of_nat (card B) = aff_dim S + 1"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1518
    using aff_dim_inner_basis_exists[of S] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1519
  then have "int (card B) \<le> aff_dim T + 1"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1520
    using assms independent_card_le_aff_dim[of B T] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1521
  with B show ?thesis by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1522
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1523
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1524
lemma aff_dim_le_DIM:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1525
  fixes S :: "'n::euclidean_space set"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1526
  shows "aff_dim S \<le> int (DIM('n))"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1527
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1528
  have "aff_dim (UNIV :: 'n::euclidean_space set) = int(DIM('n))"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1529
    using aff_dim_UNIV by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1530
  then show "aff_dim (S:: 'n::euclidean_space set) \<le> int(DIM('n))"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1531
    using aff_dim_subset[of S "(UNIV :: ('n::euclidean_space) set)"] subset_UNIV by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1532
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1533
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1534
lemma affine_dim_equal:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1535
  fixes S :: "'n::euclidean_space set"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1536
  assumes "affine S" "affine T" "S \<noteq> {}" "S \<subseteq> T" "aff_dim S = aff_dim T"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1537
  shows "S = T"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1538
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1539
  obtain a where "a \<in> S" using assms by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1540
  then have "a \<in> T" using assms by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1541
  define LS where "LS = {y. \<exists>x \<in> S. (-a) + x = y}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1542
  then have ls: "subspace LS" "affine_parallel S LS"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1543
    using assms parallel_subspace_explicit[of S a LS] \<open>a \<in> S\<close> by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1544
  then have h1: "int(dim LS) = aff_dim S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1545
    using assms aff_dim_affine[of S LS] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1546
  have "T \<noteq> {}" using assms by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1547
  define LT where "LT = {y. \<exists>x \<in> T. (-a) + x = y}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1548
  then have lt: "subspace LT \<and> affine_parallel T LT"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1549
    using assms parallel_subspace_explicit[of T a LT] \<open>a \<in> T\<close> by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1550
  then have "int(dim LT) = aff_dim T"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1551
    using assms aff_dim_affine[of T LT] \<open>T \<noteq> {}\<close> by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1552
  then have "dim LS = dim LT"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1553
    using h1 assms by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1554
  moreover have "LS \<le> LT"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1555
    using LS_def LT_def assms by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1556
  ultimately have "LS = LT"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1557
    using subspace_dim_equal[of LS LT] ls lt by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1558
  moreover have "S = {x. \<exists>y \<in> LS. a+y=x}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1559
    using LS_def by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1560
  moreover have "T = {x. \<exists>y \<in> LT. a+y=x}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1561
    using LT_def by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1562
  ultimately show ?thesis by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1563
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1564
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1565
lemma aff_dim_eq_0:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1566
  fixes S :: "'a::euclidean_space set"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1567
  shows "aff_dim S = 0 \<longleftrightarrow> (\<exists>a. S = {a})"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1568
proof (cases "S = {}")
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1569
  case True
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1570
  then show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1571
    by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1572
next
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1573
  case False
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1574
  then obtain a where "a \<in> S" by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1575
  show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1576
  proof safe
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1577
    assume 0: "aff_dim S = 0"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1578
    have "\<not> {a,b} \<subseteq> S" if "b \<noteq> a" for b
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1579
      by (metis "0" aff_dim_2 aff_dim_subset not_one_le_zero that)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1580
    then show "\<exists>a. S = {a}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1581
      using \<open>a \<in> S\<close> by blast
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1582
  qed auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1583
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1584
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1585
lemma affine_hull_UNIV:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1586
  fixes S :: "'n::euclidean_space set"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1587
  assumes "aff_dim S = int(DIM('n))"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1588
  shows "affine hull S = (UNIV :: ('n::euclidean_space) set)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1589
proof -
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1590
  have "S \<noteq> {}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1591
    using assms aff_dim_empty[of S] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1592
  have h0: "S \<subseteq> affine hull S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1593
    using hull_subset[of S _] by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1594
  have h1: "aff_dim (UNIV :: ('n::euclidean_space) set) = aff_dim S"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1595
    using aff_dim_UNIV assms by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1596
  then have h2: "aff_dim (affine hull S) \<le> aff_dim (UNIV :: ('n::euclidean_space) set)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1597
    using aff_dim_le_DIM[of "affine hull S"] assms h0 by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1598
  have h3: "aff_dim S \<le> aff_dim (affine hull S)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1599
    using h0 aff_dim_subset[of S "affine hull S"] assms by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1600
  then have h4: "aff_dim (affine hull S) = aff_dim (UNIV :: ('n::euclidean_space) set)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1601
    using h0 h1 h2 by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1602
  then show ?thesis
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1603
    using affine_dim_equal[of "affine hull S" "(UNIV :: ('n::euclidean_space) set)"]
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1604
      affine_affine_hull[of S] affine_UNIV assms h4 h0 \<open>S \<noteq> {}\<close>
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1605
    by auto
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1606
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1607
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1608
lemma disjoint_affine_hull:
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1609
  fixes s :: "'n::euclidean_space set"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1610
  assumes "\<not> affine_dependent s" "t \<subseteq> s" "u \<subseteq> s" "t \<inter> u = {}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1611
    shows "(affine hull t) \<inter> (affine hull u) = {}"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1612
proof -
73372
10b9b3341c26 tuned proofs;
wenzelm
parents: 72567
diff changeset
  1613
  from assms(1) have "finite s"
10b9b3341c26 tuned proofs;
wenzelm
parents: 72567
diff changeset
  1614
    by (simp add: aff_independent_finite)
10b9b3341c26 tuned proofs;
wenzelm
parents: 72567
diff changeset
  1615
  with assms(2,3) have "finite t" "finite u"
10b9b3341c26 tuned proofs;
wenzelm
parents: 72567
diff changeset
  1616
    by (blast intro: finite_subset)+
10b9b3341c26 tuned proofs;
wenzelm
parents: 72567
diff changeset
  1617
  have False if "y \<in> affine hull t" and "y \<in> affine hull u" for y
10b9b3341c26 tuned proofs;
wenzelm
parents: 72567
diff changeset
  1618
  proof -
10b9b3341c26 tuned proofs;
wenzelm
parents: 72567
diff changeset
  1619
    from that obtain a b
10b9b3341c26 tuned proofs;
wenzelm
parents: 72567
diff changeset
  1620
      where a1 [simp]: "sum a t = 1"
10b9b3341c26 tuned proofs;
wenzelm
parents: 72567
diff changeset
  1621
        and [simp]: "sum (\<lambda>v. a v *\<^sub>R v) t = y"
10b9b3341c26 tuned proofs;
wenzelm
parents: 72567
diff changeset
  1622
        and [simp]: "sum b u = 1" "sum (\<lambda>v. b v *\<^sub>R v) u = y"
71242
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1623
      by (auto simp: affine_hull_finite \<open>finite t\<close> \<open>finite u\<close>)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1624
    define c where "c x = (if x \<in> t then a x else if x \<in> u then -(b x) else 0)" for x
73372
10b9b3341c26 tuned proofs;
wenzelm
parents: 72567
diff changeset
  1625
    from assms(2,3,4) have [simp]: "s \<inter> t = t" "s \<inter> - t \<inter> u = u"
10b9b3341c26 tuned proofs;
wenzelm
parents: 72567
diff changeset
  1626
      by auto
71242
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1627
    have "sum c s = 0"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1628
      by (simp add: c_def comm_monoid_add_class.sum.If_cases \<open>finite s\<close> sum_negf)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1629
    moreover have "\<not> (\<forall>v\<in>s. c v = 0)"
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1630
      by (metis (no_types) IntD1 \<open>s \<inter> t = t\<close> a1 c_def sum.neutral zero_neq_one)
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1631
    moreover have "(\<Sum>v\<in>s. c v *\<^sub>R v) = 0"
73372
10b9b3341c26 tuned proofs;
wenzelm
parents: 72567
diff changeset
  1632
      by (simp add: c_def if_smult sum_negf comm_monoid_add_class.sum.If_cases \<open>finite s\<close>)
10b9b3341c26 tuned proofs;
wenzelm
parents: 72567
diff changeset
  1633
    ultimately show ?thesis
10b9b3341c26 tuned proofs;
wenzelm
parents: 72567
diff changeset
  1634
      using assms(1) \<open>finite s\<close> by (auto simp: affine_dependent_explicit)
10b9b3341c26 tuned proofs;
wenzelm
parents: 72567
diff changeset
  1635
  qed
71242
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1636
  then show ?thesis by blast
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1637
qed
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1638
ec5090faf541 separated Affine theory from Convex
nipkow
parents:
diff changeset
  1639
end