| author | huffman | 
| Mon, 02 Nov 2009 18:49:53 -0800 | |
| changeset 33402 | d9a25a87da4a | 
| parent 32960 | 69916a850301 | 
| child 44190 | fe5504984937 | 
| permissions | -rw-r--r-- | 
| 31795 
be3e1cc5005c
standard naming conventions for session and theories;
 wenzelm parents: 
30729diff
changeset | 1 | (* Title: HOL/Hahn_Banach/Subspace.thy | 
| 7566 | 2 | Author: Gertrud Bauer, TU Munich | 
| 3 | *) | |
| 7535 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 wenzelm parents: diff
changeset | 4 | |
| 9035 | 5 | header {* Subspaces *}
 | 
| 7808 | 6 | |
| 27612 | 7 | theory Subspace | 
| 31795 
be3e1cc5005c
standard naming conventions for session and theories;
 wenzelm parents: 
30729diff
changeset | 8 | imports Vector_Space | 
| 27612 | 9 | begin | 
| 7535 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 wenzelm parents: diff
changeset | 10 | |
| 9035 | 11 | subsection {* Definition *}
 | 
| 7535 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 wenzelm parents: diff
changeset | 12 | |
| 10687 | 13 | text {*
 | 
| 14 |   A non-empty subset @{text U} of a vector space @{text V} is a
 | |
| 15 |   \emph{subspace} of @{text V}, iff @{text U} is closed under addition
 | |
| 16 | and scalar multiplication. | |
| 17 | *} | |
| 7917 | 18 | |
| 29234 | 19 | locale subspace = | 
| 20 |   fixes U :: "'a\<Colon>{minus, plus, zero, uminus} set" and V
 | |
| 13515 | 21 |   assumes non_empty [iff, intro]: "U \<noteq> {}"
 | 
| 22 | and subset [iff]: "U \<subseteq> V" | |
| 23 | and add_closed [iff]: "x \<in> U \<Longrightarrow> y \<in> U \<Longrightarrow> x + y \<in> U" | |
| 24 | and mult_closed [iff]: "x \<in> U \<Longrightarrow> a \<cdot> x \<in> U" | |
| 7535 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 wenzelm parents: diff
changeset | 25 | |
| 21210 | 26 | notation (symbols) | 
| 19736 | 27 | subspace (infix "\<unlhd>" 50) | 
| 14254 
342634f38451
Isar/Locales: <loc>.intro and <loc>.axioms no longer intro? and elim? by
 ballarin parents: 
13547diff
changeset | 28 | |
| 19736 | 29 | declare vectorspace.intro [intro?] subspace.intro [intro?] | 
| 7535 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 wenzelm parents: diff
changeset | 30 | |
| 13515 | 31 | lemma subspace_subset [elim]: "U \<unlhd> V \<Longrightarrow> U \<subseteq> V" | 
| 32 | by (rule subspace.subset) | |
| 7566 | 33 | |
| 13515 | 34 | lemma (in subspace) subsetD [iff]: "x \<in> U \<Longrightarrow> x \<in> V" | 
| 35 | using subset by blast | |
| 7566 | 36 | |
| 13515 | 37 | lemma subspaceD [elim]: "U \<unlhd> V \<Longrightarrow> x \<in> U \<Longrightarrow> x \<in> V" | 
| 38 | by (rule subspace.subsetD) | |
| 7535 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 wenzelm parents: diff
changeset | 39 | |
| 13515 | 40 | lemma rev_subspaceD [elim?]: "x \<in> U \<Longrightarrow> U \<unlhd> V \<Longrightarrow> x \<in> V" | 
| 41 | by (rule subspace.subsetD) | |
| 42 | ||
| 13547 | 43 | lemma (in subspace) diff_closed [iff]: | 
| 27611 | 44 | assumes "vectorspace V" | 
| 27612 | 45 | assumes x: "x \<in> U" and y: "y \<in> U" | 
| 46 | shows "x - y \<in> U" | |
| 27611 | 47 | proof - | 
| 29234 | 48 | interpret vectorspace V by fact | 
| 27612 | 49 | from x y show ?thesis by (simp add: diff_eq1 negate_eq1) | 
| 27611 | 50 | qed | 
| 7917 | 51 | |
| 13515 | 52 | text {*
 | 
| 53 | \medskip Similar as for linear spaces, the existence of the zero | |
| 54 | element in every subspace follows from the non-emptiness of the | |
| 55 | carrier set and by vector space laws. | |
| 56 | *} | |
| 57 | ||
| 13547 | 58 | lemma (in subspace) zero [intro]: | 
| 27611 | 59 | assumes "vectorspace V" | 
| 13547 | 60 | shows "0 \<in> U" | 
| 10687 | 61 | proof - | 
| 30729 
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
 wenzelm parents: 
29252diff
changeset | 62 | interpret V: vectorspace V by fact | 
| 29234 | 63 |   have "U \<noteq> {}" by (rule non_empty)
 | 
| 13515 | 64 | then obtain x where x: "x \<in> U" by blast | 
| 27612 | 65 | then have "x \<in> V" .. then have "0 = x - x" by simp | 
| 29234 | 66 | also from `vectorspace V` x x have "\<dots> \<in> U" by (rule diff_closed) | 
| 13515 | 67 | finally show ?thesis . | 
| 9035 | 68 | qed | 
| 7535 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 wenzelm parents: diff
changeset | 69 | |
| 13547 | 70 | lemma (in subspace) neg_closed [iff]: | 
| 27611 | 71 | assumes "vectorspace V" | 
| 27612 | 72 | assumes x: "x \<in> U" | 
| 73 | shows "- x \<in> U" | |
| 27611 | 74 | proof - | 
| 29234 | 75 | interpret vectorspace V by fact | 
| 27612 | 76 | from x show ?thesis by (simp add: negate_eq1) | 
| 27611 | 77 | qed | 
| 13515 | 78 | |
| 10687 | 79 | text {* \medskip Further derived laws: every subspace is a vector space. *}
 | 
| 7535 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 wenzelm parents: diff
changeset | 80 | |
| 13547 | 81 | lemma (in subspace) vectorspace [iff]: | 
| 27611 | 82 | assumes "vectorspace V" | 
| 13547 | 83 | shows "vectorspace U" | 
| 27611 | 84 | proof - | 
| 29234 | 85 | interpret vectorspace V by fact | 
| 27612 | 86 | show ?thesis | 
| 87 | proof | |
| 27611 | 88 |     show "U \<noteq> {}" ..
 | 
| 89 | fix x y z assume x: "x \<in> U" and y: "y \<in> U" and z: "z \<in> U" | |
| 90 | fix a b :: real | |
| 91 | from x y show "x + y \<in> U" by simp | |
| 92 | from x show "a \<cdot> x \<in> U" by simp | |
| 93 | from x y z show "(x + y) + z = x + (y + z)" by (simp add: add_ac) | |
| 94 | from x y show "x + y = y + x" by (simp add: add_ac) | |
| 95 | from x show "x - x = 0" by simp | |
| 96 | from x show "0 + x = x" by simp | |
| 97 | from x y show "a \<cdot> (x + y) = a \<cdot> x + a \<cdot> y" by (simp add: distrib) | |
| 98 | from x show "(a + b) \<cdot> x = a \<cdot> x + b \<cdot> x" by (simp add: distrib) | |
| 99 | from x show "(a * b) \<cdot> x = a \<cdot> b \<cdot> x" by (simp add: mult_assoc) | |
| 100 | from x show "1 \<cdot> x = x" by simp | |
| 101 | from x show "- x = - 1 \<cdot> x" by (simp add: negate_eq1) | |
| 102 | from x y show "x - y = x + - y" by (simp add: diff_eq1) | |
| 103 | qed | |
| 9035 | 104 | qed | 
| 7535 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 wenzelm parents: diff
changeset | 105 | |
| 13515 | 106 | |
| 9035 | 107 | text {* The subspace relation is reflexive. *}
 | 
| 7917 | 108 | |
| 13515 | 109 | lemma (in vectorspace) subspace_refl [intro]: "V \<unlhd> V" | 
| 10687 | 110 | proof | 
| 13515 | 111 |   show "V \<noteq> {}" ..
 | 
| 10687 | 112 | show "V \<subseteq> V" .. | 
| 13515 | 113 | fix x y assume x: "x \<in> V" and y: "y \<in> V" | 
| 114 | fix a :: real | |
| 115 | from x y show "x + y \<in> V" by simp | |
| 116 | from x show "a \<cdot> x \<in> V" by simp | |
| 9035 | 117 | qed | 
| 7535 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 wenzelm parents: diff
changeset | 118 | |
| 9035 | 119 | text {* The subspace relation is transitive. *}
 | 
| 7917 | 120 | |
| 13515 | 121 | lemma (in vectorspace) subspace_trans [trans]: | 
| 122 | "U \<unlhd> V \<Longrightarrow> V \<unlhd> W \<Longrightarrow> U \<unlhd> W" | |
| 10687 | 123 | proof | 
| 13515 | 124 | assume uv: "U \<unlhd> V" and vw: "V \<unlhd> W" | 
| 125 |   from uv show "U \<noteq> {}" by (rule subspace.non_empty)
 | |
| 126 | show "U \<subseteq> W" | |
| 127 | proof - | |
| 128 | from uv have "U \<subseteq> V" by (rule subspace.subset) | |
| 129 | also from vw have "V \<subseteq> W" by (rule subspace.subset) | |
| 130 | finally show ?thesis . | |
| 9035 | 131 | qed | 
| 13515 | 132 | fix x y assume x: "x \<in> U" and y: "y \<in> U" | 
| 133 | from uv and x y show "x + y \<in> U" by (rule subspace.add_closed) | |
| 134 | from uv and x show "\<And>a. a \<cdot> x \<in> U" by (rule subspace.mult_closed) | |
| 9035 | 135 | qed | 
| 7535 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 wenzelm parents: diff
changeset | 136 | |
| 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 wenzelm parents: diff
changeset | 137 | |
| 9035 | 138 | subsection {* Linear closure *}
 | 
| 7808 | 139 | |
| 10687 | 140 | text {*
 | 
| 141 |   The \emph{linear closure} of a vector @{text x} is the set of all
 | |
| 142 |   scalar multiples of @{text x}.
 | |
| 143 | *} | |
| 7535 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 wenzelm parents: diff
changeset | 144 | |
| 19736 | 145 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 146 |   lin :: "('a::{minus, plus, zero}) \<Rightarrow> 'a set" where
 | 
| 19736 | 147 |   "lin x = {a \<cdot> x | a. True}"
 | 
| 7535 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 wenzelm parents: diff
changeset | 148 | |
| 13515 | 149 | lemma linI [intro]: "y = a \<cdot> x \<Longrightarrow> y \<in> lin x" | 
| 27612 | 150 | unfolding lin_def by blast | 
| 7535 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 wenzelm parents: diff
changeset | 151 | |
| 13515 | 152 | lemma linI' [iff]: "a \<cdot> x \<in> lin x" | 
| 27612 | 153 | unfolding lin_def by blast | 
| 13515 | 154 | |
| 27612 | 155 | lemma linE [elim]: "x \<in> lin v \<Longrightarrow> (\<And>a::real. x = a \<cdot> v \<Longrightarrow> C) \<Longrightarrow> C" | 
| 156 | unfolding lin_def by blast | |
| 13515 | 157 | |
| 7656 | 158 | |
| 9035 | 159 | text {* Every vector is contained in its linear closure. *}
 | 
| 7917 | 160 | |
| 13515 | 161 | lemma (in vectorspace) x_lin_x [iff]: "x \<in> V \<Longrightarrow> x \<in> lin x" | 
| 162 | proof - | |
| 163 | assume "x \<in> V" | |
| 27612 | 164 | then have "x = 1 \<cdot> x" by simp | 
| 13515 | 165 | also have "\<dots> \<in> lin x" .. | 
| 166 | finally show ?thesis . | |
| 167 | qed | |
| 168 | ||
| 169 | lemma (in vectorspace) "0_lin_x" [iff]: "x \<in> V \<Longrightarrow> 0 \<in> lin x" | |
| 170 | proof | |
| 171 | assume "x \<in> V" | |
| 27612 | 172 | then show "0 = 0 \<cdot> x" by simp | 
| 13515 | 173 | qed | 
| 7535 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 wenzelm parents: diff
changeset | 174 | |
| 9035 | 175 | text {* Any linear closure is a subspace. *}
 | 
| 7917 | 176 | |
| 13515 | 177 | lemma (in vectorspace) lin_subspace [intro]: | 
| 178 | "x \<in> V \<Longrightarrow> lin x \<unlhd> V" | |
| 9035 | 179 | proof | 
| 13515 | 180 | assume x: "x \<in> V" | 
| 27612 | 181 |   then show "lin x \<noteq> {}" by (auto simp add: x_lin_x)
 | 
