diff -r 316256709a8c -r 083eedb37a37 src/HOL/Multivariate_Analysis/Linear_Algebra.thy --- a/src/HOL/Multivariate_Analysis/Linear_Algebra.thy Thu Aug 25 11:56:20 2011 -0700 +++ b/src/HOL/Multivariate_Analysis/Linear_Algebra.thy Thu Aug 25 11:57:42 2011 -0700 @@ -628,6 +628,9 @@ apply (rule_tac x="c *\<^sub>R x" in bexI, auto) done +lemma subspace_linear_vimage: "linear f \ subspace S \ subspace (f -` S)" + by (auto simp add: subspace_def linear_def linear_0[of f]) + lemma subspace_linear_preimage: "linear f ==> subspace S ==> subspace {x. f x \ S}" by (auto simp add: subspace_def linear_def linear_0[of f]) @@ -637,6 +640,11 @@ lemma (in real_vector) subspace_inter: "subspace A \ subspace B ==> subspace (A \ B)" by (simp add: subspace_def) +lemma subspace_Times: "\subspace A; subspace B\ \ subspace (A \ B)" + unfolding subspace_def zero_prod_def by simp + +text {* Properties of span. *} + lemma (in real_vector) span_mono: "A \ B ==> span A \ span B" by (metis span_def hull_mono) @@ -655,8 +663,16 @@ by (metis span_def hull_subset subset_eq) (metis subspace_span subspace_def)+ -lemma (in real_vector) span_induct: assumes SP: "\x. x \ S ==> x \ P" - and P: "subspace P" and x: "x \ span S" shows "x \ P" +lemma span_unique: + "\S \ T; subspace T; \T'. \S \ T'; subspace T'\ \ T \ T'\ \ span S = T" + unfolding span_def by (rule hull_unique) + +lemma span_minimal: "S \ T \ subspace T \ span S \ T" + unfolding span_def by (rule hull_minimal) + +lemma (in real_vector) span_induct: + assumes x: "x \ span S" and P: "subspace P" and SP: "\x. x \ S ==> x \ P" + shows "x \ P" proof- from SP have SP': "S \ P" by (simp add: subset_eq) from x hull_minimal[where S=subspace, OF SP' P, unfolded span_def[symmetric]] @@ -788,105 +804,76 @@ text {* Mapping under linear image. *} -lemma span_linear_image: assumes lf: "linear f" +lemma image_subset_iff_subset_vimage: "f ` A \ B \ A \ f -` B" + by auto (* TODO: move *) + +lemma span_linear_image: + assumes lf: "linear f" shows "span (f ` S) = f ` (span S)" -proof- - {fix x - assume x: "x \ span (f ` S)" - have "x \ f ` span S" - apply (rule span_induct[where x=x and S = "f ` S"]) - apply (clarsimp simp add: image_iff) - apply (frule span_superset) - apply blast - apply (rule subspace_linear_image[OF lf]) - apply (rule subspace_span) - apply (rule x) - done} - moreover - {fix x assume x: "x \ span S" - have "x \ {x. f x \ span (f ` S)}" - apply (rule span_induct[where S=S]) - apply simp - apply (rule span_superset) - apply simp - apply (rule subspace_linear_preimage[OF lf subspace_span, of "f ` S"]) - apply (rule x) - done - hence "f x \ span (f ` S)" by simp - } - ultimately show ?thesis by blast +proof (rule span_unique) + show "f ` S \ f ` span S" + by (intro image_mono span_inc) + show "subspace (f ` span S)" + using lf subspace_span by (rule subspace_linear_image) +next + fix T assume "f ` S \ T" and "subspace T" thus "f ` span S \ T" + unfolding image_subset_iff_subset_vimage + by (intro span_minimal subspace_linear_vimage lf) +qed + +lemma span_union: "span (A \ B) = (\(a, b). a + b) ` (span A \ span B)" +proof (rule span_unique) + show "A \ B \ (\(a, b). a + b) ` (span A \ span B)" + by safe (force intro: span_clauses)+ +next + have "linear (\(a, b). a + b)" + by (simp add: linear_def scaleR_add_right) + moreover have "subspace (span A \ span B)" + by (intro subspace_Times subspace_span) + ultimately show "subspace ((\(a, b). a + b) ` (span A \ span B))" + by (rule subspace_linear_image) +next + fix T assume "A \ B \ T" and "subspace T" + thus "(\(a, b). a + b) ` (span A \ span B) \ T" + by (auto intro!: subspace_add elim: span_induct) qed text {* The key breakdown property. *} +lemma span_singleton: "span {x} = range (\k. k *\<^sub>R x)" +proof (rule span_unique) + show "{x} \ range (\k. k *\<^sub>R x)" + by (fast intro: scaleR_one [symmetric]) + show "subspace (range (\k. k *\<^sub>R x))" + unfolding subspace_def + by (auto intro: scaleR_add_left [symmetric]) + fix T assume "{x} \ T" and "subspace T" thus "range (\k. k *\<^sub>R x) \ T" + unfolding subspace_def by auto +qed + +lemma span_insert: + "span (insert a S) = {x. \k. (x - k *\<^sub>R a) \ span S}" +proof - + have "span ({a} \ S) = {x. \k. (x - k *\<^sub>R a) \ span S}" + unfolding span_union span_singleton + apply safe + apply (rule_tac x=k in exI, simp) + apply (erule rev_image_eqI [OF SigmaI [OF rangeI]]) + apply simp + apply (rule right_minus) + done + thus ?thesis by simp +qed + lemma span_breakdown: assumes bS: "b \ S" and aS: "a \ span S" - shows "\k. a - k *\<^sub>R b \ span (S - {b})" (is "?P a") -proof- - {fix x assume xS: "x \ S" - {assume ab: "x = b" - then have "?P x" - apply simp - apply (rule exI[where x="1"], simp) - by (rule span_0)} - moreover - {assume ab: "x \ b" - then have "?P x" using xS - apply - - apply (rule exI[where x=0]) - apply (rule span_superset) - by simp} - ultimately have "x \ Collect ?P" by blast} - moreover have "subspace (Collect ?P)" - unfolding subspace_def - apply auto - apply (rule exI[where x=0]) - using span_0[of "S - {b}"] - apply simp - apply (rule_tac x="k + ka" in exI) - apply (subgoal_tac "x + y - (k + ka) *\<^sub>R b = (x - k*\<^sub>R b) + (y - ka *\<^sub>R b)") - apply (simp only: ) - apply (rule span_add) - apply assumption+ - apply (simp add: algebra_simps) - apply (rule_tac x= "c*k" in exI) - apply (subgoal_tac "c *\<^sub>R x - (c * k) *\<^sub>R b = c*\<^sub>R (x - k*\<^sub>R b)") - apply (simp only: ) - apply (rule span_mul) - apply assumption - by (simp add: algebra_simps) - ultimately have "a \ Collect ?P" using aS by (rule span_induct) - thus "?P a" by simp -qed + shows "\k. a - k *\<^sub>R b \ span (S - {b})" + using assms span_insert [of b "S - {b}"] + by (simp add: insert_absorb) lemma span_breakdown_eq: - "x \ span (insert a S) \ (\k. (x - k *\<^sub>R a) \ span S)" (is "?lhs \ ?rhs") -proof- - {assume x: "x \ span (insert a S)" - from x span_breakdown[of "a" "insert a S" "x"] - have ?rhs apply clarsimp - apply (rule_tac x= "k" in exI) - apply (rule set_rev_mp[of _ "span (S - {a})" _]) - apply assumption - apply (rule span_mono) - apply blast - done} - moreover - { fix k assume k: "x - k *\<^sub>R a \ span S" - have eq: "x = (x - k *\<^sub>R a) + k *\<^sub>R a" by simp - have "(x - k *\<^sub>R a) + k *\<^sub>R a \ span (insert a S)" - apply (rule span_add) - apply (rule set_rev_mp[of _ "span S" _]) - apply (rule k) - apply (rule span_mono) - apply blast - apply (rule span_mul) - apply (rule span_superset) - apply blast - done - then have ?lhs using eq by metis} - ultimately show ?thesis by blast -qed + "x \ span (insert a S) \ (\k. (x - k *\<^sub>R a) \ span S)" + by (simp add: span_insert) text {* Hence some "reversal" results. *} @@ -943,26 +930,16 @@ text {* Transitivity property. *} +lemma span_redundant: "x \ span S \ span (insert x S) = span S" + unfolding span_def by (rule hull_redundant) + lemma span_trans: assumes x: "x \ span S" and y: "y \ span (insert x S)" shows "y \ span S" -proof- - from span_breakdown[of x "insert x S" y, OF insertI1 y] - obtain k where k: "y -k*\<^sub>R x \ span (S - {x})" by auto - have eq: "y = (y - k *\<^sub>R x) + k *\<^sub>R x" by simp - show ?thesis - apply (subst eq) - apply (rule span_add) - apply (rule set_rev_mp) - apply (rule k) - apply (rule span_mono) - apply blast - apply (rule span_mul) - by (rule x) -qed + using assms by (simp only: span_redundant) lemma span_insert_0[simp]: "span (insert 0 S) = span S" - using span_mono[of S "insert 0 S"] by (auto intro: span_trans span_0) + by (simp only: span_redundant span_0) text {* An explicit expansion is sometimes needed. *} @@ -1094,43 +1071,6 @@ ultimately show ?thesis by blast qed -lemma Int_Un_cancel: "(A \ B) \ A = A" "(A \ B) \ B = B" by auto - -lemma span_union: "span (A \ B) = (\(a, b). a + b) ` (span A \ span B)" -proof safe - fix x assume "x \ span (A \ B)" - then obtain S u where S: "finite S" "S \ A \ B" and x: "x = (\v\S. u v *\<^sub>R v)" - unfolding span_explicit by auto - - let ?Sa = "\v\S\A. u v *\<^sub>R v" - let ?Sb = "(\v\S\(B - A). u v *\<^sub>R v)" - show "x \ (\(a, b). a + b) ` (span A \ span B)" - proof - show "x = (case (?Sa, ?Sb) of (a, b) \ a + b)" - unfolding x using S - by (simp, subst setsum_Un_disjoint[symmetric]) (auto intro!: setsum_cong) - - from S have "?Sa \ span A" unfolding span_explicit - by (auto intro!: exI[of _ "S \ A"]) - moreover from S have "?Sb \ span B" unfolding span_explicit - by (auto intro!: exI[of _ "S \ (B - A)"]) - ultimately show "(?Sa, ?Sb) \ span A \ span B" by simp - qed -next - fix a b assume "a \ span A" and "b \ span B" - then obtain Sa ua Sb ub where span: - "finite Sa" "Sa \ A" "a = (\v\Sa. ua v *\<^sub>R v)" - "finite Sb" "Sb \ B" "b = (\v\Sb. ub v *\<^sub>R v)" - unfolding span_explicit by auto - let "?u v" = "(if v \ Sa then ua v else 0) + (if v \ Sb then ub v else 0)" - from span have "finite (Sa \ Sb)" "Sa \ Sb \ A \ B" - and "a + b = (\v\(Sa\Sb). ?u v *\<^sub>R v)" - unfolding setsum_addf scaleR_left_distrib - by (auto simp add: if_distrib cond_application_beta setsum_cases Int_Un_cancel) - thus "a + b \ span (A \ B)" - unfolding span_explicit by (auto intro!: exI[of _ ?u]) -qed - text {* This is useful for building a basis step-by-step. *} lemma independent_insert: