| author | paulson <lp15@cam.ac.uk> | 
| Mon, 21 Aug 2023 18:38:25 +0100 | |
| changeset 78517 | 28c1f4f5335f | 
| parent 78480 | b22f39c54e8c | 
| child 80914 | d97fdabd9e2b | 
| 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 | |
| 69665 | 9 | section "Linear Algebra on Finite Cartesian Products" | 
| 10 | ||
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 11 | theory Cartesian_Space | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 12 | imports | 
| 73648 | 13 | "HOL-Combinatorics.Transposition" | 
| 14 | Finite_Cartesian_Product | |
| 15 | Linear_Algebra | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 16 | begin | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 17 | |
| 70136 | 18 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Type @{typ \<open>'a ^ 'n\<close>} and fields as vector spaces\<close> (*much of the following
 | 
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 19 | is really basic linear algebra, check for overlap? rename subsection? *) | 
| 69667 | 20 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 21 | definition "cart_basis = {axis i 1 | i. i\<in>UNIV}"
 | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 22 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 23 | lemma finite_cart_basis: "finite (cart_basis)" unfolding cart_basis_def | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 24 | 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 | 25 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 26 | lemma card_cart_basis: "card (cart_basis::('a::zero_neq_one^'i) set) = CARD('i)"
 | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 27 | 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 | 28 | 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 | 29 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 30 | interpretation vec: vector_space "(*s) " | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 31 | 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 | 32 | |
| 78480 | 33 | lemma independent_cart_basis: "vec.independent (cart_basis)" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 34 | 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 | 35 | 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 | 36 |   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 | 37 | 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 | 38 | 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 | 39 |   have sum_eq_0: "(\<Sum>x\<in>(cart_basis) - {x}. f x * (x $ i)) = 0"
 | 
| 78480 | 40 | proof (intro sum.neutral ballI) | 
| 41 |     fix y assume y: "y \<in> cart_basis - {x}"
 | |
| 42 | obtain a where a: "y = axis a 1" and a_not_i: "a \<noteq> i" | |
| 43 | using y x unfolding cart_basis_def by auto | |
| 44 | have "y $ i = 0" unfolding a axis_def using a_not_i by auto | |
| 45 | thus "f y * y $ i = 0" by simp | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 46 | qed | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 47 | have "0 = (\<Sum>x\<in>cart_basis. f x *s x) $ i" using eq_0 by simp | 
| 78480 | 48 | also have "\<dots> = (\<Sum>x\<in>cart_basis. (f x *s x) $ i)" unfolding sum_component .. | 
| 49 | also have "\<dots> = (\<Sum>x\<in>cart_basis. f x * (x $ i))" unfolding vector_smult_component .. | |
| 50 |   also have "\<dots> = f x * (x $ i) + (\<Sum>x\<in>(cart_basis) - {x}. f x * (x $ i))"
 | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 51 | by (rule sum.remove[OF finite_cart_basis x_in]) | 
| 78480 | 52 | also have "\<dots> = f x * (x $ i)" unfolding sum_eq_0 by simp | 
| 53 | also have "\<dots> = f x" unfolding x axis_def by auto | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 54 | 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 | 55 | qed | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 56 | |
| 78480 | 57 | lemma span_cart_basis [simp]: "vec.span (cart_basis) = UNIV" | 
| 58 | proof - | |
| 59 |   have "x \<in> vec.span cart_basis" for x :: "('a, 'b) vec"
 | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 60 | proof - | 
| 78480 | 61 | let ?f="\<lambda>v. x $ (THE i. v = axis i 1)" | 
| 62 | have "x $ i = (\<Sum>v\<in>cart_basis. x $ (THE i. v = axis i 1) *s v) $ i" for i::'b | |
| 63 | proof - | |
| 64 | let ?w = "axis i (1::'a)" | |
| 65 | have the_eq_i: "(THE a. ?w = axis a 1) = i" | |
| 66 | by (rule the_equality, auto simp: axis_eq_axis) | |
| 67 |       have sum_eq_0: "(\<Sum>v\<in>(cart_basis) - {?w}. x $ (THE i. v = axis i 1) * v $ i) = 0"
 | |
| 68 | proof (intro sum.neutral ballI) | |
| 69 |         fix y:: "('a, 'b) vec"
 | |
| 70 |         assume y: "y \<in> cart_basis - {?w}"
 | |
| 71 | obtain j where j: "y = axis j 1" and i_not_j: "i \<noteq> j" | |
| 72 | using y unfolding cart_basis_def by auto | |
| 73 | have the_eq_j: "(THE i. y = axis i 1) = j" | |
| 74 | by (simp add: axis_eq_axis j) | |
| 75 | show "x $ (THE i. y = axis i 1) * y $ i = 0" | |
| 76 | by (simp add: axis_def i_not_j j) | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 77 | qed | 
| 78480 | 78 | have "(\<Sum>v\<in>cart_basis. x $ (THE i. v = axis i 1) *s v) $ i | 
| 79 | = (\<Sum>v\<in>cart_basis. x $ (THE i. v = axis i 1) * v $ i)" | |
| 80 | by force | |
| 81 |       also have "\<dots> = x $ (THE a. ?w = axis a 1) * ?w $ i + (\<Sum>v\<in>(cart_basis) - {?w}. x $ (THE i. v = axis i 1) * v $ i)"
 | |