| 10687 | 182 | show "lin x \<subseteq> V" | 
| 13515 | 183 | proof | 
| 184 | fix x' assume "x' \<in> lin x" | |
| 185 | then obtain a where "x' = a \<cdot> x" .. | |
| 186 | with x show "x' \<in> V" by simp | |
| 9035 | 187 | qed | 
| 13515 | 188 | fix x' x'' assume x': "x' \<in> lin x" and x'': "x'' \<in> lin x" | 
| 189 | show "x' + x'' \<in> lin x" | |
| 190 | proof - | |
| 191 | from x' obtain a' where "x' = a' \<cdot> x" .. | |
| 192 | moreover from x'' obtain a'' where "x'' = a'' \<cdot> x" .. | |
| 193 | ultimately have "x' + x'' = (a' + a'') \<cdot> x" | |
| 194 | using x by (simp add: distrib) | |
| 195 | also have "\<dots> \<in> lin x" .. | |
| 196 | finally show ?thesis . | |
| 9035 | 197 | qed | 
| 13515 | 198 | fix a :: real | 
| 199 | show "a \<cdot> x' \<in> lin x" | |
| 200 | proof - | |
| 201 | from x' obtain a' where "x' = a' \<cdot> x" .. | |
| 202 | with x have "a \<cdot> x' = (a * a') \<cdot> x" by (simp add: mult_assoc) | |
| 203 | also have "\<dots> \<in> lin x" .. | |
| 204 | finally show ?thesis . | |
| 10687 | 205 | qed | 
| 9035 | 206 | qed | 
| 7535 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 wenzelm parents: diff
changeset | 207 | |
| 13515 | 208 | |
| 9035 | 209 | text {* Any linear closure is a vector space. *}
 | 
| 7917 | 210 | |
| 13515 | 211 | lemma (in vectorspace) lin_vectorspace [intro]: | 
| 23378 | 212 | assumes "x \<in> V" | 
| 213 | shows "vectorspace (lin x)" | |
| 214 | proof - | |
| 215 | from `x \<in> V` have "subspace (lin x) V" | |
| 216 | by (rule lin_subspace) | |
| 26199 | 217 | from this and vectorspace_axioms show ?thesis | 
| 23378 | 218 | by (rule subspace.vectorspace) | 
| 219 | qed | |
| 7808 | 220 | |
| 221 | ||
| 9035 | 222 | subsection {* Sum of two vectorspaces *}
 | 
| 7808 | 223 | |
| 10687 | 224 | text {*
 | 
| 225 |   The \emph{sum} of two vectorspaces @{text U} and @{text V} is the
 | |
| 226 |   set of all sums of elements from @{text U} and @{text V}.
 | |
| 227 | *} | |
| 7535 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 wenzelm parents: diff
changeset | 228 | |
| 27612 | 229 | instantiation "fun" :: (type, type) plus | 
| 230 | begin | |
| 7917 | 231 | |
| 27612 | 232 | definition | 
| 233 |   sum_def: "plus_fun U V = {u + v | u v. u \<in> U \<and> v \<in> V}"  (* FIXME not fully general!? *)
 | |
| 234 | ||
| 235 | instance .. | |
| 236 | ||
| 237 | end | |
| 7917 | 238 | |
| 13515 | 239 | lemma sumE [elim]: | 
| 240 | "x \<in> U + V \<Longrightarrow> (\<And>u v. x = u + v \<Longrightarrow> u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> C) \<Longrightarrow> C" | |
| 27612 | 241 | unfolding sum_def by blast | 
| 7535 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 wenzelm parents: diff
changeset | 242 | |
| 13515 | 243 | lemma sumI [intro]: | 
| 244 | "u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> x = u + v \<Longrightarrow> x \<in> U + V" | |
| 27612 | 245 | unfolding sum_def by blast | 
| 7566 | 246 | |
| 13515 | 247 | lemma sumI' [intro]: | 
| 248 | "u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> u + v \<in> U + V" | |
| 27612 | 249 | unfolding sum_def by blast | 
| 7917 | 250 | |
| 10687 | 251 | text {* @{text U} is a subspace of @{text "U + V"}. *}
 | 
| 7535 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 wenzelm parents: diff
changeset | 252 | |
| 13515 | 253 | lemma subspace_sum1 [iff]: | 
| 27611 | 254 | assumes "vectorspace U" "vectorspace V" | 
| 13515 | 255 | shows "U \<unlhd> U + V" | 
| 27611 | 256 | proof - | 
| 29234 | 257 | interpret vectorspace U by fact | 
| 258 | interpret vectorspace V by fact | |
| 27612 | 259 | show ?thesis | 
| 260 | proof | |
| 27611 | 261 |     show "U \<noteq> {}" ..
 | 
| 262 | show "U \<subseteq> U + V" | |
| 263 | proof | |
| 264 | fix x assume x: "x \<in> U" | |
| 265 | moreover have "0 \<in> V" .. | |
| 266 | ultimately have "x + 0 \<in> U + V" .. | |
| 267 | with x show "x \<in> U + V" by simp | |
| 268 | qed | |
| 269 | fix x y assume x: "x \<in> U" and "y \<in> U" | |
| 27612 | 270 | then show "x + y \<in> U" by simp | 
| 27611 | 271 | from x show "\<And>a. a \<cdot> x \<in> U" by simp | 
| 9035 | 272 | qed | 
| 273 | qed | |
| 7535 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 wenzelm parents: diff
changeset | 274 | |
| 13515 | 275 | text {* The sum of two subspaces is again a subspace. *}
 | 
| 7917 | 276 | |
| 13515 | 277 | lemma sum_subspace [intro?]: | 
| 27611 | 278 | assumes "subspace U E" "vectorspace E" "subspace V E" | 
| 13515 | 279 | shows "U + V \<unlhd> E" | 
| 27611 | 280 | proof - | 
| 29234 | 281 | interpret subspace U E by fact | 
| 282 | interpret vectorspace E by fact | |
| 283 | interpret subspace V E by fact | |
| 27612 | 284 | show ?thesis | 
| 285 | proof | |
| 27611 | 286 | have "0 \<in> U + V" | 
| 287 | proof | |
| 288 | show "0 \<in> U" using `vectorspace E` .. | |
| 289 | show "0 \<in> V" using `vectorspace E` .. | |
| 290 | show "(0::'a) = 0 + 0" by simp | |
| 291 | qed | |
| 27612 | 292 |     then show "U + V \<noteq> {}" by blast
 | 
| 27611 | 293 | show "U + V \<subseteq> E" | 
| 294 | proof | |
| 295 | fix x assume "x \<in> U + V" | |
| 296 | then obtain u v where "x = u + v" and | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31795diff
changeset | 297 | "u \<in> U" and "v \<in> V" .. | 
| 27611 | 298 | then show "x \<in> E" by simp | 
| 299 | qed | |
| 300 | fix x y assume x: "x \<in> U + V" and y: "y \<in> U + V" | |
| 301 | show "x + y \<in> U + V" | |
| 302 | proof - | |
| 303 | from x obtain ux vx where "x = ux + vx" and "ux \<in> U" and "vx \<in> V" .. | |
| 304 | moreover | |
| 305 | from y obtain uy vy where "y = uy + vy" and "uy \<in> U" and "vy \<in> V" .. | |
| 306 | ultimately | |
| 307 | have "ux + uy \<in> U" | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31795diff
changeset | 308 | and "vx + vy \<in> V" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31795diff
changeset | 309 | and "x + y = (ux + uy) + (vx + vy)" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31795diff
changeset | 310 | using x y by (simp_all add: add_ac) | 
| 27612 | 311 | then show ?thesis .. | 
| 27611 | 312 | qed | 
| 313 | fix a show "a \<cdot> x \<in> U + V" | |
| 314 | proof - | |
| 315 | from x obtain u v where "x = u + v" and "u \<in> U" and "v \<in> V" .. | |
| 27612 | 316 | then have "a \<cdot> u \<in> U" and "a \<cdot> v \<in> V" | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31795diff
changeset | 317 | and "a \<cdot> x = (a \<cdot> u) + (a \<cdot> v)" by (simp_all add: distrib) | 
| 27612 | 318 | then show ?thesis .. | 
| 27611 | 319 | qed | 
| 9035 | 320 | qed | 
| 321 | qed | |
| 7535 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 wenzelm parents: diff
changeset | 322 | |
| 9035 | 323 | text{* The sum of two subspaces is a vectorspace. *}
 | 
