1(*  Title:      HOL/Hahn_Banach/Subspace.thy
2    Author:     Gertrud Bauer, TU Munich
3*)
4
5section \<open>Subspaces\<close>
6
7theory Subspace
8imports Vector_Space "HOL-Library.Set_Algebras"
9begin
10
11subsection \<open>Definition\<close>
12
13text \<open>
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.
16\<close>
17
18locale subspace =
19  fixes U :: "'a::{minus, plus, zero, uminus} set" and V
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"
24
25notation (symbols)
26  subspace  (infix "\<unlhd>" 50)
27
28declare vectorspace.intro [intro?] subspace.intro [intro?]
29
30lemma subspace_subset [elim]: "U \<unlhd> V \<Longrightarrow> U \<subseteq> V"
31  by (rule subspace.subset)
32
33lemma (in subspace) subsetD [iff]: "x \<in> U \<Longrightarrow> x \<in> V"
34  using subset by blast
35
36lemma subspaceD [elim]: "U \<unlhd> V \<Longrightarrow> x \<in> U \<Longrightarrow> x \<in> V"
37  by (rule subspace.subsetD)
38
39lemma rev_subspaceD [elim?]: "x \<in> U \<Longrightarrow> U \<unlhd> V \<Longrightarrow> x \<in> V"
40  by (rule subspace.subsetD)
41
42lemma (in subspace) diff_closed [iff]:
43  assumes "vectorspace V"
44  assumes x: "x \<in> U" and y: "y \<in> U"
45  shows "x - y \<in> U"
46proof -
47  interpret vectorspace V by fact
48  from x y show ?thesis by (simp add: diff_eq1 negate_eq1)
49qed
50
51text \<open>
52  \<^medskip>
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.
56\<close>
57
58lemma (in subspace) zero [intro]:
59  assumes "vectorspace V"
60  shows "0 \<in> U"
61proof -
62  interpret V: vectorspace V by fact
63  have "U \<noteq> {}" by (rule non_empty)
64  then obtain x where x: "x \<in> U" by blast
65  then have "x \<in> V" .. then have "0 = x - x" by simp
66  also from \<open>vectorspace V\<close> x x have "\<dots> \<in> U" by (rule diff_closed)
67  finally show ?thesis .
68qed
69
70lemma (in subspace) neg_closed [iff]:
71  assumes "vectorspace V"
72  assumes x: "x \<in> U"
73  shows "- x \<in> U"
74proof -
75  interpret vectorspace V by fact
76  from x show ?thesis by (simp add: negate_eq1)
77qed
78
79text \<open>\<^medskip> Further derived laws: every subspace is a vector space.\<close>
80
81lemma (in subspace) vectorspace [iff]:
82  assumes "vectorspace V"
83  shows "vectorspace U"
84proof -
85  interpret vectorspace V by fact
86  show ?thesis
87  proof
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
104qed
105
106
107text \<open>The subspace relation is reflexive.\<close>
108
109lemma (in vectorspace) subspace_refl [intro]: "V \<unlhd> V"
110proof
111  show "V \<noteq> {}" ..
112  show "V \<subseteq> V" ..
113next
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
118qed
119
120text \<open>The subspace relation is transitive.\<close>
121
122lemma (in vectorspace) subspace_trans [trans]:
123  "U \<unlhd> V \<Longrightarrow> V \<unlhd> W \<Longrightarrow> U \<unlhd> W"
124proof
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 .
132  qed
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)
135  from uv and x show "a \<cdot> x \<in> U" for a by (rule subspace.mult_closed)
136qed
137
138
139subsection \<open>Linear closure\<close>
140
141text \<open>
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>.
144\<close>
145
146definition lin :: "('a::{minus,plus,zero}) \<Rightarrow> 'a set"
147  where "lin x = {a \<cdot> x | a. True}"
148
149lemma linI [intro]: "y = a \<cdot> x \<Longrightarrow> y \<in> lin x"
150  unfolding lin_def by blast
151
152lemma linI' [iff]: "a \<cdot> x \<in> lin x"
153  unfolding lin_def by blast
154
155lemma 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
159
160
161text \<open>Every vector is contained in its linear closure.\<close>
162
163lemma (in vectorspace) x_lin_x [iff]: "x \<in> V \<Longrightarrow> x \<in> lin x"
164proof -
165  assume "x \<in> V"
166  then have "x = 1 \<cdot> x" by simp
167  also have "\<dots> \<in> lin x" ..
168  finally show ?thesis .
169qed
170
171lemma (in vectorspace) "0_lin_x" [iff]: "x \<in> V \<Longrightarrow> 0 \<in> lin x"
172proof
173  assume "x \<in> V"
174  then show "0 = 0 \<cdot> x" by simp
175qed
176
177text \<open>Any linear closure is a subspace.\<close>
178
179lemma (in vectorspace) lin_subspace [intro]:
180  assumes x: "x \<in> V"
181  shows "lin x \<unlhd> V"
182proof
183  from x show "lin x \<noteq> {}" by auto
184next
185  show "lin x \<subseteq> V"
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
190  qed
191next
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 .
201  qed
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 .
209  qed
210qed
211
212
213text \<open>Any linear closure is a vector space.\<close>
214
215lemma (in vectorspace) lin_vectorspace [intro]:
216  assumes "x \<in> V"
217  shows "vectorspace (lin x)"
218proof -
219  from \<open>x \<in> V\<close> have "subspace (lin x) V"
220    by (rule lin_subspace)
221  from this and vectorspace_axioms show ?thesis
222    by (rule subspace.vectorspace)
223qed
224
225
226subsection \<open>Sum of two vectorspaces\<close>
227
228text \<open>
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>.
231\<close>
232
233lemma sum_def: "U + V = {u + v | u v. u \<in> U \<and> v \<in> V}"
234  unfolding set_plus_def by auto
235
236lemma sumE [elim]:
237    "x \<in> U + V \<Longrightarrow> (\<And>u v. x = u + v \<Longrightarrow> u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> C) \<Longrightarrow> C"
238  unfolding sum_def by blast
239
240lemma sumI [intro]:
241    "u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> x = u + v \<Longrightarrow> x \<in> U + V"
242  unfolding sum_def by blast
243
244lemma sumI' [intro]:
245    "u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> u + v \<in> U + V"
246  unfolding sum_def by blast
247
248text \<open>\<open>U\<close> is a subspace of \<open>U + V\<close>.\<close>
249
250lemma subspace_sum1 [iff]:
251  assumes "vectorspace U" "vectorspace V"
252  shows "U \<unlhd> U + V"
253proof -
254  interpret vectorspace U by fact
255  interpret vectorspace V by fact
256  show ?thesis
257  proof
258    show "U \<noteq> {}" ..
259    show "U \<subseteq> U + V"
260    proof
261      fix x assume x: "x \<in> U"
262      moreover have "0 \<in> V" ..
263      ultimately have "x + 0 \<in> U + V" ..
264      with x show "x \<in> U + V" by simp
265    qed
266    fix x y assume x: "x \<in> U" and "y \<in> U"
267    then show "x + y \<in> U" by simp
268    from x show "a \<cdot> x \<in> U" for a by simp
269  qed
270qed
271
272text \<open>The sum of two subspaces is again a subspace.\<close>
273
274lemma sum_subspace [intro?]:
275  assumes "subspace U E" "vectorspace E" "subspace V E"
276  shows "U + V \<unlhd> E"
277proof -
278  interpret subspace U E by fact
279  interpret vectorspace E by fact
280  interpret subspace V E by fact
281  show ?thesis
282  proof
283    have "0 \<in> U + V"
284    proof
285      show "0 \<in> U" using \<open>vectorspace E\<close> ..
286      show "0 \<in> V" using \<open>vectorspace E\<close> ..
287      show "(0::'a) = 0 + 0" by simp
288    qed
289    then show "U + V \<noteq> {}" by blast
290    show "U + V \<subseteq> E"
291    proof
292      fix x assume "x \<in> U + V"
293      then obtain u v where "x = u + v" and
294        "u \<in> U" and "v \<in> V" ..
295      then show "x \<in> E" by simp
296    qed
297  next
298    fix x y assume x: "x \<in> U + V" and y: "y \<in> U + V"
299    show "x + y \<in> U + V"
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"
306        and "vx + vy \<in> V"
307        and "x + y = (ux + uy) + (vx + vy)"
308        using x y by (simp_all add: add_ac)
309      then show ?thesis ..
310    qed
311    fix a show "a \<cdot> x \<in> U + V"
312    proof -
313      from x obtain u v where "x = u + v" and "u \<in> U" and "v \<in> V" ..
314      then have "a \<cdot> u \<in> U" and "a \<cdot> v \<in> V"
315        and "a \<cdot> x = (a \<cdot> u) + (a \<cdot> v)" by (simp_all add: distrib)
316      then show ?thesis ..
317    qed
318  qed
319qed
320
321text \<open>The sum of two subspaces is a vectorspace.\<close>
322
323lemma sum_vs [intro?]:
324    "U \<unlhd> E \<Longrightarrow> V \<unlhd> E \<Longrightarrow> vectorspace E \<Longrightarrow> vectorspace (U + V)"
325  by (rule subspace.vectorspace) (rule sum_subspace)
326
327
328subsection \<open>Direct sums\<close>
329
330text \<open>
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.
335\<close>
336
337lemma decomp:
338  assumes "vectorspace E" "subspace U E" "subspace V E"
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"
344proof -
345  interpret vectorspace E by fact
346  interpret subspace U E by fact
347  interpret subspace V E by fact
348  show ?thesis
349  proof
350    have U: "vectorspace U"  (* FIXME: use interpret *)
351      using \<open>subspace U E\<close> \<open>vectorspace E\<close> by (rule subspace.vectorspace)
352    have V: "vectorspace V"
353      using \<open>subspace V E\<close> \<open>vectorspace E\<close> by (rule subspace.vectorspace)
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
375  qed
376qed
377
378text \<open>
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.
383\<close>
384
385lemma decomp_H':
386  assumes "vectorspace E" "subspace H E"
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"
391proof -
392  interpret vectorspace E by fact
393  interpret subspace H E by fact
394  show ?thesis
395  proof
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}"
401      proof
402        show "H \<inter> lin x' \<subseteq> {0}"
403        proof
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)
416            with \<open>x' \<notin> H\<close> show ?thesis by contradiction
417          qed
418          then show "x \<in> {0}" ..
419        qed
420        show "{0} \<subseteq> H \<inter> lin x'"
421        proof -
422          have "0 \<in> H" using \<open>vectorspace E\<close> ..
423          moreover have "0 \<in> lin x'" using \<open>x' \<in> E\<close> ..
424          ultimately show ?thesis by blast
425        qed
426      qed
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)
429    then show "y1 = y2" ..
430    from c have "a1 \<cdot> x' = a2 \<cdot> x'" ..
431    with x' show "a1 = a2" by (simp add: mult_right_cancel)
432  qed
433qed
434
435text \<open>
436  Since for any element \<open>y + a \<cdot> x'\<close> of the direct sum of a vectorspace \<open>H\<close>
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>.
439\<close>
440
441lemma decomp_H'_H:
442  assumes "vectorspace E" "subspace H E"
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)"
446proof -
447  interpret vectorspace E by fact
448  interpret subspace H E by fact
449  show ?thesis
450  proof (rule, simp_all only: split_paired_all split_conv)
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" ..
457    qed (rule \<open>vectorspace E\<close>, rule \<open>subspace H E\<close>, rule t, (rule x')+)
458    with t x' show "(y, a) = (y + a \<cdot> x', 0)" by simp
459  qed
460qed
461
462text \<open>
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.
465\<close>
466
467lemma h'_definite:
468  fixes H
469  assumes h'_def:
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)"
473    and x: "x = y + a \<cdot> x'"
474  assumes "vectorspace E" "subspace H E"
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"
478proof -
479  interpret vectorspace E by fact
480  interpret subspace H E by fact
481  from x y x' have "x \<in> H + lin x'" by auto
482  have "\<exists>!(y, a). x = y + a \<cdot> x' \<and> y \<in> H" (is "\<exists>!p. ?P p")
483  proof (rule ex_ex1I)
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"
487    proof -
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
498      qed (rule \<open>vectorspace E\<close>, rule \<open>subspace H E\<close>, (rule x')+)
499      then show ?thesis by (cases p, cases q) simp
500    qed
501  qed
502  then have eq: "(SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H) = (y, a)"
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)
505qed
506
507end
508