diff -r c968bce3921e -r 4a108f280dc2 src/HOL/Multivariate_Analysis/Linear_Algebra.thy --- a/src/HOL/Multivariate_Analysis/Linear_Algebra.thy Fri Apr 22 15:18:46 2016 +0200 +++ b/src/HOL/Multivariate_Analysis/Linear_Algebra.thy Fri Apr 22 17:22:29 2016 +0200 @@ -268,6 +268,9 @@ lemma span_minimal: "S \ T \ subspace T \ span S \ T" unfolding span_def by (rule hull_minimal) +lemma span_UNIV: "span UNIV = UNIV" + by (intro span_unique) auto + lemma (in real_vector) span_induct: assumes x: "x \ span S" and P: "subspace P" @@ -399,6 +402,10 @@ lemma span_inc: "S \ span S" by (metis subset_eq span_superset) +lemma span_eq: "span S = span T \ S \ span T \ T \ span S" + using span_inc[unfolded subset_eq] using span_mono[of T "span S"] span_mono[of S "span T"] + by (auto simp add: span_span) + lemma (in real_vector) dependent_0: assumes "0 \ A" shows "dependent A" @@ -477,6 +484,12 @@ by (intro span_minimal subspace_linear_vimage lf) qed +lemma spans_image: + assumes lf: "linear f" + and VB: "V \ span B" + shows "f ` V \ span (f ` B)" + unfolding span_linear_image[OF lf] by (metis VB image_mono) + 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)" @@ -980,6 +993,30 @@ "independent B \ \g. linear g \ (\x\B. g x = f x)" using linear_independent_extend_subspace[of B f] by auto +text \Linear functions are equal on a subspace if they are on a spanning set.\ + +lemma subspace_kernel: + assumes lf: "linear f" + shows "subspace {x. f x = 0}" + apply (simp add: subspace_def) + apply (simp add: linear_add[OF lf] linear_cmul[OF lf] linear_0[OF lf]) + done + +lemma linear_eq_0_span: + assumes lf: "linear f" and f0: "\x\B. f x = 0" + shows "\x \ span B. f x = 0" + using f0 subspace_kernel[OF lf] + by (rule span_induct') + +lemma linear_eq_0: "linear f \ S \ span B \ \x\B. f x = 0 \ \x\S. f x = 0" + using linear_eq_0_span[of f B] by auto + +lemma linear_eq_span: "linear f \ linear g \ \x\B. f x = g x \ \x \ span B. f x = g x" + using linear_eq_0_span[of "\x. f x - g x" B] by (auto simp: linear_compose_sub) + +lemma linear_eq: "linear f \ linear g \ S \ span B \ \x\B. f x = g x \ \x\S. f x = g x" + using linear_eq_span[of f g B] by auto + text \The degenerate case of the Exchange Lemma.\ lemma spanning_subset_independent: @@ -1019,6 +1056,114 @@ then show "A \ B" by blast qed +text \Relation between bases and injectivity/surjectivity of map.\ + +lemma spanning_surjective_image: + assumes us: "UNIV \ span S" + and lf: "linear f" + and sf: "surj f" + shows "UNIV \ span (f ` S)" +proof - + have "UNIV \ f ` UNIV" + using sf by (auto simp add: surj_def) + also have " \ \ span (f ` S)" + using spans_image[OF lf us] . + finally show ?thesis . +qed + +lemma independent_inj_on_image: + assumes iS: "independent S" + and lf: "linear f" + and fi: "inj_on f (span S)" + shows "independent (f ` S)" +proof - + { + fix a + assume a: "a \ S" "f a \ span (f ` S - {f a})" + have eq: "f ` S - {f a} = f ` (S - {a})" + using fi \a\S\ by (auto simp add: inj_on_def span_superset) + from a have "f a \ f ` span (S - {a})" + unfolding eq span_linear_image[OF lf, of "S - {a}"] by blast + then have "a \ span (S - {a})" + by (rule inj_on_image_mem_iff_alt[OF fi, rotated]) + (insert span_mono[of "S - {a}" S], auto intro: span_superset \a\S\) + with a(1) iS have False + by (simp add: dependent_def) + } + then show ?thesis + unfolding dependent_def by blast +qed + +lemma independent_injective_image: + "independent S \ linear f \ inj f \ independent (f ` S)" + using independent_inj_on_image[of S f] by (auto simp: subset_inj_on) + +text \Detailed theorems about left and right invertibility in general case.\ + +lemma linear_inj_on_left_inverse: + assumes lf: "linear f" and fi: "inj_on f (span S)" + shows "\g. range g \ span S \ linear g \ (\x\span S. g (f x) = x)" +proof - + obtain B where "independent B" "B \ S" "S \ span B" + using maximal_independent_subset[of S] by auto + then have "span S = span B" + unfolding span_eq by (auto simp: span_superset) + with linear_independent_extend_subspace[OF independent_inj_on_image, OF \independent B\ lf] fi + obtain g where g: "linear g" "\x\f ` B. g x = inv_into B f x" "range g = span (inv_into B f ` f ` B)" + by fastforce + have fB: "inj_on f B" + using fi by (auto simp: \span S = span B\ intro: subset_inj_on span_superset) + + have "\x\span B. g (f x) = x" + proof (intro linear_eq_span) + show "linear (\x. x)" "linear (\x. g (f x))" + using linear_id linear_compose[OF \linear f\ \linear g\] by (auto simp: id_def comp_def) + show "\x \ B. g (f x) = x" + using g fi \span S = span B\ by (auto simp: fB) + qed + moreover + have "inv_into B f ` f ` B \ B" + by (auto simp: fB) + then have "range g \ span S" + unfolding g \span S = span B\ by (intro span_mono) + ultimately show ?thesis + using \span S = span B\ \linear g\ by auto +qed + +lemma linear_injective_left_inverse: "linear f \ inj f \ \g. linear g \ g \ f = id" + using linear_inj_on_left_inverse[of f UNIV] by (auto simp: fun_eq_iff span_UNIV) + +lemma linear_surj_right_inverse: + assumes lf: "linear f" and sf: "span T \ f`span S" + shows "\g. range g \ span S \ linear g \ (\x\span T. f (g x) = x)" +proof - + obtain B where "independent B" "B \ T" "T \ span B" + using maximal_independent_subset[of T] by auto + then have "span T = span B" + unfolding span_eq by (auto simp: span_superset) + + from linear_independent_extend_subspace[OF \independent B\, of "inv_into (span S) f"] + obtain g where g: "linear g" "\x\B. g x = inv_into (span S) f x" "range g = span (inv_into (span S) f`B)" + by auto + moreover have "x \ B \ f (inv_into (span S) f x) = x" for x + using \B \ T\ \span T \ f`span S\ by (intro f_inv_into_f) (auto intro: span_superset) + ultimately have "\x\B. f (g x) = x" + by auto + then have "\x\span B. f (g x) = x" + using linear_id linear_compose[OF \linear g\ \linear f\] + by (intro linear_eq_span) (auto simp: id_def comp_def) + moreover have "inv_into (span S) f ` B \ span S" + using \B \ T\ \span T \ f`span S\ by (auto intro: inv_into_into span_superset) + then have "range g \ span S" + unfolding g by (intro span_minimal subspace_span) auto + ultimately show ?thesis + using \linear g\ \span T = span B\ by auto +qed + +lemma linear_surjective_right_inverse: "linear f \ surj f \ \g. linear g \ f \ g = id" + using linear_surj_right_inverse[of f UNIV UNIV] + by (auto simp: span_UNIV fun_eq_iff) + text \The general case of the Exchange Lemma, the key to what follows.\ lemma exchange_lemma: @@ -1026,7 +1171,6 @@ and i: "independent s" and sp: "s \ span t" shows "\t'. card t' = card t \ finite t' \ s \ t' \ t' \ s \ t \ s \ span t'" - thm maximal_independent_subset_extend[OF _ i, of "s \ t"] using f i sp proof (induct "card (t - s)" arbitrary: s t rule: less_induct) case less @@ -2035,12 +2179,6 @@ shows "span S = span T \ dim S = dim T" by (metis dim_span) -lemma spans_image: - assumes lf: "linear f" - and VB: "V \ span B" - shows "f ` V \ span (f ` B)" - unfolding span_linear_image[OF lf] by (metis VB image_mono) - lemma dim_image_le: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes lf: "linear f" @@ -2060,43 +2198,6 @@ finally show ?thesis . qed -text \Relation between bases and injectivity/surjectivity of map.\ - -lemma spanning_surjective_image: - assumes us: "UNIV \ span S" - and lf: "linear f" - and sf: "surj f" - shows "UNIV \ span (f ` S)" -proof - - have "UNIV \ f ` UNIV" - using sf by (auto simp add: surj_def) - also have " \ \ span (f ` S)" - using spans_image[OF lf us] . - finally show ?thesis . -qed - -lemma independent_injective_image: - assumes iS: "independent S" - and lf: "linear f" - and fi: "inj f" - shows "independent (f ` S)" -proof - - { - fix a - assume a: "a \ S" "f a \ span (f ` S - {f a})" - have eq: "f ` S - {f a} = f ` (S - {a})" - using fi by (auto simp add: inj_on_def) - from a have "f a \ f ` span (S - {a})" - unfolding eq span_linear_image[OF lf, of "S - {a}"] by blast - then have "a \ span (S - {a})" - using fi by (auto simp add: inj_on_def) - with a(1) iS have False - by (simp add: dependent_def) - } - then show ?thesis - unfolding dependent_def by blast -qed - text \Picking an orthogonal replacement for a spanning set.\ lemma vector_sub_project_orthogonal: @@ -2206,10 +2307,6 @@ by auto qed -lemma span_eq: "span S = span T \ S \ span T \ T \ span S" - using span_inc[unfolded subset_eq] using span_mono[of T "span S"] span_mono[of S "span T"] - by (auto simp add: span_span) - text \Low-dimensional subset is in a hyperplane (weak orthogonal complement).\ lemma span_not_univ_orthogonal: @@ -2412,41 +2509,6 @@ by blast qed -text \Linear functions are equal on a subspace if they are on a spanning set.\ - -lemma subspace_kernel: - assumes lf: "linear f" - shows "subspace {x. f x = 0}" - apply (simp add: subspace_def) - apply (simp add: linear_add[OF lf] linear_cmul[OF lf] linear_0[OF lf]) - done - -lemma linear_eq_0_span: - assumes lf: "linear f" and f0: "\x\B. f x = 0" - shows "\x \ span B. f x = 0" - using f0 subspace_kernel[OF lf] - by (rule span_induct') - -lemma linear_eq_0: - assumes lf: "linear f" - and SB: "S \ span B" - and f0: "\x\B. f x = 0" - shows "\x\S. f x = 0" - by (metis linear_eq_0_span[OF lf] subset_eq SB f0) - -lemma linear_eq: - assumes lf: "linear f" - and lg: "linear g" - and S: "S \ span B" - and fg: "\x\B. f x = g x" - shows "\x\S. f x = g x" -proof - - let ?h = "\x. f x - g x" - from fg have fg': "\x\ B. ?h x = 0" by simp - from linear_eq_0[OF linear_compose_sub[OF lf lg] S fg'] - show ?thesis by simp -qed - lemma linear_eq_stdbasis: fixes f :: "'a::euclidean_space \ _" assumes lf: "linear f" @@ -2493,43 +2555,6 @@ shows "f = g" using bilinear_eq[OF bf bg equalityD2[OF span_Basis] equalityD2[OF span_Basis] fg] by blast -text \Detailed theorems about left and right invertibility in general case.\ - -lemma linear_injective_left_inverse: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes lf: "linear f" - and fi: "inj f" - shows "\g. linear g \ g \ f = id" -proof - - from linear_independent_extend[OF independent_injective_image, OF independent_Basis, OF lf fi] - obtain h :: "'b \ 'a" where h: "linear h" "\x \ f ` Basis. h x = inv f x" - by blast - from h(2) have th: "\i\Basis. (h \ f) i = id i" - using inv_o_cancel[OF fi, unfolded fun_eq_iff id_def o_def] - by auto - from linear_eq_stdbasis[OF linear_compose[OF lf h(1)] linear_id th] - have "h \ f = id" . - then show ?thesis - using h(1) by blast -qed - -lemma linear_surjective_right_inverse: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes lf: "linear f" - and sf: "surj f" - shows "\g. linear g \ f \ g = id" -proof - - from linear_independent_extend[OF independent_Basis[where 'a='b],of "inv f"] - obtain h :: "'b \ 'a" where h: "linear h" "\x\Basis. h x = inv f x" - by blast - from h(2) have th: "\i\Basis. (f \ h) i = id i" - using sf by (auto simp add: surj_iff_all) - from linear_eq_stdbasis[OF linear_compose[OF h(1) lf] linear_id th] - have "f \ h = id" . - then show ?thesis - using h(1) by blast -qed - text \An injective map @{typ "'a::euclidean_space \ 'b::euclidean_space"} is also surjective.\ lemma linear_injective_imp_surjective: