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