| 82 | by (rule sum.remove[OF finite_cart_basis], auto simp add: cart_basis_def) | |
| 83 | also have "\<dots> = x $ (THE a. ?w = axis a 1) * ?w $ i" | |
| 84 | unfolding sum_eq_0 by simp | |
| 85 | also have "\<dots> = x $ i" | |
| 86 | unfolding the_eq_i unfolding axis_def by auto | |
| 87 | finally show ?thesis by simp | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 88 | qed | 
| 78480 | 89 | then show "x \<in> vec.span (cart_basis)" | 
| 90 | by (metis (no_types, lifting) vec.span_base vec.span_scale vec.span_sum vec_eq_iff) | |
| 91 | qed | |
| 92 | then show ?thesis by auto | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 93 | qed | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 94 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 95 | (*Some interpretations:*) | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 96 | interpretation vec: finite_dimensional_vector_space "(*s)" "cart_basis" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 97 | 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 | 98 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 99 | lemma matrix_vector_mul_linear_gen[intro, simp]: | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 100 | "Vector_Spaces.linear (*s) (*s) ((*v) A)" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 101 | by unfold_locales | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 102 | (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 | 103 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 104 | lemma span_vec_eq: "vec.span X = span X" | 
| 68074 | 105 | and dim_vec_eq: "vec.dim X = dim X" | 
| 106 | and dependent_vec_eq: "vec.dependent X = dependent X" | |
| 107 | and subspace_vec_eq: "vec.subspace X = subspace X" | |
| 108 | for X::"(real^'n) set" | |
| 109 | unfolding span_raw_def dim_raw_def dependent_raw_def subspace_raw_def | |
| 110 | by (auto simp: scalar_mult_eq_scaleR) | |
| 111 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 112 | lemma linear_componentwise: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 113 | fixes f:: "'a::field ^'m \<Rightarrow> 'a ^ 'n" | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 114 | assumes lf: "Vector_Spaces.linear (*s) (*s) f" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 115 | shows "(f x)$j = sum (\<lambda>i. (x$i) * (f (axis i 1)$j)) (UNIV :: 'm set)" (is "?lhs = ?rhs") | 
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 116 | proof - | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 117 | interpret lf: Vector_Spaces.linear "(*s)" "(*s)" f | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 118 | using lf . | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 119 | 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 | 120 | 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 | 121 | 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 | 122 | 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 | 123 | 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 | 124 | then show ?thesis | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 125 | 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 | 126 | 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 | 127 | qed | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 128 | |
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 129 | interpretation vec: Vector_Spaces.linear "(*s)" "(*s)" "(*v) A" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 130 | 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 | 131 | |
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 132 | interpretation vec: finite_dimensional_vector_space_pair "(*s)" cart_basis "(*s)" cart_basis .. | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 133 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 134 | lemma matrix_works: | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 135 | assumes lf: "Vector_Spaces.linear (*s) (*s) f" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 136 | shows "matrix f *v x = f (x::'a::field ^ 'n)" | 
| 78480 | 137 | proof - | 
| 138 | have "\<forall>i. (\<Sum>j\<in>UNIV. x $ j * f (axis j 1) $ i) = f x $ i" | |
| 139 | by (simp add: Cartesian_Space.linear_componentwise lf) | |
| 140 | then show ?thesis | |
| 141 | by (simp add: matrix_def matrix_vector_mult_def vec_eq_iff mult.commute) | |
| 142 | qed | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 143 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 144 | lemma matrix_of_matrix_vector_mul[simp]: "matrix(\<lambda>x. A *v (x :: 'a::field ^ 'n)) = A" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 145 | 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 | 146 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 147 | lemma matrix_compose_gen: | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 148 |   assumes lf: "Vector_Spaces.linear (*s) (*s) (f::'a::{field}^'n \<Rightarrow> 'a^'m)"
 | 
| 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 149 | and lg: "Vector_Spaces.linear (*s) (*s) (g::'a^'m \<Rightarrow> 'a^_)" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 150 | 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 | 151 | 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 | 152 | 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 | 153 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 154 | lemma matrix_compose: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 155 | 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 | 156 | 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 | 157 | 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 | 158 | 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 | 159 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 160 | lemma left_invertible_transpose: | 
| 68074 | 161 | "(\<exists>(B). B ** transpose (A) = mat (1::'a::comm_semiring_1)) \<longleftrightarrow> (\<exists>(B). A ** B = mat 1)" | 
| 162 | by (metis matrix_transpose_mul transpose_mat transpose_transpose) | |
| 163 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 164 | lemma right_invertible_transpose: | 
| 68074 | 165 | "(\<exists>(B). transpose (A) ** B = mat (1::'a::comm_semiring_1)) \<longleftrightarrow> (\<exists>(B). B ** A = mat 1)" | 
| 166 | by (metis matrix_transpose_mul transpose_mat transpose_transpose) | |
| 167 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 168 | lemma linear_matrix_vector_mul_eq: | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 169 | "Vector_Spaces.linear (*s) (*s) f \<longleftrightarrow> linear (f :: real^'n \<Rightarrow> real ^'m)" | 
| 68074 | 170 | by (simp add: scalar_mult_eq_scaleR linear_def) | 
| 171 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 172 | lemma matrix_vector_mul[simp]: | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 173 | "Vector_Spaces.linear (*s) (*s) g \<Longrightarrow> (\<lambda>y. matrix g *v y) = g" | 
| 68074 | 174 | "linear f \<Longrightarrow> (\<lambda>x. matrix f *v x) = f" | 
| 175 | "bounded_linear f \<Longrightarrow> (\<lambda>x. matrix f *v x) = f" | |
| 176 | for f :: "real^'n \<Rightarrow> real ^'m" | |
| 177 | by (simp_all add: ext matrix_works linear_matrix_vector_mul_eq linear_linear) | |
| 178 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 179 | lemma matrix_left_invertible_injective: | 
| 68074 | 180 | fixes A :: "'a::field^'n^'m" | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 181 | shows "(\<exists>B. B ** A = mat 1) \<longleftrightarrow> inj ((*v) A)" | 
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 182 | proof safe | 
| 68074 | 183 | fix B | 
| 184 | assume B: "B ** A = mat 1" | |
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 185 | show "inj ((*v) A)" | 
| 68074 | 186 | unfolding inj_on_def | 
| 187 | by (metis B matrix_vector_mul_assoc matrix_vector_mul_lid) | |
| 188 | next | |
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 189 | assume "inj ((*v) A)" | 
| 68074 | 190 | from vec.linear_injective_left_inverse[OF matrix_vector_mul_linear_gen this] | 
| 78480 | 191 | obtain g where "Vector_Spaces.linear (*s) (*s) g" and "g \<circ> (*v) A = id" | 
| 68074 | 192 | by blast | 
| 78480 | 193 | then have "matrix g ** A = mat 1" | 
| 194 | by (metis matrix_compose_gen matrix_id_mat_1 matrix_of_matrix_vector_mul vec.linear_axioms) | |
| 68074 | 195 | then show "\<exists>B. B ** A = mat 1" | 
| 196 | by metis | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 197 | qed | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 198 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 199 | lemma matrix_left_invertible_ker: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 200 |   "(\<exists>B. (B::'a::{field} ^'m^'n) ** (A::'a::{field}^'n^'m) = mat 1) \<longleftrightarrow> (\<forall>x. A *v x = 0 \<longrightarrow> x = 0)"
 | 
| 78480 | 201 | by (simp add: matrix_left_invertible_injective vec.inj_iff_eq_0) | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 202 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 203 | lemma matrix_right_invertible_surjective: | 
| 68074 | 204 | "(\<exists>B. (A::'a::field^'n^'m) ** (B::'a::field^'m^'n) = mat 1) \<longleftrightarrow> surj (\<lambda>x. A *v x)" | 
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 205 | proof - | 
| 78480 | 206 | have "\<And>B x. A ** B = mat 1 \<Longrightarrow> \<exists>y. x = A *v y" | 
| 207 | by (metis matrix_vector_mul_assoc matrix_vector_mul_lid) | |
| 208 | moreover have "\<forall>x. \<exists>xa. x = A *v xa \<Longrightarrow> \<exists>B. A ** B = mat 1" | |
| 209 | by (metis (mono_tags, lifting) matrix_compose_gen matrix_id_mat_1 matrix_of_matrix_vector_mul surj_def vec.linear_axioms vec.linear_surjective_right_inverse) | |
| 210 | ultimately show ?thesis | |
| 211 | by (auto simp: image_def set_eq_iff) | |
| 68074 | 212 | qed | 
| 213 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 214 | lemma matrix_left_invertible_independent_columns: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 215 |   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 | 216 | 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 | 217 | (\<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 | 218 | (is "?lhs \<longleftrightarrow> ?rhs") | 
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 219 | proof - | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 220 | let ?U = "UNIV :: 'n set" | 
| 78480 | 221 | have "c i = 0" | 
| 222 | if "\<forall>x. A *v x = 0 \<longrightarrow> x = 0" "sum (\<lambda>i. c i *s column i A) ?U = 0" for c i | |
| 223 | by (metis (no_types) UNIV_I matrix_mult_sum vec_lambda_eta vec_nth_cases zero_vec_def that) | |
| 224 | moreover have "x = 0" if "A *v x = 0" ?rhs for x | |
| 225 | by (metis (full_types) matrix_mult_sum that vec_eq_iff zero_index) | |
| 226 | ultimately show ?thesis | |
| 227 | unfolding matrix_left_invertible_ker by auto | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 228 | qed | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 229 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 230 | lemma matrix_right_invertible_independent_rows: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 231 |   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 | 232 | shows "(\<exists>(B::'a^'m^'n). A ** B = mat 1) \<longleftrightarrow> | 
| 78480 | 233 | (\<forall>c. sum (\<lambda>i::'m. c i *s row i A) UNIV = 0 \<longrightarrow> (\<forall>i. c i = 0))" | 
| 234 | by (simp add: matrix_left_invertible_independent_columns flip: left_invertible_transpose) | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 235 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 236 | lemma matrix_right_invertible_span_columns: | 
| 68074 | 237 | "(\<exists>(B::'a::field ^'n^'m). (A::'a ^'m^'n) ** B = mat 1) \<longleftrightarrow> | 
| 238 | vec.span (columns A) = UNIV" (is "?lhs = ?rhs") | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 239 | proof - | 
| 68074 | 240 | let ?U = "UNIV :: 'm set" | 
| 241 | have fU: "finite ?U" by simp | |
| 242 | have lhseq: "?lhs \<longleftrightarrow> (\<forall>y. \<exists>(x::'a^'m). sum (\<lambda>i. (x$i) *s column i A) ?U = y)" | |
| 243 | unfolding matrix_right_invertible_surjective matrix_mult_sum surj_def | |
| 244 | by (simp add: eq_commute) | |
| 245 | have rhseq: "?rhs \<longleftrightarrow> (\<forall>x. x \<in> vec.span (columns A))" by blast | |
| 246 |   { assume h: ?lhs
 | |
| 247 |     { fix x:: "'a ^'n"
 | |
| 78480 | 248 | obtain y :: "'a ^'m" where y: "sum (\<lambda>i. (y$i) *s column i A) ?U = x" | 
| 249 | using h lhseq by blast | |
| 250 | then have "x \<in> vec.span (columns A)" | |
| 251 | by (metis (mono_tags, lifting) columns_def mem_Collect_eq vec.span_base vec.span_scale vec.span_sum) | |
| 68074 | 252 | } | 
| 253 | then have ?rhs unfolding rhseq by blast } | |
| 254 | moreover | |
| 255 |   { assume h:?rhs
 | |
| 256 | let ?P = "\<lambda>(y::'a ^'n). \<exists>(x::'a^'m). sum (\<lambda>i. (x$i) *s column i A) ?U = y" | |
| 257 |     { fix y
 | |
| 258 | have "y \<in> vec.span (columns A)" | |
| 259 | unfolding h by blast | |
| 260 | then have "?P y" | |
| 261 | proof (induction rule: vec.span_induct_alt) | |
| 262 | case base | |
| 263 | then show ?case | |
| 264 | by (metis (full_types) matrix_mult_sum matrix_vector_mult_0_right) | |
| 265 | next | |
| 266 | case (step c y1 y2) | |
| 267 | from step obtain i where i: "i \<in> ?U" "y1 = column i A" | |
| 268 | unfolding columns_def by blast | |
| 269 | obtain x:: "'a ^'m" where x: "sum (\<lambda>i. (x$i) *s column i A) ?U = y2" | |
| 270 | using step by blast | |
| 271 | let ?x = "(\<chi> j. if j = i then c + (x$i) else (x$j))::'a^'m" | |
| 272 | show ?case | |
| 273 | proof (rule exI[where x= "?x"], vector, auto simp add: i x[symmetric] if_distrib distrib_left if_distribR cong del: if_weak_cong) | |
| 274 | fix j | |
| 275 | have th: "\<forall>xa \<in> ?U. (if xa = i then (c + (x$i)) * ((column xa A)$j) | |
| 276 | else (x$xa) * ((column xa A$j))) = (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))" | |
| 277 | using i(1) by (simp add: field_simps) | |
| 278 | have "sum (\<lambda>xa. if xa = i then (c + (x$i)) * ((column xa A)$j) | |
| 279 | 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" | |
| 78480 | 280 | using th by force | 
| 68074 | 281 | 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" | 
| 282 | by (simp add: sum.distrib) | |
| 283 | also have "\<dots> = c * ((column i A)$j) + sum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U" | |
| 78480 | 284 | unfolding sum.delta[OF fU] using i(1) by simp | 
| 68074 | 285 | finally show "sum (\<lambda>xa. if xa = i then (c + (x$i)) * ((column xa A)$j) | 
| 78480 | 286 | else (x$xa) * ((column xa A$j))) ?U | 
| 287 | = c * ((column i A)$j) + sum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U" . | |
| 68074 | 288 | qed | 
| 289 | qed | |
| 290 | } | |
| 291 | then have ?lhs unfolding lhseq .. | |
| 292 | } | |
| 293 | ultimately show ?thesis by blast | |
| 294 | qed | |
| 295 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 296 | lemma matrix_left_invertible_span_rows_gen: | 
| 68074 | 297 | "(\<exists>(B::'a^'m^'n). B ** (A::'a::field^'n^'m) = mat 1) \<longleftrightarrow> vec.span (rows A) = UNIV" | 
| 78480 | 298 | by (metis columns_transpose matrix_right_invertible_span_columns right_invertible_transpose) | 
| 68074 | 299 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 300 | lemma matrix_left_invertible_span_rows: | 
| 68074 | 301 | "(\<exists>(B::real^'m^'n). B ** (A::real^'n^'m) = mat 1) \<longleftrightarrow> span (rows A) = UNIV" | 
| 302 | using matrix_left_invertible_span_rows_gen[of A] by (simp add: span_vec_eq) | |
| 303 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 304 | lemma matrix_left_right_inverse: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 305 |   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 | 306 | shows "A ** A' = mat 1 \<longleftrightarrow> A' ** A = mat 1" | 
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 307 | proof - | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 308 |   { 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 | 309 | assume AA': "A ** A' = mat 1" | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 310 | 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 | 311 | 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 | 312 | obtain f' :: "'a ^'n \<Rightarrow> 'a ^'n" | 
| 78480 | 313 | where f': "Vector_Spaces.linear (*s) (*s) f'" "\<forall>x. f' (A *v x) = x" "\<forall>x. A *v f' x = x" | 
| 314 | using sA vec.linear_surjective_isomorphism by blast | |
| 315 | have "matrix f' ** A = mat 1" | |
| 316 | by (metis f' matrix_eq matrix_vector_mul_assoc matrix_vector_mul_lid matrix_works) | |
| 317 | hence "A' ** A = mat 1" | |
| 318 | by (metis AA' matrix_mul_assoc matrix_mul_lid) | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 319 | } | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 320 | 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 | 321 | qed | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 322 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 323 | lemma invertible_left_inverse: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 324 |   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 | 325 | 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 | 326 | 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 | 327 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 328 | 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 | 329 |   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 | 330 | 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 | 331 | 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 | 332 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 333 | lemma invertible_mult: | 
| 68073 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 334 | 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 | 335 | 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 | 336 | shows "invertible (A**B)" | 
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 337 | proof - | 
| 68073 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 338 | 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 | 339 | 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 | 340 | 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 | 341 | using inv_B unfolding invertible_def by blast | 
| 78480 | 342 | have "A ** B ** (B' ** A') = mat 1" | 
| 343 | by (metis AA' BB' matrix_mul_assoc matrix_mul_rid) | |
| 344 | moreover have "B' ** A' ** (A ** B) = mat 1" | |
| 345 | by (metis A'A B'B matrix_mul_assoc matrix_mul_rid) | |
| 346 | ultimately show ?thesis | |
| 347 | using invertible_def by blast | |
| 68073 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 348 | qed | 
| 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 349 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 350 | lemma transpose_invertible: | 
| 68073 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 351 | 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 | 352 | 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 | 353 | 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 | 354 | 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 | 355 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 356 | lemma vector_matrix_mul_assoc: | 
| 68073 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 357 |   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 | 358 | shows "(v v* M) v* N = v v* (M ** N)" | 
| 78480 | 359 | by (metis (no_types, opaque_lifting) matrix_transpose_mul matrix_vector_mul_assoc transpose_matrix_vector) | 
| 68073 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 360 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 361 | lemma matrix_scaleR_vector_ac: | 
| 68073 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 362 |   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 | 363 | 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 | 364 | 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 | 365 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 366 | lemma scaleR_matrix_vector_assoc: | 
| 68073 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immler parents: 
68072diff
changeset | 367 |   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 | 368 | 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 | 369 | 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 | 370 | |
| 78480 | 371 | subsection \<open>Some interesting theorems and interpretations\<close> | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 372 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 373 | 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 | 374 | 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 | 375 | 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 | 376 |   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 | 377 |   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 | 378 |   and BasisB :: "('b set)"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 379 |   and f :: "('b=>'c)"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 380 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 381 | lemma vec_dim_card: "vec.dim (UNIV::('a::{field}^'n) set) = CARD ('n)"
 | 
| 78480 | 382 | by (simp add: card_cart_basis) | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 383 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 384 | interpretation vector_space_over_itself: vector_space "(*) :: 'a::field \<Rightarrow> 'a \<Rightarrow> 'a" | 
| 78480 | 385 | by unfold_locales (simp_all add: algebra_simps) | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 386 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 387 | 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 | 388 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 389 | interpretation vector_space_over_itself: finite_dimensional_vector_space | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 390 |   "(*) :: 'a::field => 'a => 'a" "{1}"
 | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 391 | 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 | 392 | |
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 393 | 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 | 394 | 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 | 395 | |
| 69666 | 396 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 397 | lemma dim_subset_UNIV_cart_gen: | 
| 69666 | 398 |   fixes S :: "('a::field^'n) set"
 | 
| 399 |   shows "vec.dim S \<le> CARD('n)"
 | |
| 400 | by (metis vec.dim_eq_full vec.dim_subset_UNIV vec.span_UNIV vec_dim_card) | |
| 401 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 402 | lemma dim_subset_UNIV_cart: | 
| 69666 | 403 | fixes S :: "(real^'n) set" | 
| 404 |   shows "dim S \<le> CARD('n)"
 | |
| 405 | using dim_subset_UNIV_cart_gen[of S] by (simp add: dim_vec_eq) | |
| 406 | ||
| 407 | text\<open>Two sometimes fruitful ways of looking at matrix-vector multiplication.\<close> | |
| 408 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 409 | lemma matrix_mult_dot: "A *v x = (\<chi> i. inner (A$i) x)" | 
| 69666 | 410 | by (simp add: matrix_vector_mult_def inner_vec_def) | 
| 411 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 412 | lemma adjoint_matrix: "adjoint(\<lambda>x. (A::real^'n^'m) *v x) = (\<lambda>x. transpose A *v x)" | 
| 78480 | 413 | by (metis adjoint_unique dot_lmul_matrix vector_transpose_matrix) | 
| 69666 | 414 | |
| 78480 | 415 | lemma matrix_adjoint: | 
| 416 | assumes lf: "linear (f :: real^'n \<Rightarrow> real ^'m)" | |
| 69666 | 417 | shows "matrix(adjoint f) = transpose(matrix f)" | 
| 78480 | 418 | by (metis adjoint_matrix assms matrix_of_matrix_vector_mul matrix_vector_mul(2)) | 
| 69666 | 419 | |
| 420 | ||
| 69683 | 421 | subsection\<open> Rank of a matrix\<close> | 
| 69666 | 422 | |
| 423 | text\<open>Equivalence of row and column rank is taken from George Mackiw's paper, Mathematics Magazine 1995, p. 285.\<close> | |
| 424 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 425 | lemma matrix_vector_mult_in_columnspace_gen: | 
| 69666 | 426 | fixes A :: "'a::field^'n^'m" | 
| 427 | shows "(A *v x) \<in> vec.span(columns A)" | |
| 78480 | 428 | unfolding columns_def | 
| 429 | by (metis (mono_tags, lifting) matrix_mult_sum mem_Collect_eq vec.span_base vec.span_scale vec.span_sum) | |
| 69666 | 430 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 431 | lemma matrix_vector_mult_in_columnspace: | 
| 69666 | 432 | fixes A :: "real^'n^'m" | 
| 433 | shows "(A *v x) \<in> span(columns A)" | |
| 434 | using matrix_vector_mult_in_columnspace_gen[of A x] by (simp add: span_vec_eq) | |
| 435 | ||
| 436 | lemma subspace_orthogonal_to_vector: "subspace {y. orthogonal x y}"
 | |
| 437 | by (simp add: subspace_def orthogonal_clauses) | |
| 438 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 439 | lemma orthogonal_nullspace_rowspace: | 
| 69666 | 440 | fixes A :: "real^'n^'m" | 
| 441 | assumes 0: "A *v x = 0" and y: "y \<in> span(rows A)" | |
| 442 | shows "orthogonal x y" | |
| 443 | using y | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 444 | proof (induction rule: span_induct) | 
| 69666 | 445 | case base | 
| 446 | then show ?case | |
| 447 | by (simp add: subspace_orthogonal_to_vector) | |
| 448 | next | |
| 449 | case (step v) | |
| 450 | then obtain i where "v = row i A" | |
| 451 | by (auto simp: rows_def) | |
| 452 | with 0 show ?case | |
| 453 | unfolding orthogonal_def inner_vec_def matrix_vector_mult_def row_def | |
| 454 | by (simp add: mult.commute) (metis (no_types) vec_lambda_beta zero_index) | |
| 455 | qed | |
| 456 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 457 | lemma nullspace_inter_rowspace: | 
| 69666 | 458 | fixes A :: "real^'n^'m" | 
| 459 | shows "A *v x = 0 \<and> x \<in> span(rows A) \<longleftrightarrow> x = 0" | |
| 460 | using orthogonal_nullspace_rowspace orthogonal_self span_zero matrix_vector_mult_0_right | |
| 461 | by blast | |
| 462 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 463 | lemma matrix_vector_mul_injective_on_rowspace: | 
| 69666 | 464 | fixes A :: "real^'n^'m" | 
| 465 | shows "\<lbrakk>A *v x = A *v y; x \<in> span(rows A); y \<in> span(rows A)\<rbrakk> \<Longrightarrow> x = y" | |
| 466 | using nullspace_inter_rowspace [of A "x-y"] | |
| 467 | by (metis diff_eq_diff_eq diff_self matrix_vector_mult_diff_distrib span_diff) | |
| 468 | ||
| 70136 | 469 | definition\<^marker>\<open>tag important\<close> rank :: "'a::field^'n^'m=>nat" | 
| 69666 | 470 | where row_rank_def_gen: "rank A \<equiv> vec.dim(rows A)" | 
| 471 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 472 | lemma row_rank_def: "rank A = dim (rows A)" for A::"real^'n^'m" | 
| 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 473 | by (auto simp: row_rank_def_gen dim_vec_eq) | 
| 69666 | 474 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 475 | lemma dim_rows_le_dim_columns: | 
| 69666 | 476 | fixes A :: "real^'n^'m" | 
| 477 | shows "dim(rows A) \<le> dim(columns A)" | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 478 | proof - | 
| 69666 | 479 | have "dim (span (rows A)) \<le> dim (span (columns A))" | 
| 480 | proof - | |
| 481 | obtain B where "independent B" "span(rows A) \<subseteq> span B" | |
| 482 | and B: "B \<subseteq> span(rows A)""card B = dim (span(rows A))" | |
| 483 | using basis_exists [of "span(rows A)"] by metis | |
| 484 | with span_subspace have eq: "span B = span(rows A)" | |
| 485 | by auto | |
| 486 | then have inj: "inj_on ((*v) A) (span B)" | |
| 487 | by (simp add: inj_on_def matrix_vector_mul_injective_on_rowspace) | |
| 488 | then have ind: "independent ((*v) A ` B)" | |
| 489 | by (rule linear_independent_injective_image [OF Finite_Cartesian_Product.matrix_vector_mul_linear \<open>independent B\<close>]) | |
| 490 | have "dim (span (rows A)) \<le> card ((*v) A ` B)" | |
| 78480 | 491 | by (metis B(2) card_image inj inj_on_subset order.refl span_superset) | 
| 69666 | 492 | also have "\<dots> \<le> dim (span (columns A))" | 
| 493 | using _ ind | |
| 494 | by (rule independent_card_le_dim) (auto intro!: matrix_vector_mult_in_columnspace) | |
| 495 | finally show ?thesis . | |
| 496 | qed | |
| 497 | then show ?thesis | |
| 71044 | 498 | by (simp) | 
| 69666 | 499 | qed | 
| 500 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 501 | lemma column_rank_def: | 
| 69666 | 502 | fixes A :: "real^'n^'m" | 
| 503 | shows "rank A = dim(columns A)" | |
| 504 | unfolding row_rank_def | |
| 505 | by (metis columns_transpose dim_rows_le_dim_columns le_antisym rows_transpose) | |
| 506 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 507 | lemma rank_transpose: | 
| 69666 | 508 | fixes A :: "real^'n^'m" | 
| 509 | shows "rank(transpose A) = rank A" | |
| 510 | by (metis column_rank_def row_rank_def rows_transpose) | |
| 511 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 512 | lemma matrix_vector_mult_basis: | 
| 69666 | 513 | fixes A :: "real^'n^'m" | 
| 514 | shows "A *v (axis k 1) = column k A" | |
| 515 | by (simp add: cart_eq_inner_axis column_def matrix_mult_dot) | |
| 516 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 517 | lemma columns_image_basis: | 
| 69666 | 518 | fixes A :: "real^'n^'m" | 
| 519 | shows "columns A = (*v) A ` (range (\<lambda>i. axis i 1))" | |
| 520 | by (force simp: columns_def matrix_vector_mult_basis [symmetric]) | |
| 521 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 522 | lemma rank_dim_range: | 
| 69666 | 523 | fixes A :: "real^'n^'m" | 
| 524 | shows "rank A = dim(range (\<lambda>x. A *v x))" | |
| 525 | unfolding column_rank_def | |
| 78480 | 526 | by (smt (verit, best) columns_image_basis dim_span image_subset_iff iso_tuple_UNIV_I matrix_vector_mult_in_columnspace span_eq) | 
| 69666 | 527 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 528 | lemma rank_bound: | 
| 69666 | 529 | fixes A :: "real^'n^'m" | 
| 530 |   shows "rank A \<le> min CARD('m) (CARD('n))"
 | |
| 531 | by (metis (mono_tags, lifting) dim_subset_UNIV_cart min.bounded_iff | |
| 532 | column_rank_def row_rank_def) | |
| 533 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 534 | lemma full_rank_injective: | 
| 69666 | 535 | fixes A :: "real^'n^'m" | 
| 536 |   shows "rank A = CARD('n) \<longleftrightarrow> inj ((*v) A)"
 | |
| 537 | by (simp add: matrix_left_invertible_injective [symmetric] matrix_left_invertible_span_rows row_rank_def | |
| 538 | dim_eq_full [symmetric] card_cart_basis vec.dimension_def) | |
| 539 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 540 | lemma full_rank_surjective: | 
| 69666 | 541 | fixes A :: "real^'n^'m" | 
| 542 |   shows "rank A = CARD('m) \<longleftrightarrow> surj ((*v) A)"
 | |
| 78480 | 543 | by (metis (no_types, opaque_lifting) dim_eq_full dim_vec_eq rank_dim_range span_vec_eq vec.span_UNIV vec.span_image vec_dim_card) | 
| 69666 | 544 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 545 | lemma rank_I: "rank(mat 1::real^'n^'n) = CARD('n)"
 | 
| 69666 | 546 | by (simp add: full_rank_injective inj_on_def) | 
| 547 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 548 | lemma less_rank_noninjective: | 
| 69666 | 549 | fixes A :: "real^'n^'m" | 
| 550 |   shows "rank A < CARD('n) \<longleftrightarrow> \<not> inj ((*v) A)"
 | |
| 78480 | 551 | using less_le rank_bound by (auto simp: full_rank_injective [symmetric]) | 
| 69666 | 552 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 553 | lemma matrix_nonfull_linear_equations_eq: | 
| 69666 | 554 | fixes A :: "real^'n^'m" | 
| 555 |   shows "(\<exists>x. (x \<noteq> 0) \<and> A *v x = 0) \<longleftrightarrow> rank A \<noteq> CARD('n)"
 | |
| 556 | by (meson matrix_left_invertible_injective full_rank_injective matrix_left_invertible_ker) | |
| 557 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 558 | lemma rank_eq_0: "rank A = 0 \<longleftrightarrow> A = 0" and rank_0 [simp]: "rank (0::real^'n^'m) = 0" | 
| 69666 | 559 | for A :: "real^'n^'m" | 
| 560 | by (auto simp: rank_dim_range matrix_eq) | |
| 561 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 562 | lemma rank_mul_le_right: | 
| 69666 | 563 | fixes A :: "real^'n^'m" and B :: "real^'p^'n" | 
| 564 | shows "rank(A ** B) \<le> rank B" | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 565 | proof - | 
| 69666 | 566 | have "rank(A ** B) \<le> dim ((*v) A ` range ((*v) B))" | 
| 567 | by (auto simp: rank_dim_range image_comp o_def matrix_vector_mul_assoc) | |
| 568 | also have "\<dots> \<le> rank B" | |
| 569 | by (simp add: rank_dim_range dim_image_le) | |
| 570 | finally show ?thesis . | |
| 571 | qed | |
| 572 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 573 | lemma rank_mul_le_left: | 
| 69666 | 574 | fixes A :: "real^'n^'m" and B :: "real^'p^'n" | 
| 575 | shows "rank(A ** B) \<le> rank A" | |
| 576 | by (metis matrix_transpose_mul rank_mul_le_right rank_transpose) | |
| 577 | ||
| 69669 | 578 | |
| 70136 | 579 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Lemmas for working on \<open>real^1/2/3/4\<close>\<close> | 
| 69669 | 580 | |
| 581 | lemma exhaust_2: | |
| 582 | fixes x :: 2 | |
| 583 | shows "x = 1 \<or> x = 2" | |
| 584 | proof (induct x) | |
| 585 | case (of_int z) | |
| 78480 | 586 | then have "z = 0 | z = 1" | 
| 587 | by fastforce | |
| 588 | then show ?case | |
| 589 | by auto | |
| 69669 | 590 | qed | 
| 591 | ||
| 592 | lemma forall_2: "(\<forall>i::2. P i) \<longleftrightarrow> P 1 \<and> P 2" | |
| 593 | by (metis exhaust_2) | |
| 594 | ||
| 595 | lemma exhaust_3: | |
| 596 | fixes x :: 3 | |
| 597 | shows "x = 1 \<or> x = 2 \<or> x = 3" | |
| 598 | proof (induct x) | |
| 599 | case (of_int z) | |
| 78480 | 600 | then have "z = 0 \<or> z = 1 \<or> z = 2" by fastforce | 
| 69669 | 601 | then show ?case by auto | 
| 602 | qed | |
| 603 | ||
| 604 | lemma forall_3: "(\<forall>i::3. P i) \<longleftrightarrow> P 1 \<and> P 2 \<and> P 3" | |
| 605 | by (metis exhaust_3) | |
| 606 | ||
| 69918 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69723diff
changeset | 607 | lemma exhaust_4: | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69723diff
changeset | 608 | fixes x :: 4 | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69723diff
changeset | 609 | shows "x = 1 \<or> x = 2 \<or> x = 3 \<or> x = 4" | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69723diff
changeset | 610 | proof (induct x) | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69723diff
changeset | 611 | case (of_int z) | 
| 78480 | 612 | then have "z = 0 \<or> z = 1 \<or> z = 2 \<or> z = 3" by fastforce | 
| 69918 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69723diff
changeset | 613 | then show ?case by auto | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69723diff
changeset | 614 | qed | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69723diff
changeset | 615 | |
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69723diff
changeset | 616 | lemma forall_4: "(\<forall>i::4. P i) \<longleftrightarrow> P 1 \<and> P 2 \<and> P 3 \<and> P 4" | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69723diff
changeset | 617 | by (metis exhaust_4) | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69723diff
changeset | 618 | |
| 69669 | 619 | lemma UNIV_1 [simp]: "UNIV = {1::1}"
 | 
| 620 | by (auto simp add: num1_eq_iff) | |
| 621 | ||
| 622 | lemma UNIV_2: "UNIV = {1::2, 2::2}"
 | |
| 623 | using exhaust_2 by auto | |
| 624 | ||
| 625 | lemma UNIV_3: "UNIV = {1::3, 2::3, 3::3}"
 | |
| 626 | using exhaust_3 by auto | |
| 627 | ||
| 69918 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69723diff
changeset | 628 | lemma UNIV_4: "UNIV = {1::4, 2::4, 3::4, 4::4}"
 | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69723diff
changeset | 629 | using exhaust_4 by auto | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69723diff
changeset | 630 | |
| 69669 | 631 | lemma sum_1: "sum f (UNIV::1 set) = f 1" | 
| 632 | unfolding UNIV_1 by simp | |
| 633 | ||
| 634 | lemma sum_2: "sum f (UNIV::2 set) = f 1 + f 2" | |
| 635 | unfolding UNIV_2 by simp | |
| 636 | ||
| 637 | lemma sum_3: "sum f (UNIV::3 set) = f 1 + f 2 + f 3" | |
| 638 | unfolding UNIV_3 by (simp add: ac_simps) | |
| 639 | ||
| 69918 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69723diff
changeset | 640 | lemma sum_4: "sum f (UNIV::4 set) = f 1 + f 2 + f 3 + f 4" | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69723diff
changeset | 641 | unfolding UNIV_4 by (simp add: ac_simps) | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69723diff
changeset | 642 | |
| 70136 | 643 | subsection\<^marker>\<open>tag unimportant\<close>\<open>The collapse of the general concepts to dimension one\<close> | 
| 69669 | 644 | |
| 645 | lemma vector_one: "(x::'a ^1) = (\<chi> i. (x$1))" | |
| 646 | by (simp add: vec_eq_iff) | |
| 647 | ||
| 648 | lemma forall_one: "(\<forall>(x::'a ^1). P x) \<longleftrightarrow> (\<forall>x. P(\<chi> i. x))" | |
| 78480 | 649 | by (metis vector_one) | 
| 69669 | 650 | |
| 651 | lemma norm_vector_1: "norm (x :: _^1) = norm (x$1)" | |
| 652 | by (simp add: norm_vec_def) | |
| 653 | ||
| 654 | lemma dist_vector_1: | |
| 655 | fixes x :: "'a::real_normed_vector^1" | |
| 656 | shows "dist x y = dist (x$1) (y$1)" | |
| 657 | by (simp add: dist_norm norm_vector_1) | |
| 658 | ||
| 659 | lemma norm_real: "norm(x::real ^ 1) = \<bar>x$1\<bar>" | |
| 660 | by (simp add: norm_vector_1) | |
| 661 | ||
| 662 | lemma dist_real: "dist(x::real ^ 1) y = \<bar>(x$1) - (y$1)\<bar>" | |
| 663 | by (auto simp add: norm_real dist_norm) | |
| 664 | ||
| 665 | ||
| 70136 | 666 | subsection\<^marker>\<open>tag unimportant\<close>\<open>Routine results connecting the types \<^typ>\<open>real^1\<close> and \<^typ>\<open>real\<close>\<close> | 
| 69669 | 667 | |
| 668 | lemma vector_one_nth [simp]: | |
| 669 | fixes x :: "'a^1" shows "vec (x $ 1) = x" | |
| 670 | by (metis vec_def vector_one) | |
| 671 | ||
| 672 | lemma tendsto_at_within_vector_1: | |
| 673 | fixes S :: "'a :: metric_space set" | |
| 674 | assumes "(f \<longlongrightarrow> fx) (at x within S)" | |
| 675 | shows "((\<lambda>y::'a^1. \<chi> i. f (y $ 1)) \<longlongrightarrow> (vec fx::'a^1)) (at (vec x) within vec ` S)" | |
| 676 | proof (rule topological_tendstoI) | |
| 677 |   fix T :: "('a^1) set"
 | |
| 678 | assume "open T" "vec fx \<in> T" | |
| 679 | have "\<forall>\<^sub>F x in at x within S. f x \<in> (\<lambda>x. x $ 1) ` T" | |
| 680 | using \<open>open T\<close> \<open>vec fx \<in> T\<close> assms open_image_vec_nth tendsto_def by fastforce | |
| 681 | then show "\<forall>\<^sub>F x::'a^1 in at (vec x) within vec ` S. (\<chi> i. f (x $ 1)) \<in> T" | |
| 682 | unfolding eventually_at dist_norm [symmetric] | |
| 683 | by (rule ex_forward) | |
| 684 | (use \<open>open T\<close> in | |
| 685 | \<open>fastforce simp: dist_norm dist_vec_def L2_set_def image_iff vector_one open_vec_def\<close>) | |
| 686 | qed | |
| 687 | ||
| 688 | lemma has_derivative_vector_1: | |
| 689 | assumes der_g: "(g has_derivative (\<lambda>x. x * g' a)) (at a within S)" | |
| 690 | shows "((\<lambda>x. vec (g (x $ 1))) has_derivative (*\<^sub>R) (g' a)) | |
| 691 | (at ((vec a)::real^1) within vec ` S)" | |
| 692 | using der_g | |
| 78480 | 693 | apply (clarsimp simp: Deriv.has_derivative_within bounded_linear_scaleR_right norm_vector_1) | 
| 69669 | 694 | apply (drule tendsto_at_within_vector_1, vector) | 
| 695 | apply (auto simp: algebra_simps eventually_at tendsto_def) | |
| 696 | done | |
| 697 | ||
| 698 | ||
| 70136 | 699 | subsection\<^marker>\<open>tag unimportant\<close>\<open>Explicit vector construction from lists\<close> | 
| 69669 | 700 | |
| 701 | definition "vector l = (\<chi> i. foldr (\<lambda>x f n. fun_upd (f (n+1)) n x) l (\<lambda>n x. 0) 1 i)" | |
| 702 | ||
| 703 | lemma vector_1 [simp]: "(vector[x]) $1 = x" | |
| 704 | unfolding vector_def by simp | |
| 705 | ||
| 706 | lemma vector_2 [simp]: "(vector[x,y]) $1 = x" "(vector[x,y] :: 'a^2)$2 = (y::'a::zero)" | |
| 707 | unfolding vector_def by simp_all | |
| 708 | ||
| 709 | lemma vector_3 [simp]: | |
| 710 |  "(vector [x,y,z] ::('a::zero)^3)$1 = x"
 | |
| 711 |  "(vector [x,y,z] ::('a::zero)^3)$2 = y"
 | |
| 712 |  "(vector [x,y,z] ::('a::zero)^3)$3 = z"
 | |
| 713 | unfolding vector_def by simp_all | |
| 714 | ||
| 715 | lemma forall_vector_1: "(\<forall>v::'a::zero^1. P v) \<longleftrightarrow> (\<forall>x. P(vector[x]))" | |
| 716 | by (metis vector_1 vector_one) | |
| 717 | ||
| 718 | lemma forall_vector_2: "(\<forall>v::'a::zero^2. P v) \<longleftrightarrow> (\<forall>x y. P(vector[x, y]))" | |
| 78480 | 719 | proof - | 
| 720 | have "P v" if "\<And>x y. P (vector [x, y])" for v | |
| 721 | proof - | |
| 722 | have "vector [v$1, v$2] = v" | |
| 723 | by (smt (verit, best) exhaust_2 vec_eq_iff vector_2) | |
| 724 | then show ?thesis | |
| 725 | by (metis that) | |
| 726 | qed | |
| 727 | then show ?thesis by auto | |
| 728 | qed | |
| 69669 | 729 | |
| 730 | lemma forall_vector_3: "(\<forall>v::'a::zero^3. P v) \<longleftrightarrow> (\<forall>x y z. P(vector[x, y, z]))" | |
| 78480 | 731 | proof - | 
| 732 | have "P v" if "\<And>x y z. P (vector [x, y, z])" for v | |
| 733 | proof - | |
| 734 | have "vector [v$1, v$2, v$3] = v" | |
| 735 | by (smt (verit, best) exhaust_3 vec_eq_iff vector_3) | |
| 736 | then show ?thesis | |
| 737 | by (metis that) | |
| 738 | qed | |
| 739 | then show ?thesis by auto | |
| 740 | qed | |
| 69669 | 741 | |
| 70136 | 742 | subsection\<^marker>\<open>tag unimportant\<close> \<open>lambda skolemization on cartesian products\<close> | 
| 69669 | 743 | |
| 78480 | 744 | lemma lambda_skolem: "(\<forall>i. \<exists>x. P i x) \<longleftrightarrow> (\<exists>x::'a ^ 'n. \<forall>i. P i (x $ i))" | 
| 745 | by (metis vec_lambda_beta) | |
| 69669 | 746 | |
| 747 | ||
| 748 | text \<open>The same result in terms of square matrices.\<close> | |
| 749 | ||
| 750 | ||
| 751 | text \<open>Considering an n-element vector as an n-by-1 or 1-by-n matrix.\<close> | |
| 752 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 753 | definition "rowvector v = (\<chi> i j. (v$j))" | 
| 69669 | 754 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 755 | definition "columnvector v = (\<chi> i j. (v$i))" | 
| 69669 | 756 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 757 | lemma transpose_columnvector: "transpose(columnvector v) = rowvector v" | 
| 69669 | 758 | by (simp add: transpose_def rowvector_def columnvector_def vec_eq_iff) | 
| 759 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 760 | lemma transpose_rowvector: "transpose(rowvector v) = columnvector v" | 
| 69669 | 761 | by (simp add: transpose_def columnvector_def rowvector_def vec_eq_iff) | 
| 762 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 763 | lemma dot_rowvector_columnvector: "columnvector (A *v v) = A ** columnvector v" | 
| 69669 | 764 | by (vector columnvector_def matrix_matrix_mult_def matrix_vector_mult_def) | 
| 765 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 766 | lemma dot_matrix_product: | 
| 69669 | 767 | "(x::real^'n) \<bullet> y = (((rowvector x ::real^'n^1) ** (columnvector y :: real^1^'n))$1)$1" | 
| 768 | by (vector matrix_matrix_mult_def rowvector_def columnvector_def inner_vec_def) | |
| 769 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 770 | lemma dot_matrix_vector_mul: | 
| 69669 | 771 | fixes A B :: "real ^'n ^'n" and x y :: "real ^'n" | 
| 772 | shows "(A *v x) \<bullet> (B *v y) = | |
| 773 | (((rowvector x :: real^'n^1) ** ((transpose A ** B) ** (columnvector y :: real ^1^'n)))$1)$1" | |
| 78480 | 774 | by (metis dot_lmul_matrix dot_matrix_product dot_rowvector_columnvector matrix_mul_assoc vector_transpose_matrix) | 
| 69669 | 775 | |
| 776 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 777 | lemma dim_substandard_cart: "vec.dim {x::'a::field^'n. \<forall>i. i \<notin> d \<longrightarrow> x$i = 0} = card d"
 | 
| 69669 | 778 | (is "vec.dim ?A = _") | 
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 779 | proof (rule vec.dim_unique) | 
| 69669 | 780 | let ?B = "((\<lambda>x. axis x 1) ` d)" | 
| 781 | have subset_basis: "?B \<subseteq> cart_basis" | |
| 782 | by (auto simp: cart_basis_def) | |
| 783 | show "?B \<subseteq> ?A" | |
| 784 | by (auto simp: axis_def) | |
| 785 | show "vec.independent ((\<lambda>x. axis x 1) ` d)" | |
| 786 | using subset_basis | |
| 787 | by (rule vec.independent_mono[OF vec.independent_Basis]) | |
| 788 | have "x \<in> vec.span ?B" if "\<forall>i. i \<notin> d \<longrightarrow> x $ i = 0" for x::"'a^'n" | |
| 789 | proof - | |
| 790 | have "finite ?B" | |
| 791 | using subset_basis finite_cart_basis | |
| 792 | by (rule finite_subset) | |
| 793 | have "x = (\<Sum>i\<in>UNIV. x $ i *s axis i 1)" | |
| 794 | by (rule basis_expansion[symmetric]) | |
| 795 | also have "\<dots> = (\<Sum>i\<in>d. (x $ i) *s axis i 1)" | |
| 796 | by (rule sum.mono_neutral_cong_right) (auto simp: that) | |
| 797 | also have "\<dots> \<in> vec.span ?B" | |
| 798 | by (simp add: vec.span_sum vec.span_clauses) | |
| 799 | finally show "x \<in> vec.span ?B" . | |
| 800 | qed | |
| 801 | then show "?A \<subseteq> vec.span ?B" by auto | |
| 802 | qed (simp add: card_image inj_on_def axis_eq_axis) | |
| 803 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 804 | lemma affinity_inverses: | 
| 69669 | 805 | assumes m0: "m \<noteq> (0::'a::field)" | 
| 806 | shows "(\<lambda>x. m *s x + c) \<circ> (\<lambda>x. inverse(m) *s x + (-(inverse(m) *s c))) = id" | |
| 807 | "(\<lambda>x. inverse(m) *s x + (-(inverse(m) *s c))) \<circ> (\<lambda>x. m *s x + c) = id" | |
| 808 | using m0 | |
| 809 | by (auto simp add: fun_eq_iff vector_add_ldistrib diff_conv_add_uminus simp del: add_uminus_conv_diff) | |
| 810 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 811 | lemma vector_affinity_eq: | 
| 69669 | 812 | assumes m0: "(m::'a::field) \<noteq> 0" | 
| 813 | shows "m *s x + c = y \<longleftrightarrow> x = inverse m *s y + -(inverse m *s c)" | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 814 | proof | 
| 69669 | 815 | assume h: "m *s x + c = y" | 
| 816 | hence "m *s x = y - c" by (simp add: field_simps) | |
| 817 | hence "inverse m *s (m *s x) = inverse m *s (y - c)" by simp | |
| 818 | then show "x = inverse m *s y + - (inverse m *s c)" | |
| 78480 | 819 | by (simp add: m0 vec.scale_right_diff_distrib) | 
| 69669 | 820 | next | 
| 821 | assume h: "x = inverse m *s y + - (inverse m *s c)" | |
| 822 | show "m *s x + c = y" unfolding h | |
| 823 | using m0 by (simp add: vector_smult_assoc vector_ssub_ldistrib) | |
| 824 | qed | |
| 825 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 826 | lemma vector_eq_affinity: | 
| 69669 | 827 | "(m::'a::field) \<noteq> 0 ==> (y = m *s x + c \<longleftrightarrow> inverse(m) *s y + -(inverse(m) *s c) = x)" | 
| 78480 | 828 | by (metis vector_affinity_eq) | 
| 69669 | 829 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 830 | lemma vector_cart: | 
| 69669 | 831 | fixes f :: "real^'n \<Rightarrow> real" | 
| 832 | shows "(\<chi> i. f (axis i 1)) = (\<Sum>i\<in>Basis. f i *\<^sub>R i)" | |
| 833 | unfolding euclidean_eq_iff[where 'a="real^'n"] | |
| 834 | by simp (simp add: Basis_vec_def inner_axis) | |
| 835 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 836 | lemma const_vector_cart:"((\<chi> i. d)::real^'n) = (\<Sum>i\<in>Basis. d *\<^sub>R i)" | 
| 69669 | 837 | by (rule vector_cart) | 
| 838 | ||
| 70136 | 839 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Explicit formulas for low dimensions\<close> | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 840 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 841 | lemma  prod_neutral_const: "prod f {(1::nat)..1} = f 1"
 | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 842 | by simp | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 843 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 844 | lemma  prod_2: "prod f {(1::nat)..2} = f 1 * f 2"
 | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 845 | by (simp add: eval_nat_numeral atLeastAtMostSuc_conv mult.commute) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 846 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 847 | lemma  prod_3: "prod f {(1::nat)..3} = f 1 * f 2 * f 3"
 | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 848 | by (simp add: eval_nat_numeral atLeastAtMostSuc_conv mult.commute) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 849 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 850 | |
| 69683 | 851 | subsection \<open>Orthogonality of a matrix\<close> | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 852 | |
| 70136 | 853 | definition\<^marker>\<open>tag important\<close> "orthogonal_matrix (Q::'a::semiring_1^'n^'n) \<longleftrightarrow> | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 854 | transpose Q ** Q = mat 1 \<and> Q ** transpose Q = mat 1" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 855 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 856 | lemma orthogonal_matrix: "orthogonal_matrix (Q:: real ^'n^'n) \<longleftrightarrow> transpose Q ** Q = mat 1" | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 857 | by (metis matrix_left_right_inverse orthogonal_matrix_def) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 858 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 859 | lemma orthogonal_matrix_id: "orthogonal_matrix (mat 1 :: _^'n^'n)" | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 860 | by (simp add: orthogonal_matrix_def) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 861 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 862 | proposition orthogonal_matrix_mul: | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 863 | fixes A :: "real ^'n^'n" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 864 | assumes "orthogonal_matrix A" "orthogonal_matrix B" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 865 | shows "orthogonal_matrix(A ** B)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 866 | using assms | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 867 | by (simp add: orthogonal_matrix matrix_transpose_mul matrix_left_right_inverse matrix_mul_assoc) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 868 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 869 | proposition orthogonal_transformation_matrix: | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 870 | fixes f:: "real^'n \<Rightarrow> real^'n" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 871 | shows "orthogonal_transformation f \<longleftrightarrow> linear f \<and> orthogonal_matrix(matrix f)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 872 | (is "?lhs \<longleftrightarrow> ?rhs") | 
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 873 | proof - | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 874 | let ?mf = "matrix f" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 875 | let ?ot = "orthogonal_transformation f" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 876 | let ?U = "UNIV :: 'n set" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 877 | have fU: "finite ?U" by simp | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 878 | let ?m1 = "mat 1 :: real ^'n^'n" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 879 |   {
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 880 | assume ot: ?ot | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 881 | from ot have lf: "Vector_Spaces.linear (*s) (*s) f" and fd: "\<And>v w. f v \<bullet> f w = v \<bullet> w" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 882 | unfolding orthogonal_transformation_def orthogonal_matrix linear_def scalar_mult_eq_scaleR | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 883 | by blast+ | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 884 |     {
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 885 | fix i j | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 886 | let ?A = "transpose ?mf ** ?mf" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 887 | have th0: "\<And>b (x::'a::comm_ring_1). (if b then 1 else 0)*x = (if b then x else 0)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 888 | "\<And>b (x::'a::comm_ring_1). x*(if b then 1 else 0) = (if b then x else 0)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 889 | by simp_all | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 890 | from fd[of "axis i 1" "axis j 1", | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 891 | simplified matrix_works[OF lf, symmetric] dot_matrix_vector_mul] | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 892 | have "?A$i$j = ?m1 $ i $ j" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 893 | by (simp add: inner_vec_def matrix_matrix_mult_def columnvector_def rowvector_def | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 894 | th0 sum.delta[OF fU] mat_def axis_def) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 895 | } | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 896 | then have "orthogonal_matrix ?mf" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 897 | unfolding orthogonal_matrix | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 898 | by vector | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 899 | with lf have ?rhs | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 900 | unfolding linear_def scalar_mult_eq_scaleR | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 901 | by blast | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 902 | } | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 903 | moreover | 
| 78480 | 904 | have ?lhs if "Vector_Spaces.linear (*s) (*s) f" and "orthogonal_matrix ?mf" | 
| 905 | using that unfolding orthogonal_matrix_def norm_eq orthogonal_transformation | |
| 906 | by (metis dot_matrix_product dot_matrix_vector_mul linear_matrix_vector_mul_eq matrix_mul_lid matrix_vector_mul(2)) | |
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 907 | ultimately show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 908 | by (auto simp: linear_def scalar_mult_eq_scaleR) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 909 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 910 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 911 | |
| 71044 | 912 | subsection \<open>Finding an Orthogonal Matrix\<close> | 
| 913 | ||
| 914 | text \<open>We can find an orthogonal matrix taking any unit vector to any other.\<close> | |
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 915 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 916 | lemma orthogonal_matrix_transpose [simp]: | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 917 | "orthogonal_matrix(transpose A) \<longleftrightarrow> orthogonal_matrix A" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 918 | by (auto simp: orthogonal_matrix_def) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 919 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 920 | lemma orthogonal_matrix_orthonormal_columns: | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 921 | fixes A :: "real^'n^'n" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 922 | shows "orthogonal_matrix A \<longleftrightarrow> | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 923 | (\<forall>i. norm(column i A) = 1) \<and> | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 924 | (\<forall>i j. i \<noteq> j \<longrightarrow> orthogonal (column i A) (column j A))" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 925 | by (auto simp: orthogonal_matrix matrix_mult_transpose_dot_column vec_eq_iff mat_def norm_eq_1 orthogonal_def) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 926 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 927 | lemma orthogonal_matrix_orthonormal_rows: | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 928 | fixes A :: "real^'n^'n" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 929 | shows "orthogonal_matrix A \<longleftrightarrow> | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 930 | (\<forall>i. norm(row i A) = 1) \<and> | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 931 | (\<forall>i j. i \<noteq> j \<longrightarrow> orthogonal (row i A) (row j A))" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 932 | using orthogonal_matrix_orthonormal_columns [of "transpose A"] by simp | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 933 | |
| 78480 | 934 | proposition orthogonal_matrix_exists_basis: | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 935 | fixes a :: "real^'n" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 936 | assumes "norm a = 1" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 937 | obtains A where "orthogonal_matrix A" "A *v (axis k 1) = a" | 
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 938 | proof - | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 939 | obtain S where "a \<in> S" "pairwise orthogonal S" and noS: "\<And>x. x \<in> S \<Longrightarrow> norm x = 1" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 940 |    and "independent S" "card S = CARD('n)" "span S = UNIV"
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 941 | using vector_in_orthonormal_basis assms by force | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 942 | then obtain f0 where "bij_betw f0 (UNIV::'n set) S" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 943 | by (metis finite_class.finite_UNIV finite_same_card_bij finiteI_independent) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 944 | then obtain f where f: "bij_betw f (UNIV::'n set) S" and a: "a = f k" | 
| 73648 | 945 | using bij_swap_iff [of f0 k "inv f0 a"] | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 946 | by (metis UNIV_I \<open>a \<in> S\<close> bij_betw_inv_into_right bij_betw_swap_iff swap_apply(1)) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 947 | show thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 948 | proof | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 949 | have [simp]: "\<And>i. norm (f i) = 1" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 950 | using bij_betwE [OF \<open>bij_betw f UNIV S\<close>] by (blast intro: noS) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 951 | have [simp]: "\<And>i j. i \<noteq> j \<Longrightarrow> orthogonal (f i) (f j)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 952 | using \<open>pairwise orthogonal S\<close> \<open>bij_betw f UNIV S\<close> | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 953 | by (auto simp: pairwise_def bij_betw_def inj_on_def) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 954 | show "orthogonal_matrix (\<chi> i j. f j $ i)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 955 | by (simp add: orthogonal_matrix_orthonormal_columns column_def) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 956 | show "(\<chi> i j. f j $ i) *v axis k 1 = a" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 957 | by (simp add: matrix_vector_mult_def axis_def a if_distrib cong: if_cong) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 958 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 959 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 960 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 961 | lemma orthogonal_transformation_exists_1: | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 962 | fixes a b :: "real^'n" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 963 | assumes "norm a = 1" "norm b = 1" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 964 | obtains f where "orthogonal_transformation f" "f a = b" | 
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 965 | proof - | 
| 78480 | 966 | obtain k A B where AB: "orthogonal_matrix A" "orthogonal_matrix B" and eq: "A *v (axis k 1) = a" "B *v (axis k 1) = b" | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 967 | using orthogonal_matrix_exists_basis assms by metis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 968 | let ?f = "\<lambda>x. (B ** transpose A) *v x" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 969 | show thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 970 | proof | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 971 | show "orthogonal_transformation ?f" | 
| 78480 | 972 | by (simp add: AB orthogonal_matrix_mul orthogonal_transformation_matrix) | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 973 | next | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 974 | show "?f a = b" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 975 | using \<open>orthogonal_matrix A\<close> unfolding orthogonal_matrix_def | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 976 | by (metis eq matrix_mul_rid matrix_vector_mul_assoc) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 977 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 978 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 979 | |
| 78480 | 980 | proposition orthogonal_transformation_exists: | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 981 | fixes a b :: "real^'n" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 982 | assumes "norm a = norm b" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 983 | obtains f where "orthogonal_transformation f" "f a = b" | 
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 984 | proof (cases "a = 0 \<or> b = 0") | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 985 | case True | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 986 | with assms show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 987 | using that by force | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 988 | next | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 989 | case False | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 990 | then obtain f where f: "orthogonal_transformation f" and eq: "f (a /\<^sub>R norm a) = (b /\<^sub>R norm b)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 991 | by (auto intro: orthogonal_transformation_exists_1 [of "a /\<^sub>R norm a" "b /\<^sub>R norm b"]) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 992 | show ?thesis | 
| 78480 | 993 | using False assms eq f orthogonal_transformation_scaleR that by fastforce | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 994 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 995 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 996 | |
| 71044 | 997 | subsection \<open>Scaling and isometry\<close> | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 998 | |
| 71044 | 999 | proposition scaling_linear: | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1000 | fixes f :: "'a::real_inner \<Rightarrow> 'a::real_inner" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1001 | assumes f0: "f 0 = 0" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1002 | and fd: "\<forall>x y. dist (f x) (f y) = c * dist x y" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1003 | shows "linear f" | 
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 1004 | proof - | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1005 |   {
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1006 | fix v w | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1007 | have "norm (f x) = c * norm x" for x | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1008 | by (metis dist_0_norm f0 fd) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1009 | then have "f v \<bullet> f w = c\<^sup>2 * (v \<bullet> w)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1010 | unfolding dot_norm_neg dist_norm[symmetric] | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1011 | by (simp add: fd power2_eq_square field_simps) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1012 | } | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1013 | then show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1014 | unfolding linear_iff vector_eq[where 'a="'a"] scalar_mult_eq_scaleR | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1015 | by (simp add: inner_add field_simps) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1016 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1017 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 1018 | lemma isometry_linear: | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1019 | "f (0::'a::real_inner) = (0::'a) \<Longrightarrow> \<forall>x y. dist(f x) (f y) = dist x y \<Longrightarrow> linear f" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1020 | by (rule scaling_linear[where c=1]) simp_all | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1021 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1022 | text \<open>Hence another formulation of orthogonal transformation\<close> | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1023 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 1024 | proposition orthogonal_transformation_isometry: | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1025 | "orthogonal_transformation f \<longleftrightarrow> f(0::'a::real_inner) = (0::'a) \<and> (\<forall>x y. dist(f x) (f y) = dist x y)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1026 | unfolding orthogonal_transformation | 
| 78480 | 1027 | by (metis dist_0_norm dist_norm isometry_linear linear_0 linear_diff) | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1028 | |
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1029 | |
| 71044 | 1030 | text \<open>Can extend an isometry from unit sphere:\<close> | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1031 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 1032 | lemma isometry_sphere_extend: | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1033 | fixes f:: "'a::real_inner \<Rightarrow> 'a" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1034 | assumes f1: "\<And>x. norm x = 1 \<Longrightarrow> norm (f x) = 1" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1035 | and fd1: "\<And>x y. \<lbrakk>norm x = 1; norm y = 1\<rbrakk> \<Longrightarrow> dist (f x) (f y) = dist x y" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1036 | shows "\<exists>g. orthogonal_transformation g \<and> (\<forall>x. norm x = 1 \<longrightarrow> g x = f x)" | 
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 1037 | proof - | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1038 |   {
 | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1039 | fix x y x' y' u v u' v' :: "'a" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1040 | assume H: "x = norm x *\<^sub>R u" "y = norm y *\<^sub>R v" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1041 | "x' = norm x *\<^sub>R u'" "y' = norm y *\<^sub>R v'" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1042 | and J: "norm u = 1" "norm u' = 1" "norm v = 1" "norm v' = 1" "norm(u' - v') = norm(u - v)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1043 | then have *: "u \<bullet> v = u' \<bullet> v' + v' \<bullet> u' - v \<bullet> u " | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1044 | by (simp add: norm_eq norm_eq_1 inner_add inner_diff) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1045 | have "norm (norm x *\<^sub>R u' - norm y *\<^sub>R v') = norm (norm x *\<^sub>R u - norm y *\<^sub>R v)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1046 | using J by (simp add: norm_eq norm_eq_1 inner_diff * field_simps) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1047 | then have "norm(x' - y') = norm(x - y)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1048 | using H by metis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1049 | } | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1050 | note norm_eq = this | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1051 | let ?g = "\<lambda>x. if x = 0 then 0 else norm x *\<^sub>R f (x /\<^sub>R norm x)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1052 | have thfg: "?g x = f x" if "norm x = 1" for x | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1053 | using that by auto | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1054 | have thd: "dist (?g x) (?g y) = dist x y" for x y | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1055 | proof (cases "x=0 \<or> y=0") | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1056 | case False | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1057 | show "dist (?g x) (?g y) = dist x y" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1058 | unfolding dist_norm | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1059 | proof (rule norm_eq) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1060 | show "x = norm x *\<^sub>R (x /\<^sub>R norm x)" "y = norm y *\<^sub>R (y /\<^sub>R norm y)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1061 | "norm (f (x /\<^sub>R norm x)) = 1" "norm (f (y /\<^sub>R norm y)) = 1" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1062 | using False f1 by auto | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1063 | qed (use False in \<open>auto simp: field_simps intro: f1 fd1[unfolded dist_norm]\<close>) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1064 | qed (auto simp: f1) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1065 | show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1066 | unfolding orthogonal_transformation_isometry | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1067 | by (rule exI[where x= ?g]) (metis thfg thd) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1068 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1069 | |
| 69683 | 1070 | subsection\<open>Induction on matrix row operations\<close> | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1071 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 1072 | lemma induct_matrix_row_operations: | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1073 | fixes P :: "real^'n^'n \<Rightarrow> bool" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1074 | assumes zero_row: "\<And>A i. row i A = 0 \<Longrightarrow> P A" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1075 | and diagonal: "\<And>A. (\<And>i j. i \<noteq> j \<Longrightarrow> A$i$j = 0) \<Longrightarrow> P A" | 
| 73648 | 1076 | and swap_cols: "\<And>A m n. \<lbrakk>P A; m \<noteq> n\<rbrakk> \<Longrightarrow> P(\<chi> i j. A $ i $ Transposition.transpose m n j)" | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1077 | and row_op: "\<And>A m n c. \<lbrakk>P A; m \<noteq> n\<rbrakk> | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1078 | \<Longrightarrow> P(\<chi> i. if i = m then row m A + c *\<^sub>R row n A else row i A)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1079 | shows "P A" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1080 | proof - | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1081 | have "P A" if "(\<And>i j. \<lbrakk>j \<in> -K; i \<noteq> j\<rbrakk> \<Longrightarrow> A$i$j = 0)" for A K | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1082 | proof - | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1083 | have "finite K" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1084 | by simp | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1085 | then show ?thesis using that | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1086 | proof (induction arbitrary: A rule: finite_induct) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1087 | case empty | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1088 | with diagonal show ?case | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1089 | by simp | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1090 | next | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1091 | case (insert k K) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1092 | note insertK = insert | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1093 | have "P A" if kk: "A$k$k \<noteq> 0" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1094 | and 0: "\<And>i j. \<lbrakk>j \<in> - insert k K; i \<noteq> j\<rbrakk> \<Longrightarrow> A$i$j = 0" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1095 | "\<And>i. \<lbrakk>i \<in> -L; i \<noteq> k\<rbrakk> \<Longrightarrow> A$i$k = 0" for A L | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1096 | proof - | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1097 | have "finite L" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1098 | by simp | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1099 | then show ?thesis using 0 kk | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1100 | proof (induction arbitrary: A rule: finite_induct) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1101 | case (empty B) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1102 | show ?case | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1103 | proof (rule insertK) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1104 | fix i j | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1105 | assume "i \<in> - K" "j \<noteq> i" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1106 | show "B $ j $ i = 0" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1107 | using \<open>j \<noteq> i\<close> \<open>i \<in> - K\<close> empty | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1108 | by (metis ComplD ComplI Compl_eq_Diff_UNIV Diff_empty UNIV_I insert_iff) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1109 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1110 | next | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1111 | case (insert l L B) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1112 | show ?case | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1113 | proof (cases "k = l") | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1114 | case True | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1115 | with insert show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1116 | by auto | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1117 | next | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1118 | case False | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1119 | let ?C = "\<chi> i. if i = l then row l B - (B $ l $ k / B $ k $ k) *\<^sub>R row k B else row i B" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1120 | have 1: "\<lbrakk>j \<in> - insert k K; i \<noteq> j\<rbrakk> \<Longrightarrow> ?C $ i $ j = 0" for j i | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1121 | by (auto simp: insert.prems(1) row_def) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1122 | have 2: "?C $ i $ k = 0" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1123 | if "i \<in> - L" "i \<noteq> k" for i | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1124 | proof (cases "i=l") | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1125 | case True | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1126 | with that insert.prems show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1127 | by (simp add: row_def) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1128 | next | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1129 | case False | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1130 | with that show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1131 | by (simp add: insert.prems(2) row_def) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1132 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1133 | have 3: "?C $ k $ k \<noteq> 0" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1134 | by (auto simp: insert.prems row_def \<open>k \<noteq> l\<close>) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1135 | have PC: "P ?C" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1136 | using insert.IH [OF 1 2 3] by auto | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1137 | have eqB: "(\<chi> i. if i = l then row l ?C + (B $ l $ k / B $ k $ k) *\<^sub>R row k ?C else row i ?C) = B" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1138 | using \<open>k \<noteq> l\<close> by (simp add: vec_eq_iff row_def) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1139 | show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1140 | using row_op [OF PC, of l k, where c = "B$l$k / B$k$k"] eqB \<open>k \<noteq> l\<close> | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1141 | by (simp add: cong: if_cong) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1142 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1143 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1144 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1145 | then have nonzero_hyp: "P A" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1146 | if kk: "A$k$k \<noteq> 0" and zeroes: "\<And>i j. j \<in> - insert k K \<and> i\<noteq>j \<Longrightarrow> A$i$j = 0" for A | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1147 | by (auto simp: intro!: kk zeroes) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1148 | show ?case | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1149 | proof (cases "row k A = 0") | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1150 | case True | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1151 | with zero_row show ?thesis by auto | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1152 | next | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1153 | case False | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1154 | then obtain l where l: "A$k$l \<noteq> 0" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1155 | by (auto simp: row_def zero_vec_def vec_eq_iff) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1156 | show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1157 | proof (cases "k = l") | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1158 | case True | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1159 | with l nonzero_hyp insert.prems show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1160 | by blast | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1161 | next | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1162 | case False | 
| 73648 | 1163 | have *: "A $ i $ Transposition.transpose k l j = 0" if "j \<noteq> k" "j \<notin> K" "i \<noteq> j" for i j | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1164 | using False l insert.prems that | 
| 73648 | 1165 | by (auto simp add: Transposition.transpose_def) | 
| 1166 | have "P (\<chi> i j. (\<chi> i j. A $ i $ Transposition.transpose k l j) $ i $ Transposition.transpose k l j)" | |
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1167 | by (rule swap_cols [OF nonzero_hyp False]) (auto simp: l *) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1168 | moreover | 
| 73648 | 1169 | have "(\<chi> i j. (\<chi> i j. A $ i $ Transposition.transpose k l j) $ i $ Transposition.transpose k l j) = A" | 
| 1170 | by simp | |
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1171 | ultimately show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1172 | by simp | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1173 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1174 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1175 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1176 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1177 | then show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1178 | by blast | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1179 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1180 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 1181 | lemma induct_matrix_elementary: | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1182 | fixes P :: "real^'n^'n \<Rightarrow> bool" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1183 | assumes mult: "\<And>A B. \<lbrakk>P A; P B\<rbrakk> \<Longrightarrow> P(A ** B)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1184 | and zero_row: "\<And>A i. row i A = 0 \<Longrightarrow> P A" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1185 | and diagonal: "\<And>A. (\<And>i j. i \<noteq> j \<Longrightarrow> A$i$j = 0) \<Longrightarrow> P A" | 
| 73648 | 1186 | and swap1: "\<And>m n. m \<noteq> n \<Longrightarrow> P(\<chi> i j. mat 1 $ i $ Transposition.transpose m n j)" | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1187 | and idplus: "\<And>m n c. m \<noteq> n \<Longrightarrow> P(\<chi> i j. if i = m \<and> j = n then c else of_bool (i = j))" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1188 | shows "P A" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1189 | proof - | 
| 73648 | 1190 | have swap: "P (\<chi> i j. A $ i $ Transposition.transpose m n j)" (is "P ?C") | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1191 | if "P A" "m \<noteq> n" for A m n | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1192 | proof - | 
| 73648 | 1193 | have "A ** (\<chi> i j. mat 1 $ i $ Transposition.transpose m n j) = ?C" | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1194 | by (simp add: matrix_matrix_mult_def mat_def vec_eq_iff if_distrib sum.delta_remove) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1195 | then show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1196 | using mult swap1 that by metis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1197 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1198 | have row: "P (\<chi> i. if i = m then row m A + c *\<^sub>R row n A else row i A)" (is "P ?C") | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1199 | if "P A" "m \<noteq> n" for A m n c | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1200 | proof - | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1201 | let ?B = "\<chi> i j. if i = m \<and> j = n then c else of_bool (i = j)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1202 | have "?B ** A = ?C" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1203 | using \<open>m \<noteq> n\<close> unfolding matrix_matrix_mult_def row_def of_bool_def | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1204 | by (auto simp: vec_eq_iff if_distrib [of "\<lambda>x. x * y" for y] sum.remove cong: if_cong) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1205 | then show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1206 | by (rule subst) (auto simp: that mult idplus) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1207 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1208 | show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1209 | by (rule induct_matrix_row_operations [OF zero_row diagonal swap row]) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1210 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1211 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 1212 | lemma induct_matrix_elementary_alt: | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1213 | fixes P :: "real^'n^'n \<Rightarrow> bool" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1214 | assumes mult: "\<And>A B. \<lbrakk>P A; P B\<rbrakk> \<Longrightarrow> P(A ** B)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1215 | and zero_row: "\<And>A i. row i A = 0 \<Longrightarrow> P A" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1216 | and diagonal: "\<And>A. (\<And>i j. i \<noteq> j \<Longrightarrow> A$i$j = 0) \<Longrightarrow> P A" | 
| 73648 | 1217 | and swap1: "\<And>m n. m \<noteq> n \<Longrightarrow> P(\<chi> i j. mat 1 $ i $ Transposition.transpose m n j)" | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1218 | and idplus: "\<And>m n. m \<noteq> n \<Longrightarrow> P(\<chi> i j. of_bool (i = m \<and> j = n \<or> i = j))" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1219 | shows "P A" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1220 | proof - | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1221 | have *: "P (\<chi> i j. if i = m \<and> j = n then c else of_bool (i = j))" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1222 | if "m \<noteq> n" for m n c | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1223 | proof (cases "c = 0") | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1224 | case True | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1225 | with diagonal show ?thesis by auto | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1226 | next | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1227 | case False | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1228 | then have eq: "(\<chi> i j. if i = m \<and> j = n then c else of_bool (i = j)) = | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1229 | (\<chi> i j. if i = j then (if j = n then inverse c else 1) else 0) ** | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1230 | (\<chi> i j. of_bool (i = m \<and> j = n \<or> i = j)) ** | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1231 | (\<chi> i j. if i = j then if j = n then c else 1 else 0)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1232 | using \<open>m \<noteq> n\<close> | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1233 | apply (simp add: matrix_matrix_mult_def vec_eq_iff of_bool_def if_distrib [of "\<lambda>x. y * x" for y] cong: if_cong) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1234 | apply (simp add: if_if_eq_conj sum.neutral conj_commute cong: conj_cong) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1235 | done | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1236 | show ?thesis | 
| 78480 | 1237 | unfolding eq by (intro mult idplus that) (auto intro: diagonal) | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1238 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1239 | show ?thesis | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1240 | by (rule induct_matrix_elementary) (auto intro: assms *) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1241 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1242 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 1243 | lemma matrix_vector_mult_matrix_matrix_mult_compose: | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1244 | "(*v) (A ** B) = (*v) A \<circ> (*v) B" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1245 | by (auto simp: matrix_vector_mul_assoc) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1246 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 1247 | lemma induct_linear_elementary: | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1248 | fixes f :: "real^'n \<Rightarrow> real^'n" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1249 | assumes "linear f" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1250 | and comp: "\<And>f g. \<lbrakk>linear f; linear g; P f; P g\<rbrakk> \<Longrightarrow> P(f \<circ> g)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1251 | and zeroes: "\<And>f i. \<lbrakk>linear f; \<And>x. (f x) $ i = 0\<rbrakk> \<Longrightarrow> P f" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1252 | and const: "\<And>c. P(\<lambda>x. \<chi> i. c i * x$i)" | 
| 73648 | 1253 | and swap: "\<And>m n::'n. m \<noteq> n \<Longrightarrow> P(\<lambda>x. \<chi> i. x $ Transposition.transpose m n i)" | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1254 | and idplus: "\<And>m n::'n. m \<noteq> n \<Longrightarrow> P(\<lambda>x. \<chi> i. if i = m then x$m + x$n else x$i)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1255 | shows "P f" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1256 | proof - | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1257 | have "P ((*v) A)" for A | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1258 | proof (rule induct_matrix_elementary_alt) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1259 | fix A B | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1260 | assume "P ((*v) A)" and "P ((*v) B)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1261 | then show "P ((*v) (A ** B))" | 
| 71044 | 1262 | by (auto simp add: matrix_vector_mult_matrix_matrix_mult_compose intro!: comp) | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1263 | next | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1264 | fix A :: "real^'n^'n" and i | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1265 | assume "row i A = 0" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1266 | show "P ((*v) A)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1267 | using matrix_vector_mul_linear | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1268 | by (rule zeroes[where i=i]) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1269 | (metis \<open>row i A = 0\<close> inner_zero_left matrix_vector_mul_component row_def vec_lambda_eta) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1270 | next | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1271 | fix A :: "real^'n^'n" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1272 | assume 0: "\<And>i j. i \<noteq> j \<Longrightarrow> A $ i $ j = 0" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1273 | have "A $ i $ i * x $ i = (\<Sum>j\<in>UNIV. A $ i $ j * x $ j)" for x and i :: "'n" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1274 | by (simp add: 0 comm_monoid_add_class.sum.remove [where x=i]) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1275 | then have "(\<lambda>x. \<chi> i. A $ i $ i * x $ i) = ((*v) A)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1276 | by (auto simp: 0 matrix_vector_mult_def) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1277 | then show "P ((*v) A)" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1278 | using const [of "\<lambda>i. A $ i $ i"] by simp | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1279 | next | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1280 | fix m n :: "'n" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1281 | assume "m \<noteq> n" | 
| 73648 | 1282 | have eq: "(\<Sum>j\<in>UNIV. if i = Transposition.transpose m n j then x $ j else 0) = | 
| 1283 | (\<Sum>j\<in>UNIV. if j = Transposition.transpose m n i then x $ j else 0)" | |
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1284 | for i and x :: "real^'n" | 
| 73466 | 1285 | by (rule sum.cong) (auto simp add: swap_id_eq) | 
| 73648 | 1286 | have "(\<lambda>x::real^'n. \<chi> i. x $ Transposition.transpose m n i) = ((*v) (\<chi> i j. if i = Transposition.transpose m n j then 1 else 0))" | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1287 | by (auto simp: mat_def matrix_vector_mult_def eq if_distrib [of "\<lambda>x. x * y" for y] cong: if_cong) | 
| 73648 | 1288 | with swap [OF \<open>m \<noteq> n\<close>] show "P ((*v) (\<chi> i j. mat 1 $ i $ Transposition.transpose m n j))" | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1289 | by (simp add: mat_def matrix_vector_mult_def) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1290 | next | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1291 | fix m n :: "'n" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1292 | assume "m \<noteq> n" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1293 | then have "x $ m + x $ n = (\<Sum>j\<in>UNIV. of_bool (j = n \<or> m = j) * x $ j)" for x :: "real^'n" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1294 | by (auto simp: of_bool_def if_distrib [of "\<lambda>x. x * y" for y] sum.remove cong: if_cong) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1295 | then have "(\<lambda>x::real^'n. \<chi> i. if i = m then x $ m + x $ n else x $ i) = | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1296 | ((*v) (\<chi> i j. of_bool (i = m \<and> j = n \<or> i = j)))" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1297 | unfolding matrix_vector_mult_def of_bool_def | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1298 | by (auto simp: vec_eq_iff if_distrib [of "\<lambda>x. x * y" for y] cong: if_cong) | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1299 | then show "P ((*v) (\<chi> i j. of_bool (i = m \<and> j = n \<or> i = j)))" | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1300 | using idplus [OF \<open>m \<noteq> n\<close>] by simp | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1301 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1302 | then show ?thesis | 
| 71633 | 1303 | by (metis \<open>linear f\<close> matrix_vector_mul(2)) | 
| 69675 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1304 | qed | 
| 
880ab0f27ddf
Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
 immler parents: 
69669diff
changeset | 1305 | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: diff
changeset | 1306 | end |