| 7917 | 324 | |
| 13515 | 325 | lemma sum_vs [intro?]: | 
| 326 | "U \<unlhd> E \<Longrightarrow> V \<unlhd> E \<Longrightarrow> vectorspace E \<Longrightarrow> vectorspace (U + V)" | |
| 13547 | 327 | by (rule subspace.vectorspace) (rule sum_subspace) | 
| 7535 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 wenzelm parents: diff
changeset | 328 | |
| 7808 | 329 | |
| 9035 | 330 | subsection {* Direct sums *}
 | 
| 7808 | 331 | |
| 10687 | 332 | text {*
 | 
| 333 |   The sum of @{text U} and @{text V} is called \emph{direct}, iff the
 | |
| 334 |   zero element is the only common element of @{text U} and @{text
 | |
| 335 |   V}. For every element @{text x} of the direct sum of @{text U} and
 | |
| 336 |   @{text V} the decomposition in @{text "x = u + v"} with
 | |
| 337 |   @{text "u \<in> U"} and @{text "v \<in> V"} is unique.
 | |
| 338 | *} | |
| 7808 | 339 | |
| 10687 | 340 | lemma decomp: | 
| 27611 | 341 | assumes "vectorspace E" "subspace U E" "subspace V E" | 
| 13515 | 342 |   assumes direct: "U \<inter> V = {0}"
 | 
| 343 | and u1: "u1 \<in> U" and u2: "u2 \<in> U" | |
| 344 | and v1: "v1 \<in> V" and v2: "v2 \<in> V" | |
| 345 | and sum: "u1 + v1 = u2 + v2" | |
| 346 | shows "u1 = u2 \<and> v1 = v2" | |
| 27611 | 347 | proof - | 
| 29234 | 348 | interpret vectorspace E by fact | 
| 349 | interpret subspace U E by fact | |
| 350 | interpret subspace V E by fact | |
| 27612 | 351 | show ?thesis | 
| 352 | proof | |
| 27611 | 353 | have U: "vectorspace U" (* FIXME: use interpret *) | 
| 354 | using `subspace U E` `vectorspace E` by (rule subspace.vectorspace) | |
| 355 | have V: "vectorspace V" | |
| 356 | using `subspace V E` `vectorspace E` by (rule subspace.vectorspace) | |
| 357 | from u1 u2 v1 v2 and sum have eq: "u1 - u2 = v2 - v1" | |
| 358 | by (simp add: add_diff_swap) | |
| 359 | from u1 u2 have u: "u1 - u2 \<in> U" | |
| 360 | by (rule vectorspace.diff_closed [OF U]) | |
| 361 | with eq have v': "v2 - v1 \<in> U" by (simp only:) | |
| 362 | from v2 v1 have v: "v2 - v1 \<in> V" | |
| 363 | by (rule vectorspace.diff_closed [OF V]) | |
| 364 | with eq have u': " u1 - u2 \<in> V" by (simp only:) | |
| 365 | ||
| 366 | show "u1 = u2" | |
| 367 | proof (rule add_minus_eq) | |
| 368 | from u1 show "u1 \<in> E" .. | |
| 369 | from u2 show "u2 \<in> E" .. | |
| 370 | from u u' and direct show "u1 - u2 = 0" by blast | |
| 371 | qed | |
| 372 | show "v1 = v2" | |
| 373 | proof (rule add_minus_eq [symmetric]) | |
| 374 | from v1 show "v1 \<in> E" .. | |
| 375 | from v2 show "v2 \<in> E" .. | |
| 376 | from v v' and direct show "v2 - v1 = 0" by blast | |
| 377 | qed | |
| 9035 | 378 | qed | 
| 379 | qed | |
| 7656 | 380 | |
| 10687 | 381 | text {*
 | 
| 382 | An application of the previous lemma will be used in the proof of | |
| 383 |   the Hahn-Banach Theorem (see page \pageref{decomp-H-use}): for any
 | |
| 384 |   element @{text "y + a \<cdot> x\<^sub>0"} of the direct sum of a
 | |
| 385 |   vectorspace @{text H} and the linear closure of @{text "x\<^sub>0"}
 | |
| 386 |   the components @{text "y \<in> H"} and @{text a} are uniquely
 | |
| 387 | determined. | |
| 388 | *} | |
| 7917 | 389 | |
| 10687 | 390 | lemma decomp_H': | 
| 27611 | 391 | assumes "vectorspace E" "subspace H E" | 
| 13515 | 392 | assumes y1: "y1 \<in> H" and y2: "y2 \<in> H" | 
| 393 | and x': "x' \<notin> H" "x' \<in> E" "x' \<noteq> 0" | |
| 394 | and eq: "y1 + a1 \<cdot> x' = y2 + a2 \<cdot> x'" | |
| 395 | shows "y1 = y2 \<and> a1 = a2" | |
| 27611 | 396 | proof - | 
| 29234 | 397 | interpret vectorspace E by fact | 
| 398 | interpret subspace H E by fact | |
| 27612 | 399 | show ?thesis | 
| 400 | proof | |
| 27611 | 401 | have c: "y1 = y2 \<and> a1 \<cdot> x' = a2 \<cdot> x'" | 
| 402 | proof (rule decomp) | |
| 403 | show "a1 \<cdot> x' \<in> lin x'" .. | |
| 404 | show "a2 \<cdot> x' \<in> lin x'" .. | |
| 405 |       show "H \<inter> lin x' = {0}"
 | |
