Up to index of Isabelle/HOL/HOL-Complex/HahnBanach
theory Subspace(* Title: HOL/Real/HahnBanach/Subspace.thy
ID: $Id: Subspace.thy,v 1.32 2005/06/17 14:13:10 haftmann Exp $
Author: Gertrud Bauer, TU Munich
*)
header {* Subspaces *}
theory Subspace imports VectorSpace begin
subsection {* Definition *}
text {*
A non-empty subset @{text U} of a vector space @{text V} is a
\emph{subspace} of @{text V}, iff @{text U} is closed under addition
and scalar multiplication.
*}
locale subspace = var U + var V +
assumes non_empty [iff, intro]: "U ≠ {}"
and subset [iff]: "U ⊆ V"
and add_closed [iff]: "x ∈ U ==> y ∈ U ==> x + y ∈ U"
and mult_closed [iff]: "x ∈ U ==> a · x ∈ U"
declare vectorspace.intro [intro?] subspace.intro [intro?]
syntax (symbols)
subspace :: "'a set => 'a set => bool" (infix "\<unlhd>" 50)
lemma subspace_subset [elim]: "U \<unlhd> V ==> U ⊆ V"
by (rule subspace.subset)
lemma (in subspace) subsetD [iff]: "x ∈ U ==> x ∈ V"
using subset by blast
lemma subspaceD [elim]: "U \<unlhd> V ==> x ∈ U ==> x ∈ V"
by (rule subspace.subsetD)
lemma rev_subspaceD [elim?]: "x ∈ U ==> U \<unlhd> V ==> x ∈ V"
by (rule subspace.subsetD)
lemma (in subspace) diff_closed [iff]:
includes vectorspace
shows "x ∈ U ==> y ∈ U ==> x - y ∈ U"
by (simp add: diff_eq1 negate_eq1)
text {*
\medskip Similar as for linear spaces, the existence of the zero
element in every subspace follows from the non-emptiness of the
carrier set and by vector space laws.
*}
lemma (in subspace) zero [intro]:
includes vectorspace
shows "0 ∈ U"
proof -
have "U ≠ {}" by (rule U_V.non_empty)
then obtain x where x: "x ∈ U" by blast
hence "x ∈ V" .. hence "0 = x - x" by simp
also have "... ∈ U" by (rule U_V.diff_closed)
finally show ?thesis .
qed
lemma (in subspace) neg_closed [iff]:
includes vectorspace
shows "x ∈ U ==> - x ∈ U"
by (simp add: negate_eq1)
text {* \medskip Further derived laws: every subspace is a vector space. *}
lemma (in subspace) vectorspace [iff]:
includes vectorspace
shows "vectorspace U"
proof
show "U ≠ {}" ..
fix x y z assume x: "x ∈ U" and y: "y ∈ U" and z: "z ∈ U"
fix a b :: real
from x y show "x + y ∈ U" by simp
from x show "a · x ∈ U" by simp
from x y z show "(x + y) + z = x + (y + z)" by (simp add: add_ac)
from x y show "x + y = y + x" by (simp add: add_ac)
from x show "x - x = 0" by simp
from x show "0 + x = x" by simp
from x y show "a · (x + y) = a · x + a · y" by (simp add: distrib)
from x show "(a + b) · x = a · x + b · x" by (simp add: distrib)
from x show "(a * b) · x = a · b · x" by (simp add: mult_assoc)
from x show "1 · x = x" by simp
from x show "- x = - 1 · x" by (simp add: negate_eq1)
from x y show "x - y = x + - y" by (simp add: diff_eq1)
qed
text {* The subspace relation is reflexive. *}
lemma (in vectorspace) subspace_refl [intro]: "V \<unlhd> V"
proof
show "V ≠ {}" ..
show "V ⊆ V" ..
fix x y assume x: "x ∈ V" and y: "y ∈ V"
fix a :: real
from x y show "x + y ∈ V" by simp
from x show "a · x ∈ V" by simp
qed
text {* The subspace relation is transitive. *}
lemma (in vectorspace) subspace_trans [trans]:
"U \<unlhd> V ==> V \<unlhd> W ==> U \<unlhd> W"
proof
assume uv: "U \<unlhd> V" and vw: "V \<unlhd> W"
from uv show "U ≠ {}" by (rule subspace.non_empty)
show "U ⊆ W"
proof -
from uv have "U ⊆ V" by (rule subspace.subset)
also from vw have "V ⊆ W" by (rule subspace.subset)
finally show ?thesis .
qed
fix x y assume x: "x ∈ U" and y: "y ∈ U"
from uv and x y show "x + y ∈ U" by (rule subspace.add_closed)
from uv and x show "!!a. a · x ∈ U" by (rule subspace.mult_closed)
qed
subsection {* Linear closure *}
text {*
The \emph{linear closure} of a vector @{text x} is the set of all
scalar multiples of @{text x}.
*}
constdefs
lin :: "('a::{minus, plus, zero}) => 'a set"
"lin x ≡ {a · x | a. True}"
lemma linI [intro]: "y = a · x ==> y ∈ lin x"
by (unfold lin_def) blast
lemma linI' [iff]: "a · x ∈ lin x"
by (unfold lin_def) blast
lemma linE [elim]:
"x ∈ lin v ==> (!!a::real. x = a · v ==> C) ==> C"
by (unfold lin_def) blast
text {* Every vector is contained in its linear closure. *}
lemma (in vectorspace) x_lin_x [iff]: "x ∈ V ==> x ∈ lin x"
proof -
assume "x ∈ V"
hence "x = 1 · x" by simp
also have "… ∈ lin x" ..
finally show ?thesis .
qed
lemma (in vectorspace) "0_lin_x" [iff]: "x ∈ V ==> 0 ∈ lin x"
proof
assume "x ∈ V"
thus "0 = 0 · x" by simp
qed
text {* Any linear closure is a subspace. *}
lemma (in vectorspace) lin_subspace [intro]:
"x ∈ V ==> lin x \<unlhd> V"
proof
assume x: "x ∈ V"
thus "lin x ≠ {}" by (auto simp add: x_lin_x)
show "lin x ⊆ V"
proof
fix x' assume "x' ∈ lin x"
then obtain a where "x' = a · x" ..
with x show "x' ∈ V" by simp
qed
fix x' x'' assume x': "x' ∈ lin x" and x'': "x'' ∈ lin x"
show "x' + x'' ∈ lin x"
proof -
from x' obtain a' where "x' = a' · x" ..
moreover from x'' obtain a'' where "x'' = a'' · x" ..
ultimately have "x' + x'' = (a' + a'') · x"
using x by (simp add: distrib)
also have "… ∈ lin x" ..
finally show ?thesis .
qed
fix a :: real
show "a · x' ∈ lin x"
proof -
from x' obtain a' where "x' = a' · x" ..
with x have "a · x' = (a * a') · x" by (simp add: mult_assoc)
also have "… ∈ lin x" ..
finally show ?thesis .
qed
qed
text {* Any linear closure is a vector space. *}
lemma (in vectorspace) lin_vectorspace [intro]:
"x ∈ V ==> vectorspace (lin x)"
by (rule subspace.vectorspace) (rule lin_subspace)
subsection {* Sum of two vectorspaces *}
text {*
The \emph{sum} of two vectorspaces @{text U} and @{text V} is the
set of all sums of elements from @{text U} and @{text V}.
*}
instance set :: (plus) plus ..
defs (overloaded)
sum_def: "U + V ≡ {u + v | u v. u ∈ U ∧ v ∈ V}"
lemma sumE [elim]:
"x ∈ U + V ==> (!!u v. x = u + v ==> u ∈ U ==> v ∈ V ==> C) ==> C"
by (unfold sum_def) blast
lemma sumI [intro]:
"u ∈ U ==> v ∈ V ==> x = u + v ==> x ∈ U + V"
by (unfold sum_def) blast
lemma sumI' [intro]:
"u ∈ U ==> v ∈ V ==> u + v ∈ U + V"
by (unfold sum_def) blast
text {* @{text U} is a subspace of @{text "U + V"}. *}
lemma subspace_sum1 [iff]:
includes vectorspace U + vectorspace V
shows "U \<unlhd> U + V"
proof
show "U ≠ {}" ..
show "U ⊆ U + V"
proof
fix x assume x: "x ∈ U"
moreover have "0 ∈ V" ..
ultimately have "x + 0 ∈ U + V" ..
with x show "x ∈ U + V" by simp
qed
fix x y assume x: "x ∈ U" and "y ∈ U"
thus "x + y ∈ U" by simp
from x show "!!a. a · x ∈ U" by simp
qed
text {* The sum of two subspaces is again a subspace. *}
lemma sum_subspace [intro?]:
includes subspace U E + vectorspace E + subspace V E
shows "U + V \<unlhd> E"
proof
have "0 ∈ U + V"
proof
show "0 ∈ U" ..
show "0 ∈ V" ..
show "(0::'a) = 0 + 0" by simp
qed
thus "U + V ≠ {}" by blast
show "U + V ⊆ E"
proof
fix x assume "x ∈ U + V"
then obtain u v where x: "x = u + v" and
u: "u ∈ U" and v: "v ∈ V" ..
have "U \<unlhd> E" . with u have "u ∈ E" ..
moreover have "V \<unlhd> E" . with v have "v ∈ E" ..
ultimately show "x ∈ E" using x by simp
qed
fix x y assume x: "x ∈ U + V" and y: "y ∈ U + V"
show "x + y ∈ U + V"
proof -
from x obtain ux vx where "x = ux + vx" and "ux ∈ U" and "vx ∈ V" ..
moreover
from y obtain uy vy where "y = uy + vy" and "uy ∈ U" and "vy ∈ V" ..
ultimately
have "ux + uy ∈ U"
and "vx + vy ∈ V"
and "x + y = (ux + uy) + (vx + vy)"
using x y by (simp_all add: add_ac)
thus ?thesis ..
qed
fix a show "a · x ∈ U + V"
proof -
from x obtain u v where "x = u + v" and "u ∈ U" and "v ∈ V" ..
hence "a · u ∈ U" and "a · v ∈ V"
and "a · x = (a · u) + (a · v)" by (simp_all add: distrib)
thus ?thesis ..
qed
qed
text{* The sum of two subspaces is a vectorspace. *}
lemma sum_vs [intro?]:
"U \<unlhd> E ==> V \<unlhd> E ==> vectorspace E ==> vectorspace (U + V)"
by (rule subspace.vectorspace) (rule sum_subspace)
subsection {* Direct sums *}
text {*
The sum of @{text U} and @{text V} is called \emph{direct}, iff the
zero element is the only common element of @{text U} and @{text
V}. For every element @{text x} of the direct sum of @{text U} and
@{text V} the decomposition in @{text "x = u + v"} with
@{text "u ∈ U"} and @{text "v ∈ V"} is unique.
*}
lemma decomp:
includes vectorspace E + subspace U E + subspace V E
assumes direct: "U ∩ V = {0}"
and u1: "u1 ∈ U" and u2: "u2 ∈ U"
and v1: "v1 ∈ V" and v2: "v2 ∈ V"
and sum: "u1 + v1 = u2 + v2"
shows "u1 = u2 ∧ v1 = v2"
proof
have U: "vectorspace U" by (rule subspace.vectorspace)
have V: "vectorspace V" by (rule subspace.vectorspace)
from u1 u2 v1 v2 and sum have eq: "u1 - u2 = v2 - v1"
by (simp add: add_diff_swap)
from u1 u2 have u: "u1 - u2 ∈ U"
by (rule vectorspace.diff_closed [OF U])
with eq have v': "v2 - v1 ∈ U" by (simp only:)
from v2 v1 have v: "v2 - v1 ∈ V"
by (rule vectorspace.diff_closed [OF V])
with eq have u': " u1 - u2 ∈ V" by (simp only:)
show "u1 = u2"
proof (rule add_minus_eq)
show "u1 ∈ E" ..
show "u2 ∈ E" ..
from u u' and direct show "u1 - u2 = 0" by force
qed
show "v1 = v2"
proof (rule add_minus_eq [symmetric])
show "v1 ∈ E" ..
show "v2 ∈ E" ..
from v v' and direct show "v2 - v1 = 0" by force
qed
qed
text {*
An application of the previous lemma will be used in the proof of
the Hahn-Banach Theorem (see page \pageref{decomp-H-use}): for any
element @{text "y + a · x0"} of the direct sum of a
vectorspace @{text H} and the linear closure of @{text "x0"}
the components @{text "y ∈ H"} and @{text a} are uniquely
determined.
*}
lemma decomp_H':
includes vectorspace E + subspace H E
assumes y1: "y1 ∈ H" and y2: "y2 ∈ H"
and x': "x' ∉ H" "x' ∈ E" "x' ≠ 0"
and eq: "y1 + a1 · x' = y2 + a2 · x'"
shows "y1 = y2 ∧ a1 = a2"
proof
have c: "y1 = y2 ∧ a1 · x' = a2 · x'"
proof (rule decomp)
show "a1 · x' ∈ lin x'" ..
show "a2 · x' ∈ lin x'" ..
show "H ∩ lin x' = {0}"
proof
show "H ∩ lin x' ⊆ {0}"
proof
fix x assume x: "x ∈ H ∩ lin x'"
then obtain a where xx': "x = a · x'"
by blast
have "x = 0"
proof cases
assume "a = 0"
with xx' and x' show ?thesis by simp
next
assume a: "a ≠ 0"
from x have "x ∈ H" ..
with xx' have "inverse a · a · x' ∈ H" by simp
with a and x' have "x' ∈ H" by (simp add: mult_assoc2)
thus ?thesis by contradiction
qed
thus "x ∈ {0}" ..
qed
show "{0} ⊆ H ∩ lin x'"
proof -
have "0 ∈ H" ..
moreover have "0 ∈ lin x'" ..
ultimately show ?thesis by blast
qed
qed
show "lin x' \<unlhd> E" ..
qed
thus "y1 = y2" ..
from c have "a1 · x' = a2 · x'" ..
with x' show "a1 = a2" by (simp add: mult_right_cancel)
qed
text {*
Since for any element @{text "y + a · x'"} of the direct sum of a
vectorspace @{text H} and the linear closure of @{text x'} the
components @{text "y ∈ H"} and @{text a} are unique, it follows from
@{text "y ∈ H"} that @{text "a = 0"}.
*}
lemma decomp_H'_H:
includes vectorspace E + subspace H E
assumes t: "t ∈ H"
and x': "x' ∉ H" "x' ∈ E" "x' ≠ 0"
shows "(SOME (y, a). t = y + a · x' ∧ y ∈ H) = (t, 0)"
proof (rule, simp_all only: split_paired_all split_conv)
from t x' show "t = t + 0 · x' ∧ t ∈ H" by simp
fix y and a assume ya: "t = y + a · x' ∧ y ∈ H"
have "y = t ∧ a = 0"
proof (rule decomp_H')
from ya x' show "y + a · x' = t + 0 · x'" by simp
from ya show "y ∈ H" ..
qed
with t x' show "(y, a) = (y + a · x', 0)" by simp
qed
text {*
The components @{text "y ∈ H"} and @{text a} in @{text "y + a · x'"}
are unique, so the function @{text h'} defined by
@{text "h' (y + a · x') = h y + a · ξ"} is definite.
*}
lemma h'_definite:
includes var H
assumes h'_def:
"h' ≡ (λx. let (y, a) = SOME (y, a). (x = y + a · x' ∧ y ∈ H)
in (h y) + a * xi)"
and x: "x = y + a · x'"
includes vectorspace E + subspace H E
assumes y: "y ∈ H"
and x': "x' ∉ H" "x' ∈ E" "x' ≠ 0"
shows "h' x = h y + a * xi"
proof -
from x y x' have "x ∈ H + lin x'" by auto
have "∃!p. (λ(y, a). x = y + a · x' ∧ y ∈ H) p" (is "∃!p. ?P p")
proof
from x y show "∃p. ?P p" by blast
fix p q assume p: "?P p" and q: "?P q"
show "p = q"
proof -
from p have xp: "x = fst p + snd p · x' ∧ fst p ∈ H"
by (cases p) simp
from q have xq: "x = fst q + snd q · x' ∧ fst q ∈ H"
by (cases q) simp
have "fst p = fst q ∧ snd p = snd q"
proof (rule decomp_H')
from xp show "fst p ∈ H" ..
from xq show "fst q ∈ H" ..
from xp and xq show "fst p + snd p · x' = fst q + snd q · x'"
by simp
apply_end assumption+
qed
thus ?thesis by (cases p, cases q) simp
qed
qed
hence eq: "(SOME (y, a). x = y + a · x' ∧ y ∈ H) = (y, a)"
by (rule some1_equality) (simp add: x y)
with h'_def show "h' x = h y + a * xi" by (simp add: Let_def)
qed
end
lemma subspace_subset:
subspace U V ==> U ⊆ V
lemma subsetD:
[| subspace U V; x ∈ U |] ==> x ∈ V
lemma subspaceD:
[| subspace U V; x ∈ U |] ==> x ∈ V
lemma rev_subspaceD:
[| x ∈ U; subspace U V |] ==> x ∈ V
lemma diff_closed:
[| subspace U V; vectorspace V; x ∈ U; y ∈ U |] ==> x - y ∈ U
lemma zero:
[| subspace U V; vectorspace V |] ==> (0::'a) ∈ U
lemma neg_closed:
[| subspace U V; vectorspace V; x ∈ U |] ==> - x ∈ U
lemma vectorspace:
[| subspace U V; vectorspace V |] ==> vectorspace U
lemma subspace_refl:
vectorspace V ==> subspace V V
lemma subspace_trans:
[| vectorspace V; subspace U V; subspace V W |] ==> subspace U W
lemma linI:
y = a · x ==> y ∈ lin x
lemma linI':
a · x ∈ lin x
lemma linE:
[| x ∈ lin v; !!a. x = a · v ==> C |] ==> C
lemma x_lin_x:
[| vectorspace V; x ∈ V |] ==> x ∈ lin x
lemma 0_lin_x:
[| vectorspace V; x ∈ V |] ==> (0::'a) ∈ lin x
lemma lin_subspace:
[| vectorspace V; x ∈ V |] ==> subspace (lin x) V
lemma lin_vectorspace:
[| vectorspace V; x ∈ V |] ==> vectorspace (lin x)
lemma sumE:
[| x ∈ U + V; !!u v. [| x = u + v; u ∈ U; v ∈ V |] ==> C |] ==> C
lemma sumI:
[| u ∈ U; v ∈ V; x = u + v |] ==> x ∈ U + V
lemma sumI':
[| u ∈ U; v ∈ V |] ==> u + v ∈ U + V
lemma subspace_sum1:
[| vectorspace U; vectorspace V |] ==> subspace U (U + V)
lemma sum_subspace:
[| subspace U E; vectorspace E; subspace V E |] ==> subspace (U + V) E
lemma sum_vs:
[| subspace U E; subspace V E; vectorspace E |] ==> vectorspace (U + V)
lemma decomp:
[| vectorspace E; subspace U E; subspace V E; U ∩ V = {0::'a}; u1.0 ∈ U; u2.0 ∈ U; v1.0 ∈ V; v2.0 ∈ V; u1.0 + v1.0 = u2.0 + v2.0 |] ==> u1.0 = u2.0 ∧ v1.0 = v2.0
lemma decomp_H':
[| vectorspace E; subspace H E; y1.0 ∈ H; y2.0 ∈ H; x' ∉ H; x' ∈ E; x' ≠ (0::'a); y1.0 + a1.0 · x' = y2.0 + a2.0 · x' |] ==> y1.0 = y2.0 ∧ a1.0 = a2.0
lemma decomp_H'_H:
[| vectorspace E; subspace H E; t ∈ H; x' ∉ H; x' ∈ E; x' ≠ (0::'a) |] ==> (SOME (y, a). t = y + a · x' ∧ y ∈ H) = (t, 0)
lemma h'_definite:
[| h' == %x. let (y, a) = SOME (y, a). x = y + a · x' ∧ y ∈ H in h y + a * xi; x = y + a · x'; vectorspace E; subspace H E; y ∈ H; x' ∉ H; x' ∈ E; x' ≠ (0::'a) |] ==> h' x = h y + a * xi