| author | wenzelm | 
| Sat, 26 May 2018 21:23:51 +0200 | |
| changeset 68293 | 2bc4e5d9cca6 | 
| parent 68189 | 6163c90694ef | 
| child 68833 | fde093888c16 | 
| permissions | -rw-r--r-- | 
| 68189 | 1 | (* Title: HOL/Analysis/Cartesian_Space.thy | 
| 2 | Author: Amine Chaieb, University of Cambridge | |
| 3 | Author: Jose Divasón <jose.divasonm at unirioja.es> | |
| 4 | Author: Jesús Aransay <jesus-maria.aransay at unirioja.es> | |
| 5 | Author: Johannes Hölzl, VU Amsterdam | |
| 6 | Author: Fabian Immler, TUM | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 7 | *) | 
| 68189 | 8 | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 9 | theory Cartesian_Space | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 10 | imports | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 11 | Finite_Cartesian_Product Linear_Algebra | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 12 | begin | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 13 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 14 | definition "cart_basis = {axis i 1 | i. i\<in>UNIV}"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 15 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 16 | lemma finite_cart_basis: "finite (cart_basis)" unfolding cart_basis_def | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 17 | using finite_Atleast_Atmost_nat by fastforce | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 18 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 19 | lemma card_cart_basis: "card (cart_basis::('a::zero_neq_one^'i) set) = CARD('i)"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 20 | unfolding cart_basis_def Setcompr_eq_image | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 21 | by (rule card_image) (auto simp: inj_on_def axis_eq_axis) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 22 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 23 | interpretation vec: vector_space "( *s) " | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 24 | by unfold_locales (vector algebra_simps)+ | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 25 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 26 | lemma independent_cart_basis: | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 27 | "vec.independent (cart_basis)" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 28 | proof (rule vec.independent_if_scalars_zero) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 29 | show "finite (cart_basis)" using finite_cart_basis . | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 30 |   fix f::"('a, 'b) vec \<Rightarrow> 'a" and x::"('a, 'b) vec"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 31 | assume eq_0: "(\<Sum>x\<in>cart_basis. f x *s x) = 0" and x_in: "x \<in> cart_basis" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 32 | obtain i where x: "x = axis i 1" using x_in unfolding cart_basis_def by auto | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 33 |   have sum_eq_0: "(\<Sum>x\<in>(cart_basis) - {x}. f x * (x $ i)) = 0"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 34 | proof (rule sum.neutral, rule ballI) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 35 |     fix xa assume xa: "xa \<in> cart_basis - {x}"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 36 | obtain a where a: "xa = axis a 1" and a_not_i: "a \<noteq> i" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 37 | using xa x unfolding cart_basis_def by auto | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 38 | have "xa $ i = 0" unfolding a axis_def using a_not_i by auto | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 39 | thus "f xa * xa $ i = 0" by simp | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 40 | qed | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 41 | have "0 = (\<Sum>x\<in>cart_basis. f x *s x) $ i" using eq_0 by simp | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 42 | also have "... = (\<Sum>x\<in>cart_basis. (f x *s x) $ i)" unfolding sum_component .. | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 43 | also have "... = (\<Sum>x\<in>cart_basis. f x * (x $ i))" unfolding vector_smult_component .. | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 44 |   also have "... = f x * (x $ i) + (\<Sum>x\<in>(cart_basis) - {x}. f x * (x $ i))"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 45 | by (rule sum.remove[OF finite_cart_basis x_in]) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 46 | also have "... = f x * (x $ i)" unfolding sum_eq_0 by simp | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 47 | also have "... = f x" unfolding x axis_def by auto | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 48 | finally show "f x = 0" .. | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 49 | qed | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 50 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 51 | lemma span_cart_basis: | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 52 | "vec.span (cart_basis) = UNIV" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 53 | proof (auto) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 54 |   fix x::"('a, 'b) vec"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 55 | let ?f="\<lambda>v. x $ (THE i. v = axis i 1)" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 56 | show "x \<in> vec.span (cart_basis)" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 57 | apply (unfold vec.span_finite[OF finite_cart_basis]) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 58 | apply (rule image_eqI[of _ _ ?f]) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 59 | apply (subst vec_eq_iff) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 60 | apply clarify | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 61 | proof - | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 62 | fix i::'b | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 63 | let ?w = "axis i (1::'a)" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 64 | have the_eq_i: "(THE a. ?w = axis a 1) = i" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 65 | by (rule the_equality, auto simp: axis_eq_axis) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 66 |     have sum_eq_0: "(\<Sum>v\<in>(cart_basis) - {?w}. x $ (THE i. v = axis i 1) * v $ i) = 0"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 67 | proof (rule sum.neutral, rule ballI) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 68 |       fix xa::"('a, 'b) vec"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 69 |       assume xa: "xa \<in> cart_basis - {?w}"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 70 | obtain j where j: "xa = axis j 1" and i_not_j: "i \<noteq> j" using xa unfolding cart_basis_def by auto | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 71 | have the_eq_j: "(THE i. xa = axis i 1) = j" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 72 | proof (rule the_equality) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 73 | show "xa = axis j 1" using j . | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 74 | show "\<And>i. xa = axis i 1 \<Longrightarrow> i = j" by (metis axis_eq_axis j zero_neq_one) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 75 | qed | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 76 | show "x $ (THE i. xa = axis i 1) * xa $ i = 0" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 77 | apply (subst (2) j) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 78 | unfolding the_eq_j unfolding axis_def using i_not_j by simp | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 79 | qed | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 80 | have "(\<Sum>v\<in>cart_basis. x $ (THE i. v = axis i 1) *s v) $ i = | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 81 | (\<Sum>v\<in>cart_basis. (x $ (THE i. v = axis i 1) *s v) $ i)" unfolding sum_component .. | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 82 | also have "... = (\<Sum>v\<in>cart_basis. x $ (THE i. v = axis i 1) * v $ i)" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 83 | unfolding vector_smult_component .. | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 84 |     also have "... = x $ (THE a. ?w = axis a 1) * ?w $ i + (\<Sum>v\<in>(cart_basis) - {?w}. x $ (THE i. v = axis i 1) * v $ i)"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 85 | by (rule sum.remove[OF finite_cart_basis], auto simp add: cart_basis_def) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 86 | also have "... = x $ (THE a. ?w = axis a 1) * ?w $ i" unfolding sum_eq_0 by simp | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 87 | also have "... = x $ i" unfolding the_eq_i unfolding axis_def by auto | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 88 | finally show "x $ i = (\<Sum>v\<in>cart_basis. x $ (THE i. v = axis i 1) *s v) $ i" by simp | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 89 | qed simp | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 90 | qed | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 91 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 92 | (*Some interpretations:*) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 93 | interpretation vec: finite_dimensional_vector_space "( *s)" "cart_basis" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 94 | by (unfold_locales, auto simp add: finite_cart_basis independent_cart_basis span_cart_basis) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 95 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 96 | lemma matrix_vector_mul_linear_gen[intro, simp]: | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 97 | "Vector_Spaces.linear ( *s) ( *s) (( *v) A)" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 98 | by unfold_locales | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 99 | (vector matrix_vector_mult_def sum.distrib algebra_simps)+ | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 100 | |
| 68074 | 101 | lemma span_vec_eq: "vec.span X = span X" | 
| 102 | and dim_vec_eq: "vec.dim X = dim X" | |
| 103 | and dependent_vec_eq: "vec.dependent X = dependent X" | |
| 104 | and subspace_vec_eq: "vec.subspace X = subspace X" | |
| 105 | for X::"(real^'n) set" | |
| 106 | unfolding span_raw_def dim_raw_def dependent_raw_def subspace_raw_def | |
| 107 | by (auto simp: scalar_mult_eq_scaleR) | |
| 108 | ||
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 109 | lemma linear_componentwise: | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 110 | fixes f:: "'a::field ^'m \<Rightarrow> 'a ^ 'n" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 111 | assumes lf: "Vector_Spaces.linear ( *s) ( *s) f" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 112 | shows "(f x)$j = sum (\<lambda>i. (x$i) * (f (axis i 1)$j)) (UNIV :: 'm set)" (is "?lhs = ?rhs") | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 113 | proof - | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 114 | interpret lf: Vector_Spaces.linear "( *s)" "( *s)" f | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 115 | using lf . | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 116 | let ?M = "(UNIV :: 'm set)" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 117 | let ?N = "(UNIV :: 'n set)" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 118 | have fM: "finite ?M" by simp | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 119 | have "?rhs = (sum (\<lambda>i. (x$i) *s (f (axis i 1))) ?M)$j" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 120 | unfolding sum_component by simp | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 121 | then show ?thesis | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 122 | unfolding lf.sum[symmetric] lf.scale[symmetric] | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 123 | unfolding basis_expansion by auto | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 124 | qed | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 125 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 126 | interpretation vec: Vector_Spaces.linear "( *s)" "( *s)" "( *v) A" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 127 | using matrix_vector_mul_linear_gen. | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 128 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 129 | interpretation vec: finite_dimensional_vector_space_pair "( *s)" cart_basis "( *s)" cart_basis .. | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 130 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 131 | lemma matrix_works: | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 132 | assumes lf: "Vector_Spaces.linear ( *s) ( *s) f" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 133 | shows "matrix f *v x = f (x::'a::field ^ 'n)" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 134 | apply (simp add: matrix_def matrix_vector_mult_def vec_eq_iff mult.commute) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 135 | apply clarify | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 136 | apply (rule linear_componentwise[OF lf, symmetric]) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 137 | done | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 138 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 139 | lemma matrix_of_matrix_vector_mul[simp]: "matrix(\<lambda>x. A *v (x :: 'a::field ^ 'n)) = A" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 140 | by (simp add: matrix_eq matrix_works) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 141 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 142 | lemma matrix_compose_gen: | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 143 |   assumes lf: "Vector_Spaces.linear ( *s) ( *s) (f::'a::{field}^'n \<Rightarrow> 'a^'m)"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 144 | and lg: "Vector_Spaces.linear ( *s) ( *s) (g::'a^'m \<Rightarrow> 'a^_)" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 145 | shows "matrix (g o f) = matrix g ** matrix f" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 146 | using lf lg Vector_Spaces.linear_compose[OF lf lg] matrix_works[OF Vector_Spaces.linear_compose[OF lf lg]] | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 147 | by (simp add: matrix_eq matrix_works matrix_vector_mul_assoc[symmetric] o_def) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 148 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 149 | lemma matrix_compose: | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 150 | assumes "linear (f::real^'n \<Rightarrow> real^'m)" "linear (g::real^'m \<Rightarrow> real^_)" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 151 | shows "matrix (g o f) = matrix g ** matrix f" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 152 | using matrix_compose_gen[of f g] assms | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 153 | by (simp add: linear_def scalar_mult_eq_scaleR) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 154 | |
| 68074 | 155 | lemma left_invertible_transpose: | 
| 156 | "(\<exists>(B). B ** transpose (A) = mat (1::'a::comm_semiring_1)) \<longleftrightarrow> (\<exists>(B). A ** B = mat 1)" | |
| 157 | by (metis matrix_transpose_mul transpose_mat transpose_transpose) | |
| 158 | ||
| 159 | lemma right_invertible_transpose: | |
| 160 | "(\<exists>(B). transpose (A) ** B = mat (1::'a::comm_semiring_1)) \<longleftrightarrow> (\<exists>(B). B ** A = mat 1)" | |
| 161 | by (metis matrix_transpose_mul transpose_mat transpose_transpose) | |
| 162 | ||
| 163 | lemma linear_matrix_vector_mul_eq: | |
| 164 | "Vector_Spaces.linear ( *s) ( *s) f \<longleftrightarrow> linear (f :: real^'n \<Rightarrow> real ^'m)" | |
| 165 | by (simp add: scalar_mult_eq_scaleR linear_def) | |
| 166 | ||
| 167 | lemma matrix_vector_mul[simp]: | |
| 168 | "Vector_Spaces.linear ( *s) ( *s) g \<Longrightarrow> (\<lambda>y. matrix g *v y) = g" | |
| 169 | "linear f \<Longrightarrow> (\<lambda>x. matrix f *v x) = f" | |
| 170 | "bounded_linear f \<Longrightarrow> (\<lambda>x. matrix f *v x) = f" | |
| 171 | for f :: "real^'n \<Rightarrow> real ^'m" | |
| 172 | by (simp_all add: ext matrix_works linear_matrix_vector_mul_eq linear_linear) | |
| 173 | ||
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 174 | lemma matrix_left_invertible_injective: | 
| 68074 | 175 | fixes A :: "'a::field^'n^'m" | 
| 176 | shows "(\<exists>B. B ** A = mat 1) \<longleftrightarrow> inj (( *v) A)" | |
| 177 | proof safe | |
| 178 | fix B | |
| 179 | assume B: "B ** A = mat 1" | |
| 180 | show "inj (( *v) A)" | |
| 181 | unfolding inj_on_def | |
| 182 | by (metis B matrix_vector_mul_assoc matrix_vector_mul_lid) | |
| 183 | next | |
| 184 | assume "inj (( *v) A)" | |
| 185 | from vec.linear_injective_left_inverse[OF matrix_vector_mul_linear_gen this] | |
| 186 | obtain g where "Vector_Spaces.linear ( *s) ( *s) g" and g: "g \<circ> ( *v) A = id" | |
| 187 | by blast | |
| 188 | have "matrix g ** A = mat 1" | |
| 189 | by (metis matrix_vector_mul_linear_gen \<open>Vector_Spaces.linear ( *s) ( *s) g\<close> g matrix_compose_gen | |
| 190 | matrix_eq matrix_id_mat_1 matrix_vector_mul(1)) | |
| 191 | then show "\<exists>B. B ** A = mat 1" | |
| 192 | by metis | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 193 | qed | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 194 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 195 | lemma matrix_left_invertible_ker: | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 196 |   "(\<exists>B. (B::'a::{field} ^'m^'n) ** (A::'a::{field}^'n^'m) = mat 1) \<longleftrightarrow> (\<forall>x. A *v x = 0 \<longrightarrow> x = 0)"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 197 | unfolding matrix_left_invertible_injective | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 198 | using vec.inj_on_iff_eq_0[OF vec.subspace_UNIV, of A] | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 199 | by (simp add: inj_on_def) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 200 | |
| 68074 | 201 | lemma matrix_right_invertible_surjective: | 
| 202 | "(\<exists>B. (A::'a::field^'n^'m) ** (B::'a::field^'m^'n) = mat 1) \<longleftrightarrow> surj (\<lambda>x. A *v x)" | |
| 203 | proof - | |
| 204 |   { fix B :: "'a ^'m^'n"
 | |
| 205 | assume AB: "A ** B = mat 1" | |
| 206 |     { fix x :: "'a ^ 'm"
 | |
| 207 | have "A *v (B *v x) = x" | |
| 208 | by (simp add: matrix_vector_mul_assoc AB) } | |
| 209 | hence "surj (( *v) A)" unfolding surj_def by metis } | |
| 210 | moreover | |
| 211 |   { assume sf: "surj (( *v) A)"
 | |
| 212 | from vec.linear_surjective_right_inverse[OF _ this] | |
| 213 | obtain g:: "'a ^'m \<Rightarrow> 'a ^'n" where g: "Vector_Spaces.linear ( *s) ( *s) g" "( *v) A \<circ> g = id" | |
| 214 | by blast | |
| 215 | ||
| 216 | have "A ** (matrix g) = mat 1" | |
| 217 | unfolding matrix_eq matrix_vector_mul_lid | |
| 218 | matrix_vector_mul_assoc[symmetric] matrix_works[OF g(1)] | |
| 219 | using g(2) unfolding o_def fun_eq_iff id_def | |
| 220 | . | |
| 221 | hence "\<exists>B. A ** (B::'a^'m^'n) = mat 1" by blast | |
| 222 | } | |
| 223 | ultimately show ?thesis unfolding surj_def by blast | |
| 224 | qed | |
| 225 | ||
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 226 | lemma matrix_left_invertible_independent_columns: | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 227 |   fixes A :: "'a::{field}^'n^'m"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 228 | shows "(\<exists>(B::'a ^'m^'n). B ** A = mat 1) \<longleftrightarrow> | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 229 | (\<forall>c. sum (\<lambda>i. c i *s column i A) (UNIV :: 'n set) = 0 \<longrightarrow> (\<forall>i. c i = 0))" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 230 | (is "?lhs \<longleftrightarrow> ?rhs") | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 231 | proof - | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 232 | let ?U = "UNIV :: 'n set" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 233 |   { assume k: "\<forall>x. A *v x = 0 \<longrightarrow> x = 0"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 234 |     { fix c i
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 235 | assume c: "sum (\<lambda>i. c i *s column i A) ?U = 0" and i: "i \<in> ?U" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 236 | let ?x = "\<chi> i. c i" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 237 | have th0:"A *v ?x = 0" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 238 | using c | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 239 | by (vector matrix_mult_sum) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 240 | from k[rule_format, OF th0] i | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 241 | have "c i = 0" by (vector vec_eq_iff)} | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 242 | hence ?rhs by blast } | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 243 | moreover | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 244 |   { assume H: ?rhs
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 245 |     { fix x assume x: "A *v x = 0"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 246 | let ?c = "\<lambda>i. ((x$i ):: 'a)" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 247 | from H[rule_format, of ?c, unfolded matrix_mult_sum[symmetric], OF x] | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 248 | have "x = 0" by vector } | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 249 | } | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 250 | ultimately show ?thesis unfolding matrix_left_invertible_ker by auto | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 251 | qed | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 252 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 253 | lemma matrix_right_invertible_independent_rows: | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 254 |   fixes A :: "'a::{field}^'n^'m"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 255 | shows "(\<exists>(B::'a^'m^'n). A ** B = mat 1) \<longleftrightarrow> | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 256 | (\<forall>c. sum (\<lambda>i. c i *s row i A) (UNIV :: 'm set) = 0 \<longrightarrow> (\<forall>i. c i = 0))" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 257 | unfolding left_invertible_transpose[symmetric] | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 258 | matrix_left_invertible_independent_columns | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 259 | by (simp add:) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 260 | |
| 68074 | 261 | lemma matrix_right_invertible_span_columns: | 
| 262 | "(\<exists>(B::'a::field ^'n^'m). (A::'a ^'m^'n) ** B = mat 1) \<longleftrightarrow> | |
| 263 | vec.span (columns A) = UNIV" (is "?lhs = ?rhs") | |
| 264 | proof - | |
| 265 | let ?U = "UNIV :: 'm set" | |
| 266 | have fU: "finite ?U" by simp | |
| 267 | have lhseq: "?lhs \<longleftrightarrow> (\<forall>y. \<exists>(x::'a^'m). sum (\<lambda>i. (x$i) *s column i A) ?U = y)" | |
| 268 | unfolding matrix_right_invertible_surjective matrix_mult_sum surj_def | |
| 269 | by (simp add: eq_commute) | |
| 270 | have rhseq: "?rhs \<longleftrightarrow> (\<forall>x. x \<in> vec.span (columns A))" by blast | |
| 271 |   { assume h: ?lhs
 | |
| 272 |     { fix x:: "'a ^'n"
 | |
| 273 | from h[unfolded lhseq, rule_format, of x] obtain y :: "'a ^'m" | |
| 274 | where y: "sum (\<lambda>i. (y$i) *s column i A) ?U = x" by blast | |
| 275 | have "x \<in> vec.span (columns A)" | |
| 276 | unfolding y[symmetric] scalar_mult_eq_scaleR | |
| 277 | proof (rule vec.span_sum [OF vec.span_scale]) | |
| 278 | show "column i A \<in> vec.span (columns A)" for i | |
| 279 | using columns_def vec.span_superset by auto | |
| 280 | qed | |
| 281 | } | |
| 282 | then have ?rhs unfolding rhseq by blast } | |
| 283 | moreover | |
| 284 |   { assume h:?rhs
 | |
| 285 | let ?P = "\<lambda>(y::'a ^'n). \<exists>(x::'a^'m). sum (\<lambda>i. (x$i) *s column i A) ?U = y" | |
| 286 |     { fix y
 | |
| 287 | have "y \<in> vec.span (columns A)" | |
| 288 | unfolding h by blast | |
| 289 | then have "?P y" | |
| 290 | proof (induction rule: vec.span_induct_alt) | |
| 291 | case base | |
| 292 | then show ?case | |
| 293 | by (metis (full_types) matrix_mult_sum matrix_vector_mult_0_right) | |
| 294 | next | |
| 295 | case (step c y1 y2) | |
| 296 | from step obtain i where i: "i \<in> ?U" "y1 = column i A" | |
| 297 | unfolding columns_def by blast | |
| 298 | obtain x:: "'a ^'m" where x: "sum (\<lambda>i. (x$i) *s column i A) ?U = y2" | |
| 299 | using step by blast | |
| 300 | let ?x = "(\<chi> j. if j = i then c + (x$i) else (x$j))::'a^'m" | |
| 301 | show ?case | |
| 302 | proof (rule exI[where x= "?x"], vector, auto simp add: i x[symmetric] if_distrib distrib_left if_distribR cong del: if_weak_cong) | |
| 303 | fix j | |
| 304 | have th: "\<forall>xa \<in> ?U. (if xa = i then (c + (x$i)) * ((column xa A)$j) | |
| 305 | else (x$xa) * ((column xa A$j))) = (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))" | |
| 306 | using i(1) by (simp add: field_simps) | |
| 307 | have "sum (\<lambda>xa. if xa = i then (c + (x$i)) * ((column xa A)$j) | |
| 308 | else (x$xa) * ((column xa A$j))) ?U = sum (\<lambda>xa. (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))) ?U" | |
| 309 | by (rule sum.cong[OF refl]) (use th in blast) | |
| 310 | also have "\<dots> = sum (\<lambda>xa. if xa = i then c * ((column i A)$j) else 0) ?U + sum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U" | |
| 311 | by (simp add: sum.distrib) | |
| 312 | also have "\<dots> = c * ((column i A)$j) + sum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U" | |
| 313 | unfolding sum.delta[OF fU] | |
| 314 | using i(1) by simp | |
| 315 | finally show "sum (\<lambda>xa. if xa = i then (c + (x$i)) * ((column xa A)$j) | |
| 316 | else (x$xa) * ((column xa A$j))) ?U = c * ((column i A)$j) + sum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U" . | |
| 317 | qed | |
| 318 | qed | |
| 319 | } | |
| 320 | then have ?lhs unfolding lhseq .. | |
| 321 | } | |
| 322 | ultimately show ?thesis by blast | |
| 323 | qed | |
| 324 | ||
| 325 | lemma matrix_left_invertible_span_rows_gen: | |
| 326 | "(\<exists>(B::'a^'m^'n). B ** (A::'a::field^'n^'m) = mat 1) \<longleftrightarrow> vec.span (rows A) = UNIV" | |
| 327 | unfolding right_invertible_transpose[symmetric] | |
| 328 | unfolding columns_transpose[symmetric] | |
| 329 | unfolding matrix_right_invertible_span_columns | |
| 330 | .. | |
| 331 | ||
| 332 | lemma matrix_left_invertible_span_rows: | |
| 333 | "(\<exists>(B::real^'m^'n). B ** (A::real^'n^'m) = mat 1) \<longleftrightarrow> span (rows A) = UNIV" | |
| 334 | using matrix_left_invertible_span_rows_gen[of A] by (simp add: span_vec_eq) | |
| 335 | ||
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 336 | lemma matrix_left_right_inverse: | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 337 |   fixes A A' :: "'a::{field}^'n^'n"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 338 | shows "A ** A' = mat 1 \<longleftrightarrow> A' ** A = mat 1" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 339 | proof - | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 340 |   { fix A A' :: "'a ^'n^'n"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 341 | assume AA': "A ** A' = mat 1" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 342 | have sA: "surj (( *v) A)" | 
| 68073 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 343 | using AA' matrix_right_invertible_surjective by auto | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 344 | from vec.linear_surjective_isomorphism[OF matrix_vector_mul_linear_gen sA] | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 345 | obtain f' :: "'a ^'n \<Rightarrow> 'a ^'n" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 346 | where f': "Vector_Spaces.linear ( *s) ( *s) f'" "\<forall>x. f' (A *v x) = x" "\<forall>x. A *v f' x = x" by blast | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 347 | have th: "matrix f' ** A = mat 1" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 348 | by (simp add: matrix_eq matrix_works[OF f'(1)] | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 349 | matrix_vector_mul_assoc[symmetric] f'(2)[rule_format]) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 350 | hence "(matrix f' ** A) ** A' = mat 1 ** A'" by simp | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 351 | hence "matrix f' = A'" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 352 | by (simp add: matrix_mul_assoc[symmetric] AA') | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 353 | hence "matrix f' ** A = A' ** A" by simp | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 354 | hence "A' ** A = mat 1" by (simp add: th) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 355 | } | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 356 | then show ?thesis by blast | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 357 | qed | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 358 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 359 | lemma invertible_left_inverse: | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 360 |   fixes A :: "'a::{field}^'n^'n"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 361 | shows "invertible A \<longleftrightarrow> (\<exists>(B::'a^'n^'n). B ** A = mat 1)" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 362 | by (metis invertible_def matrix_left_right_inverse) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 363 | |
| 68073 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 364 | lemma invertible_right_inverse: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 365 |   fixes A :: "'a::{field}^'n^'n"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 366 | shows "invertible A \<longleftrightarrow> (\<exists>(B::'a^'n^'n). A** B = mat 1)" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 367 | by (metis invertible_def matrix_left_right_inverse) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 368 | |
| 68073 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 369 | lemma invertible_mult: | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 370 | assumes inv_A: "invertible A" | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 371 | and inv_B: "invertible B" | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 372 | shows "invertible (A**B)" | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 373 | proof - | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 374 | obtain A' where AA': "A ** A' = mat 1" and A'A: "A' ** A = mat 1" | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 375 | using inv_A unfolding invertible_def by blast | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 376 | obtain B' where BB': "B ** B' = mat 1" and B'B: "B' ** B = mat 1" | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 377 | using inv_B unfolding invertible_def by blast | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 378 | show ?thesis | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 379 | proof (unfold invertible_def, rule exI[of _ "B'**A'"], rule conjI) | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 380 | have "A ** B ** (B' ** A') = A ** (B ** (B' ** A'))" | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 381 | using matrix_mul_assoc[of A B "(B' ** A')", symmetric] . | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 382 | also have "... = A ** (B ** B' ** A')" unfolding matrix_mul_assoc[of B "B'" "A'"] .. | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 383 | also have "... = A ** (mat 1 ** A')" unfolding BB' .. | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 384 | also have "... = A ** A'" unfolding matrix_mul_lid .. | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 385 | also have "... = mat 1" unfolding AA' .. | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 386 | finally show "A ** B ** (B' ** A') = mat (1::'a)" . | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 387 | have "B' ** A' ** (A ** B) = B' ** (A' ** (A ** B))" using matrix_mul_assoc[of B' A' "(A ** B)", symmetric] . | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 388 | also have "... = B' ** (A' ** A ** B)" unfolding matrix_mul_assoc[of A' A B] .. | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 389 | also have "... = B' ** (mat 1 ** B)" unfolding A'A .. | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 390 | also have "... = B' ** B" unfolding matrix_mul_lid .. | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 391 | also have "... = mat 1" unfolding B'B .. | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 392 | finally show "B' ** A' ** (A ** B) = mat 1" . | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 393 | qed | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 394 | qed | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 395 | |
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 396 | lemma transpose_invertible: | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 397 | fixes A :: "real^'n^'n" | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 398 | assumes "invertible A" | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 399 | shows "invertible (transpose A)" | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 400 | by (meson assms invertible_def matrix_left_right_inverse right_invertible_transpose) | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 401 | |
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 402 | lemma vector_matrix_mul_assoc: | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 403 |   fixes v :: "('a::comm_semiring_1)^'n"
 | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 404 | shows "(v v* M) v* N = v v* (M ** N)" | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 405 | proof - | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 406 | from matrix_vector_mul_assoc | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 407 | have "transpose N *v (transpose M *v v) = (transpose N ** transpose M) *v v" by fast | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 408 | thus "(v v* M) v* N = v v* (M ** N)" | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 409 | by (simp add: matrix_transpose_mul [symmetric]) | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 410 | qed | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 411 | |
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 412 | lemma matrix_scaleR_vector_ac: | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 413 |   fixes A :: "real^('m::finite)^'n"
 | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 414 | shows "A *v (k *\<^sub>R v) = k *\<^sub>R A *v v" | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 415 | by (metis matrix_vector_mult_scaleR transpose_scalar vector_scaleR_matrix_ac vector_transpose_matrix) | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 416 | |
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 417 | lemma scaleR_matrix_vector_assoc: | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 418 |   fixes A :: "real^('m::finite)^'n"
 | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 419 | shows "k *\<^sub>R (A *v v) = k *\<^sub>R A *v v" | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 420 | by (metis matrix_scaleR_vector_ac matrix_vector_mult_scaleR) | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 421 | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 422 | (*Finally, some interesting theorems and interpretations that don't appear in any file of the | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 423 | library.*) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 424 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 425 | locale linear_first_finite_dimensional_vector_space = | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 426 | l?: Vector_Spaces.linear scaleB scaleC f + | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 427 | B?: finite_dimensional_vector_space scaleB BasisB | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 428 |   for scaleB :: "('a::field => 'b::ab_group_add => 'b)" (infixr "*b" 75)
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 429 |   and scaleC :: "('a => 'c::ab_group_add => 'c)" (infixr "*c" 75)
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 430 |   and BasisB :: "('b set)"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 431 |   and f :: "('b=>'c)"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 432 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 433 | lemma vec_dim_card: "vec.dim (UNIV::('a::{field}^'n) set) = CARD ('n)"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 434 | proof - | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 435 | let ?f="\<lambda>i::'n. axis i (1::'a)" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 436 |   have "vec.dim (UNIV::('a::{field}^'n) set) = card (cart_basis::('a^'n) set)"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 437 | unfolding vec.dim_UNIV .. | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 438 |   also have "... = card ({i. i\<in> UNIV}::('n) set)"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 439 | proof (rule bij_betw_same_card[of ?f, symmetric], unfold bij_betw_def, auto) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 440 | show "inj (\<lambda>i::'n. axis i (1::'a))" by (simp add: inj_on_def axis_eq_axis) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 441 | fix i::'n | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 442 | show "axis i 1 \<in> cart_basis" unfolding cart_basis_def by auto | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 443 | fix x::"'a^'n" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 444 | assume "x \<in> cart_basis" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 445 | thus "x \<in> range (\<lambda>i. axis i 1)" unfolding cart_basis_def by auto | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 446 | qed | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 447 |   also have "... = CARD('n)" by auto
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 448 | finally show ?thesis . | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 449 | qed | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 450 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 451 | interpretation vector_space_over_itself: vector_space "( *) :: 'a::field => 'a => 'a" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 452 | by unfold_locales (simp_all add: algebra_simps) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 453 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 454 | lemmas [simp del] = vector_space_over_itself.scale_scale | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 455 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 456 | interpretation vector_space_over_itself: finite_dimensional_vector_space | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 457 |   "( *) :: 'a::field => 'a => 'a" "{1}"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 458 | by unfold_locales (auto simp: vector_space_over_itself.span_singleton) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 459 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 460 | lemma dimension_eq_1[code_unfold]: "vector_space_over_itself.dimension TYPE('a::field)= 1"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 461 | unfolding vector_space_over_itself.dimension_def by simp | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 462 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 463 | end |