wenzelm@31795: (* Title: HOL/Hahn_Banach/Subspace.thy wenzelm@7566: Author: Gertrud Bauer, TU Munich wenzelm@7566: *) wenzelm@7535: wenzelm@9035: header {* Subspaces *} wenzelm@7808: wenzelm@27612: theory Subspace huffman@44190: imports Vector_Space "~~/src/HOL/Library/Set_Algebras" wenzelm@27612: begin wenzelm@7535: wenzelm@9035: subsection {* Definition *} wenzelm@7535: wenzelm@10687: text {* wenzelm@10687: A non-empty subset @{text U} of a vector space @{text V} is a wenzelm@10687: \emph{subspace} of @{text V}, iff @{text U} is closed under addition wenzelm@10687: and scalar multiplication. wenzelm@10687: *} wenzelm@7917: ballarin@29234: locale subspace = ballarin@29234: fixes U :: "'a\{minus, plus, zero, uminus} set" and V wenzelm@13515: assumes non_empty [iff, intro]: "U \ {}" wenzelm@13515: and subset [iff]: "U \ V" wenzelm@13515: and add_closed [iff]: "x \ U \ y \ U \ x + y \ U" wenzelm@13515: and mult_closed [iff]: "x \ U \ a \ x \ U" wenzelm@7535: wenzelm@21210: notation (symbols) wenzelm@19736: subspace (infix "\" 50) ballarin@14254: wenzelm@19736: declare vectorspace.intro [intro?] subspace.intro [intro?] wenzelm@7535: wenzelm@13515: lemma subspace_subset [elim]: "U \ V \ U \ V" wenzelm@13515: by (rule subspace.subset) wenzelm@7566: wenzelm@13515: lemma (in subspace) subsetD [iff]: "x \ U \ x \ V" wenzelm@13515: using subset by blast wenzelm@7566: wenzelm@13515: lemma subspaceD [elim]: "U \ V \ x \ U \ x \ V" wenzelm@13515: by (rule subspace.subsetD) wenzelm@7535: wenzelm@13515: lemma rev_subspaceD [elim?]: "x \ U \ U \ V \ x \ V" wenzelm@13515: by (rule subspace.subsetD) wenzelm@13515: wenzelm@13547: lemma (in subspace) diff_closed [iff]: ballarin@27611: assumes "vectorspace V" wenzelm@27612: assumes x: "x \ U" and y: "y \ U" wenzelm@27612: shows "x - y \ U" ballarin@27611: proof - ballarin@29234: interpret vectorspace V by fact wenzelm@27612: from x y show ?thesis by (simp add: diff_eq1 negate_eq1) ballarin@27611: qed wenzelm@7917: wenzelm@13515: text {* wenzelm@13515: \medskip Similar as for linear spaces, the existence of the zero wenzelm@13515: element in every subspace follows from the non-emptiness of the wenzelm@13515: carrier set and by vector space laws. wenzelm@13515: *} wenzelm@13515: wenzelm@13547: lemma (in subspace) zero [intro]: ballarin@27611: assumes "vectorspace V" wenzelm@13547: shows "0 \ U" wenzelm@10687: proof - wenzelm@30729: interpret V: vectorspace V by fact ballarin@29234: have "U \ {}" by (rule non_empty) wenzelm@13515: then obtain x where x: "x \ U" by blast wenzelm@27612: then have "x \ V" .. then have "0 = x - x" by simp ballarin@29234: also from `vectorspace V` x x have "\ \ U" by (rule diff_closed) wenzelm@13515: finally show ?thesis . wenzelm@9035: qed wenzelm@7535: wenzelm@13547: lemma (in subspace) neg_closed [iff]: ballarin@27611: assumes "vectorspace V" wenzelm@27612: assumes x: "x \ U" wenzelm@27612: shows "- x \ U" ballarin@27611: proof - ballarin@29234: interpret vectorspace V by fact wenzelm@27612: from x show ?thesis by (simp add: negate_eq1) ballarin@27611: qed wenzelm@13515: wenzelm@10687: text {* \medskip Further derived laws: every subspace is a vector space. *} wenzelm@7535: wenzelm@13547: lemma (in subspace) vectorspace [iff]: ballarin@27611: assumes "vectorspace V" wenzelm@13547: shows "vectorspace U" ballarin@27611: proof - ballarin@29234: interpret vectorspace V by fact wenzelm@27612: show ?thesis wenzelm@27612: proof ballarin@27611: show "U \ {}" .. ballarin@27611: fix x y z assume x: "x \ U" and y: "y \ U" and z: "z \ U" ballarin@27611: fix a b :: real ballarin@27611: from x y show "x + y \ U" by simp ballarin@27611: from x show "a \ x \ U" by simp ballarin@27611: from x y z show "(x + y) + z = x + (y + z)" by (simp add: add_ac) ballarin@27611: from x y show "x + y = y + x" by (simp add: add_ac) ballarin@27611: from x show "x - x = 0" by simp ballarin@27611: from x show "0 + x = x" by simp ballarin@27611: from x y show "a \ (x + y) = a \ x + a \ y" by (simp add: distrib) ballarin@27611: from x show "(a + b) \ x = a \ x + b \ x" by (simp add: distrib) ballarin@27611: from x show "(a * b) \ x = a \ b \ x" by (simp add: mult_assoc) ballarin@27611: from x show "1 \ x = x" by simp ballarin@27611: from x show "- x = - 1 \ x" by (simp add: negate_eq1) ballarin@27611: from x y show "x - y = x + - y" by (simp add: diff_eq1) ballarin@27611: qed wenzelm@9035: qed wenzelm@7535: wenzelm@13515: wenzelm@9035: text {* The subspace relation is reflexive. *} wenzelm@7917: wenzelm@13515: lemma (in vectorspace) subspace_refl [intro]: "V \ V" wenzelm@10687: proof wenzelm@13515: show "V \ {}" .. wenzelm@10687: show "V \ V" .. wenzelm@44887: next wenzelm@13515: fix x y assume x: "x \ V" and y: "y \ V" wenzelm@13515: fix a :: real wenzelm@13515: from x y show "x + y \ V" by simp wenzelm@13515: from x show "a \ x \ V" by simp wenzelm@9035: qed wenzelm@7535: wenzelm@9035: text {* The subspace relation is transitive. *} wenzelm@7917: wenzelm@13515: lemma (in vectorspace) subspace_trans [trans]: wenzelm@13515: "U \ V \ V \ W \ U \ W" wenzelm@10687: proof wenzelm@13515: assume uv: "U \ V" and vw: "V \ W" wenzelm@13515: from uv show "U \ {}" by (rule subspace.non_empty) wenzelm@13515: show "U \ W" wenzelm@13515: proof - wenzelm@13515: from uv have "U \ V" by (rule subspace.subset) wenzelm@13515: also from vw have "V \ W" by (rule subspace.subset) wenzelm@13515: finally show ?thesis . wenzelm@9035: qed wenzelm@13515: fix x y assume x: "x \ U" and y: "y \ U" wenzelm@13515: from uv and x y show "x + y \ U" by (rule subspace.add_closed) wenzelm@13515: from uv and x show "\a. a \ x \ U" by (rule subspace.mult_closed) wenzelm@9035: qed wenzelm@7535: wenzelm@7535: wenzelm@9035: subsection {* Linear closure *} wenzelm@7808: wenzelm@10687: text {* wenzelm@10687: The \emph{linear closure} of a vector @{text x} is the set of all wenzelm@10687: scalar multiples of @{text x}. wenzelm@10687: *} wenzelm@7535: wenzelm@44887: definition lin :: "('a::{minus, plus, zero}) \ 'a set" wenzelm@44887: where "lin x = {a \ x | a. True}" wenzelm@7535: wenzelm@13515: lemma linI [intro]: "y = a \ x \ y \ lin x" wenzelm@27612: unfolding lin_def by blast wenzelm@7535: wenzelm@13515: lemma linI' [iff]: "a \ x \ lin x" wenzelm@27612: unfolding lin_def by blast wenzelm@13515: wenzelm@27612: lemma linE [elim]: "x \ lin v \ (\a::real. x = a \ v \ C) \ C" wenzelm@27612: unfolding lin_def by blast wenzelm@13515: wenzelm@7656: wenzelm@9035: text {* Every vector is contained in its linear closure. *} wenzelm@7917: wenzelm@13515: lemma (in vectorspace) x_lin_x [iff]: "x \ V \ x \ lin x" wenzelm@13515: proof - wenzelm@13515: assume "x \ V" wenzelm@27612: then have "x = 1 \ x" by simp wenzelm@13515: also have "\ \ lin x" .. wenzelm@13515: finally show ?thesis . wenzelm@13515: qed wenzelm@13515: wenzelm@13515: lemma (in vectorspace) "0_lin_x" [iff]: "x \ V \ 0 \ lin x" wenzelm@13515: proof wenzelm@13515: assume "x \ V" wenzelm@27612: then show "0 = 0 \ x" by simp wenzelm@13515: qed wenzelm@7535: wenzelm@9035: text {* Any linear closure is a subspace. *} wenzelm@7917: wenzelm@13515: lemma (in vectorspace) lin_subspace [intro]: wenzelm@44887: assumes x: "x \ V" wenzelm@44887: shows "lin x \ V" wenzelm@9035: proof wenzelm@44887: from x show "lin x \ {}" by auto wenzelm@44887: next wenzelm@10687: show "lin x \ V" wenzelm@13515: proof wenzelm@13515: fix x' assume "x' \ lin x" wenzelm@13515: then obtain a where "x' = a \ x" .. wenzelm@13515: with x show "x' \ V" by simp wenzelm@9035: qed wenzelm@44887: next wenzelm@13515: fix x' x'' assume x': "x' \ lin x" and x'': "x'' \ lin x" wenzelm@13515: show "x' + x'' \ lin x" wenzelm@13515: proof - wenzelm@13515: from x' obtain a' where "x' = a' \ x" .. wenzelm@13515: moreover from x'' obtain a'' where "x'' = a'' \ x" .. wenzelm@13515: ultimately have "x' + x'' = (a' + a'') \ x" wenzelm@13515: using x by (simp add: distrib) wenzelm@13515: also have "\ \ lin x" .. wenzelm@13515: finally show ?thesis . wenzelm@9035: qed wenzelm@13515: fix a :: real wenzelm@13515: show "a \ x' \ lin x" wenzelm@13515: proof - wenzelm@13515: from x' obtain a' where "x' = a' \ x" .. wenzelm@13515: with x have "a \ x' = (a * a') \ x" by (simp add: mult_assoc) wenzelm@13515: also have "\ \ lin x" .. wenzelm@13515: finally show ?thesis . wenzelm@10687: qed wenzelm@9035: qed wenzelm@7535: wenzelm@13515: wenzelm@9035: text {* Any linear closure is a vector space. *} wenzelm@7917: wenzelm@13515: lemma (in vectorspace) lin_vectorspace [intro]: wenzelm@23378: assumes "x \ V" wenzelm@23378: shows "vectorspace (lin x)" wenzelm@23378: proof - wenzelm@23378: from `x \ V` have "subspace (lin x) V" wenzelm@23378: by (rule lin_subspace) wenzelm@26199: from this and vectorspace_axioms show ?thesis wenzelm@23378: by (rule subspace.vectorspace) wenzelm@23378: qed wenzelm@7808: wenzelm@7808: wenzelm@9035: subsection {* Sum of two vectorspaces *} wenzelm@7808: wenzelm@10687: text {* wenzelm@10687: The \emph{sum} of two vectorspaces @{text U} and @{text V} is the wenzelm@10687: set of all sums of elements from @{text U} and @{text V}. wenzelm@10687: *} wenzelm@7535: krauss@47445: lemma sum_def: "U + V = {u + v | u v. u \ U \ v \ V}" huffman@44190: unfolding set_plus_def by auto wenzelm@7917: wenzelm@13515: lemma sumE [elim]: krauss@47445: "x \ U + V \ (\u v. x = u + v \ u \ U \ v \ V \ C) \ C" wenzelm@27612: unfolding sum_def by blast wenzelm@7535: wenzelm@13515: lemma sumI [intro]: krauss@47445: "u \ U \ v \ V \ x = u + v \ x \ U + V" wenzelm@27612: unfolding sum_def by blast wenzelm@7566: wenzelm@13515: lemma sumI' [intro]: krauss@47445: "u \ U \ v \ V \ u + v \ U + V" wenzelm@27612: unfolding sum_def by blast wenzelm@7917: krauss@47445: text {* @{text U} is a subspace of @{text "U + V"}. *} wenzelm@7535: wenzelm@13515: lemma subspace_sum1 [iff]: ballarin@27611: assumes "vectorspace U" "vectorspace V" krauss@47445: shows "U \ U + V" ballarin@27611: proof - ballarin@29234: interpret vectorspace U by fact ballarin@29234: interpret vectorspace V by fact wenzelm@27612: show ?thesis wenzelm@27612: proof ballarin@27611: show "U \ {}" .. krauss@47445: show "U \ U + V" ballarin@27611: proof ballarin@27611: fix x assume x: "x \ U" ballarin@27611: moreover have "0 \ V" .. krauss@47445: ultimately have "x + 0 \ U + V" .. krauss@47445: with x show "x \ U + V" by simp ballarin@27611: qed ballarin@27611: fix x y assume x: "x \ U" and "y \ U" wenzelm@27612: then show "x + y \ U" by simp ballarin@27611: from x show "\a. a \ x \ U" by simp wenzelm@9035: qed wenzelm@9035: qed wenzelm@7535: wenzelm@13515: text {* The sum of two subspaces is again a subspace. *} wenzelm@7917: wenzelm@13515: lemma sum_subspace [intro?]: ballarin@27611: assumes "subspace U E" "vectorspace E" "subspace V E" krauss@47445: shows "U + V \ E" ballarin@27611: proof - ballarin@29234: interpret subspace U E by fact ballarin@29234: interpret vectorspace E by fact ballarin@29234: interpret subspace V E by fact wenzelm@27612: show ?thesis wenzelm@27612: proof krauss@47445: have "0 \ U + V" ballarin@27611: proof ballarin@27611: show "0 \ U" using `vectorspace E` .. ballarin@27611: show "0 \ V" using `vectorspace E` .. ballarin@27611: show "(0::'a) = 0 + 0" by simp ballarin@27611: qed krauss@47445: then show "U + V \ {}" by blast krauss@47445: show "U + V \ E" ballarin@27611: proof krauss@47445: fix x assume "x \ U + V" ballarin@27611: then obtain u v where "x = u + v" and wenzelm@32960: "u \ U" and "v \ V" .. ballarin@27611: then show "x \ E" by simp ballarin@27611: qed wenzelm@44887: next krauss@47445: fix x y assume x: "x \ U + V" and y: "y \ U + V" krauss@47445: show "x + y \ U + V" ballarin@27611: proof - ballarin@27611: from x obtain ux vx where "x = ux + vx" and "ux \ U" and "vx \ V" .. ballarin@27611: moreover ballarin@27611: from y obtain uy vy where "y = uy + vy" and "uy \ U" and "vy \ V" .. ballarin@27611: ultimately ballarin@27611: have "ux + uy \ U" wenzelm@32960: and "vx + vy \ V" wenzelm@32960: and "x + y = (ux + uy) + (vx + vy)" wenzelm@32960: using x y by (simp_all add: add_ac) wenzelm@27612: then show ?thesis .. ballarin@27611: qed krauss@47445: fix a show "a \ x \ U + V" ballarin@27611: proof - ballarin@27611: from x obtain u v where "x = u + v" and "u \ U" and "v \ V" .. wenzelm@27612: then have "a \ u \ U" and "a \ v \ V" wenzelm@32960: and "a \ x = (a \ u) + (a \ v)" by (simp_all add: distrib) wenzelm@27612: then show ?thesis .. ballarin@27611: qed wenzelm@9035: qed wenzelm@9035: qed wenzelm@7535: wenzelm@9035: text{* The sum of two subspaces is a vectorspace. *} wenzelm@7917: wenzelm@13515: lemma sum_vs [intro?]: krauss@47445: "U \ E \ V \ E \ vectorspace E \ vectorspace (U + V)" wenzelm@13547: by (rule subspace.vectorspace) (rule sum_subspace) wenzelm@7535: wenzelm@7808: wenzelm@9035: subsection {* Direct sums *} wenzelm@7808: wenzelm@10687: text {* wenzelm@10687: The sum of @{text U} and @{text V} is called \emph{direct}, iff the wenzelm@10687: zero element is the only common element of @{text U} and @{text wenzelm@10687: V}. For every element @{text x} of the direct sum of @{text U} and wenzelm@10687: @{text V} the decomposition in @{text "x = u + v"} with wenzelm@10687: @{text "u \ U"} and @{text "v \ V"} is unique. wenzelm@10687: *} wenzelm@7808: wenzelm@10687: lemma decomp: ballarin@27611: assumes "vectorspace E" "subspace U E" "subspace V E" wenzelm@13515: assumes direct: "U \ V = {0}" wenzelm@13515: and u1: "u1 \ U" and u2: "u2 \ U" wenzelm@13515: and v1: "v1 \ V" and v2: "v2 \ V" wenzelm@13515: and sum: "u1 + v1 = u2 + v2" wenzelm@13515: shows "u1 = u2 \ v1 = v2" ballarin@27611: proof - ballarin@29234: interpret vectorspace E by fact ballarin@29234: interpret subspace U E by fact ballarin@29234: interpret subspace V E by fact wenzelm@27612: show ?thesis wenzelm@27612: proof ballarin@27611: have U: "vectorspace U" (* FIXME: use interpret *) ballarin@27611: using `subspace U E` `vectorspace E` by (rule subspace.vectorspace) ballarin@27611: have V: "vectorspace V" ballarin@27611: using `subspace V E` `vectorspace E` by (rule subspace.vectorspace) ballarin@27611: from u1 u2 v1 v2 and sum have eq: "u1 - u2 = v2 - v1" ballarin@27611: by (simp add: add_diff_swap) ballarin@27611: from u1 u2 have u: "u1 - u2 \ U" ballarin@27611: by (rule vectorspace.diff_closed [OF U]) ballarin@27611: with eq have v': "v2 - v1 \ U" by (simp only:) ballarin@27611: from v2 v1 have v: "v2 - v1 \ V" ballarin@27611: by (rule vectorspace.diff_closed [OF V]) ballarin@27611: with eq have u': " u1 - u2 \ V" by (simp only:) ballarin@27611: ballarin@27611: show "u1 = u2" ballarin@27611: proof (rule add_minus_eq) ballarin@27611: from u1 show "u1 \ E" .. ballarin@27611: from u2 show "u2 \ E" .. ballarin@27611: from u u' and direct show "u1 - u2 = 0" by blast ballarin@27611: qed ballarin@27611: show "v1 = v2" ballarin@27611: proof (rule add_minus_eq [symmetric]) ballarin@27611: from v1 show "v1 \ E" .. ballarin@27611: from v2 show "v2 \ E" .. ballarin@27611: from v v' and direct show "v2 - v1 = 0" by blast ballarin@27611: qed wenzelm@9035: qed wenzelm@9035: qed wenzelm@7656: wenzelm@10687: text {* wenzelm@10687: An application of the previous lemma will be used in the proof of wenzelm@10687: the Hahn-Banach Theorem (see page \pageref{decomp-H-use}): for any wenzelm@10687: element @{text "y + a \ x\<^sub>0"} of the direct sum of a wenzelm@10687: vectorspace @{text H} and the linear closure of @{text "x\<^sub>0"} wenzelm@10687: the components @{text "y \ H"} and @{text a} are uniquely wenzelm@10687: determined. wenzelm@10687: *} wenzelm@7917: wenzelm@10687: lemma decomp_H': ballarin@27611: assumes "vectorspace E" "subspace H E" wenzelm@13515: assumes y1: "y1 \ H" and y2: "y2 \ H" wenzelm@13515: and x': "x' \ H" "x' \ E" "x' \ 0" wenzelm@13515: and eq: "y1 + a1 \ x' = y2 + a2 \ x'" wenzelm@13515: shows "y1 = y2 \ a1 = a2" ballarin@27611: proof - ballarin@29234: interpret vectorspace E by fact ballarin@29234: interpret subspace H E by fact wenzelm@27612: show ?thesis wenzelm@27612: proof ballarin@27611: have c: "y1 = y2 \ a1 \ x' = a2 \ x'" ballarin@27611: proof (rule decomp) ballarin@27611: show "a1 \ x' \ lin x'" .. ballarin@27611: show "a2 \ x' \ lin x'" .. ballarin@27611: show "H \ lin x' = {0}" wenzelm@13515: proof wenzelm@32960: show "H \ lin x' \ {0}" wenzelm@32960: proof ballarin@27611: fix x assume x: "x \ H \ lin x'" ballarin@27611: then obtain a where xx': "x = a \ x'" ballarin@27611: by blast ballarin@27611: have "x = 0" ballarin@27611: proof cases ballarin@27611: assume "a = 0" ballarin@27611: with xx' and x' show ?thesis by simp ballarin@27611: next ballarin@27611: assume a: "a \ 0" ballarin@27611: from x have "x \ H" .. ballarin@27611: with xx' have "inverse a \ a \ x' \ H" by simp ballarin@27611: with a and x' have "x' \ H" by (simp add: mult_assoc2) ballarin@27611: with `x' \ H` show ?thesis by contradiction ballarin@27611: qed wenzelm@27612: then show "x \ {0}" .. wenzelm@32960: qed wenzelm@32960: show "{0} \ H \ lin x'" wenzelm@32960: proof - ballarin@27611: have "0 \ H" using `vectorspace E` .. ballarin@27611: moreover have "0 \ lin x'" using `x' \ E` .. ballarin@27611: ultimately show ?thesis by blast wenzelm@32960: qed wenzelm@9035: qed ballarin@27611: show "lin x' \ E" using `x' \ E` .. ballarin@27611: qed (rule `vectorspace E`, rule `subspace H E`, rule y1, rule y2, rule eq) wenzelm@27612: then show "y1 = y2" .. ballarin@27611: from c have "a1 \ x' = a2 \ x'" .. ballarin@27611: with x' show "a1 = a2" by (simp add: mult_right_cancel) ballarin@27611: qed wenzelm@9035: qed wenzelm@7535: wenzelm@10687: text {* wenzelm@10687: Since for any element @{text "y + a \ x'"} of the direct sum of a wenzelm@10687: vectorspace @{text H} and the linear closure of @{text x'} the wenzelm@10687: components @{text "y \ H"} and @{text a} are unique, it follows from wenzelm@10687: @{text "y \ H"} that @{text "a = 0"}. wenzelm@10687: *} wenzelm@7917: wenzelm@10687: lemma decomp_H'_H: ballarin@27611: assumes "vectorspace E" "subspace H E" wenzelm@13515: assumes t: "t \ H" wenzelm@13515: and x': "x' \ H" "x' \ E" "x' \ 0" wenzelm@13515: shows "(SOME (y, a). t = y + a \ x' \ y \ H) = (t, 0)" ballarin@27611: proof - ballarin@29234: interpret vectorspace E by fact ballarin@29234: interpret subspace H E by fact wenzelm@27612: show ?thesis wenzelm@27612: proof (rule, simp_all only: split_paired_all split_conv) ballarin@27611: from t x' show "t = t + 0 \ x' \ t \ H" by simp ballarin@27611: fix y and a assume ya: "t = y + a \ x' \ y \ H" ballarin@27611: have "y = t \ a = 0" ballarin@27611: proof (rule decomp_H') ballarin@27611: from ya x' show "y + a \ x' = t + 0 \ x'" by simp ballarin@27611: from ya show "y \ H" .. ballarin@27611: qed (rule `vectorspace E`, rule `subspace H E`, rule t, (rule x')+) ballarin@27611: with t x' show "(y, a) = (y + a \ x', 0)" by simp ballarin@27611: qed wenzelm@13515: qed wenzelm@7535: wenzelm@10687: text {* wenzelm@10687: The components @{text "y \ H"} and @{text a} in @{text "y + a \ x'"} wenzelm@10687: are unique, so the function @{text h'} defined by wenzelm@10687: @{text "h' (y + a \ x') = h y + a \ \"} is definite. wenzelm@10687: *} wenzelm@7917: bauerg@9374: lemma h'_definite: ballarin@27611: fixes H wenzelm@13515: assumes h'_def: wenzelm@44887: "h' \ \x. wenzelm@44887: let (y, a) = SOME (y, a). (x = y + a \ x' \ y \ H) wenzelm@44887: in (h y) + a * xi" wenzelm@13515: and x: "x = y + a \ x'" ballarin@27611: assumes "vectorspace E" "subspace H E" wenzelm@13515: assumes y: "y \ H" wenzelm@13515: and x': "x' \ H" "x' \ E" "x' \ 0" wenzelm@13515: shows "h' x = h y + a * xi" wenzelm@10687: proof - ballarin@29234: interpret vectorspace E by fact ballarin@29234: interpret subspace H E by fact krauss@47445: from x y x' have "x \ H + lin x'" by auto wenzelm@13515: have "\!p. (\(y, a). x = y + a \ x' \ y \ H) p" (is "\!p. ?P p") wenzelm@18689: proof (rule ex_ex1I) wenzelm@13515: from x y show "\p. ?P p" by blast wenzelm@13515: fix p q assume p: "?P p" and q: "?P q" wenzelm@13515: show "p = q" wenzelm@9035: proof - wenzelm@13515: from p have xp: "x = fst p + snd p \ x' \ fst p \ H" wenzelm@13515: by (cases p) simp wenzelm@13515: from q have xq: "x = fst q + snd q \ x' \ fst q \ H" wenzelm@13515: by (cases q) simp wenzelm@13515: have "fst p = fst q \ snd p = snd q" wenzelm@13515: proof (rule decomp_H') wenzelm@13515: from xp show "fst p \ H" .. wenzelm@13515: from xq show "fst q \ H" .. wenzelm@13515: from xp and xq show "fst p + snd p \ x' = fst q + snd q \ x'" wenzelm@13515: by simp wenzelm@23378: qed (rule `vectorspace E`, rule `subspace H E`, (rule x')+) wenzelm@27612: then show ?thesis by (cases p, cases q) simp wenzelm@9035: qed wenzelm@9035: qed wenzelm@27612: then have eq: "(SOME (y, a). x = y + a \ x' \ y \ H) = (y, a)" wenzelm@13515: by (rule some1_equality) (simp add: x y) wenzelm@13515: with h'_def show "h' x = h y + a * xi" by (simp add: Let_def) wenzelm@9035: qed wenzelm@7535: wenzelm@10687: end