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