| 13515 | 406 | proof | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31795diff
changeset | 407 |         show "H \<inter> lin x' \<subseteq> {0}"
 | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31795diff
changeset | 408 | proof | 
| 27611 | 409 | fix x assume x: "x \<in> H \<inter> lin x'" | 
| 410 | then obtain a where xx': "x = a \<cdot> x'" | |
| 411 | by blast | |
| 412 | have "x = 0" | |
| 413 | proof cases | |
| 414 | assume "a = 0" | |
| 415 | with xx' and x' show ?thesis by simp | |
| 416 | next | |
| 417 | assume a: "a \<noteq> 0" | |
| 418 | from x have "x \<in> H" .. | |
| 419 | with xx' have "inverse a \<cdot> a \<cdot> x' \<in> H" by simp | |
| 420 | with a and x' have "x' \<in> H" by (simp add: mult_assoc2) | |
| 421 | with `x' \<notin> H` show ?thesis by contradiction | |
| 422 | qed | |
| 27612 | 423 |           then show "x \<in> {0}" ..
 | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31795diff
changeset | 424 | qed | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31795diff
changeset | 425 |         show "{0} \<subseteq> H \<inter> lin x'"
 | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31795diff
changeset | 426 | proof - | 
| 27611 | 427 | have "0 \<in> H" using `vectorspace E` .. | 
| 428 | moreover have "0 \<in> lin x'" using `x' \<in> E` .. | |
| 429 | ultimately show ?thesis by blast | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31795diff
changeset | 430 | qed | 
| 9035 | 431 | qed | 
| 27611 | 432 | show "lin x' \<unlhd> E" using `x' \<in> E` .. | 
| 433 | qed (rule `vectorspace E`, rule `subspace H E`, rule y1, rule y2, rule eq) | |
| 27612 | 434 | then show "y1 = y2" .. | 
| 27611 | 435 | from c have "a1 \<cdot> x' = a2 \<cdot> x'" .. | 
| 436 | with x' show "a1 = a2" by (simp add: mult_right_cancel) | |
| 437 | qed | |
| 9035 | 438 | qed | 
| 7535 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 wenzelm parents: diff
changeset | 439 | |
| 10687 | 440 | text {*
 | 
| 441 |   Since for any element @{text "y + a \<cdot> x'"} of the direct sum of a
 | |
| 442 |   vectorspace @{text H} and the linear closure of @{text x'} the
 | |
| 443 |   components @{text "y \<in> H"} and @{text a} are unique, it follows from
 | |
| 444 |   @{text "y \<in> H"} that @{text "a = 0"}.
 | |
| 445 | *} | |
| 7917 | 446 | |
| 10687 | 447 | lemma decomp_H'_H: | 
| 27611 | 448 | assumes "vectorspace E" "subspace H E" | 
| 13515 | 449 | assumes t: "t \<in> H" | 
| 450 | and x': "x' \<notin> H" "x' \<in> E" "x' \<noteq> 0" | |
| 451 | shows "(SOME (y, a). t = y + a \<cdot> x' \<and> y \<in> H) = (t, 0)" | |
| 27611 | 452 | proof - | 
| 29234 | 453 | interpret vectorspace E by fact | 
| 454 | interpret subspace H E by fact | |
| 27612 | 455 | show ?thesis | 
| 456 | proof (rule, simp_all only: split_paired_all split_conv) | |
| 27611 | 457 | from t x' show "t = t + 0 \<cdot> x' \<and> t \<in> H" by simp | 
| 458 | fix y and a assume ya: "t = y + a \<cdot> x' \<and> y \<in> H" | |
| 459 | have "y = t \<and> a = 0" | |
| 460 | proof (rule decomp_H') | |
| 461 | from ya x' show "y + a \<cdot> x' = t + 0 \<cdot> x'" by simp | |
| 462 | from ya show "y \<in> H" .. | |
| 463 | qed (rule `vectorspace E`, rule `subspace H E`, rule t, (rule x')+) | |
| 464 | with t x' show "(y, a) = (y + a \<cdot> x', 0)" by simp | |
| 465 | qed | |
| 13515 | 466 | qed | 
| 7535 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 wenzelm parents: diff
changeset | 467 | |
| 10687 | 468 | text {*
 | 
| 469 |   The components @{text "y \<in> H"} and @{text a} in @{text "y + a \<cdot> x'"}
 | |
| 470 |   are unique, so the function @{text h'} defined by
 | |
| 471 |   @{text "h' (y + a \<cdot> x') = h y + a \<cdot> \<xi>"} is definite.
 | |
| 472 | *} | |
| 7917 | 473 | |
| 9374 | 474 | lemma h'_definite: | 
| 27611 | 475 | fixes H | 
| 13515 | 476 | assumes h'_def: | 
| 477 | "h' \<equiv> (\<lambda>x. let (y, a) = SOME (y, a). (x = y + a \<cdot> x' \<and> y \<in> H) | |
| 478 | in (h y) + a * xi)" | |
| 479 | and x: "x = y + a \<cdot> x'" | |
| 27611 | 480 | assumes "vectorspace E" "subspace H E" | 
| 13515 | 481 | assumes y: "y \<in> H" | 
| 482 | and x': "x' \<notin> H" "x' \<in> E" "x' \<noteq> 0" | |
| 483 | shows "h' x = h y + a * xi" | |
| 10687 | 484 | proof - | 
| 29234 | 485 | interpret vectorspace E by fact | 
| 486 | interpret subspace H E by fact | |
| 13515 | 487 | from x y x' have "x \<in> H + lin x'" by auto | 
| 488 | have "\<exists>!p. (\<lambda>(y, a). x = y + a \<cdot> x' \<and> y \<in> H) p" (is "\<exists>!p. ?P p") | |
| 18689 
a50587cd8414
prefer ex1I over ex_ex1I in single-step reasoning;
 wenzelm parents: 
16417diff
changeset | 489 | proof (rule ex_ex1I) | 
| 13515 | 490 | from x y show "\<exists>p. ?P p" by blast | 
| 491 | fix p q assume p: "?P p" and q: "?P q" | |
| 492 | show "p = q" | |
| 9035 | 493 | proof - | 
| 13515 | 494 | from p have xp: "x = fst p + snd p \<cdot> x' \<and> fst p \<in> H" | 
| 495 | by (cases p) simp | |
| 496 | from q have xq: "x = fst q + snd q \<cdot> x' \<and> fst q \<in> H" | |
| 497 | by (cases q) simp | |
| 498 | have "fst p = fst q \<and> snd p = snd q" | |
| 499 | proof (rule decomp_H') | |
| 500 | from xp show "fst p \<in> H" .. | |
| 501 | from xq show "fst q \<in> H" .. | |
| 502 | from xp and xq show "fst p + snd p \<cdot> x' = fst q + snd q \<cdot> x'" | |
| 503 | by simp | |
| 23378 | 504 | qed (rule `vectorspace E`, rule `subspace H E`, (rule x')+) | 
| 27612 | 505 | then show ?thesis by (cases p, cases q) simp | 
| 9035 | 506 | qed | 
| 507 | qed | |
| 27612 | 508 | then have eq: "(SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H) = (y, a)" | 
| 13515 | 509 | by (rule some1_equality) (simp add: x y) | 
| 510 | with h'_def show "h' x = h y + a * xi" by (simp add: Let_def) | |
| 9035 | 511 | qed | 
| 7535 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 wenzelm parents: diff
changeset | 512 | |
| 10687 | 513 | end |