| author | chaieb | 
| Thu, 09 Jul 2009 08:55:42 +0200 | |
| changeset 31967 | 81dbc693143b | 
| parent 29700 | 22faf21db3df | 
| child 32440 | 153965be0f4b | 
| permissions | -rw-r--r-- | 
| 14593 | 1  | 
(* Title: HOL/Matrix/Matrix.thy  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Steven Obua  | 
|
4  | 
*)  | 
|
5  | 
||
| 17915 | 6  | 
theory Matrix  | 
| 27484 | 7  | 
imports Main  | 
| 17915 | 8  | 
begin  | 
| 14940 | 9  | 
|
| 27484 | 10  | 
types 'a infmatrix = "nat \<Rightarrow> nat \<Rightarrow> 'a"  | 
11  | 
||
12  | 
definition nonzero_positions :: "(nat \<Rightarrow> nat \<Rightarrow> 'a::zero) \<Rightarrow> nat \<times> nat \<Rightarrow> bool" where  | 
|
13  | 
  "nonzero_positions A = {pos. A (fst pos) (snd pos) ~= 0}"
 | 
|
14  | 
||
15  | 
typedef 'a matrix = "{(f::(nat \<Rightarrow> nat \<Rightarrow> 'a::zero)). finite (nonzero_positions f)}"
 | 
|
16  | 
proof -  | 
|
17  | 
  have "(\<lambda>j i. 0) \<in> {(f::(nat \<Rightarrow> nat \<Rightarrow> 'a::zero)). finite (nonzero_positions f)}"
 | 
|
18  | 
by (simp add: nonzero_positions_def)  | 
|
19  | 
then show ?thesis by auto  | 
|
20  | 
qed  | 
|
21  | 
||
22  | 
declare Rep_matrix_inverse[simp]  | 
|
23  | 
||
24  | 
lemma finite_nonzero_positions : "finite (nonzero_positions (Rep_matrix A))"  | 
|
25  | 
apply (rule Abs_matrix_induct)  | 
|
26  | 
by (simp add: Abs_matrix_inverse matrix_def)  | 
|
27  | 
||
28  | 
constdefs  | 
|
29  | 
  nrows :: "('a::zero) matrix \<Rightarrow> nat"
 | 
|
30  | 
  "nrows A == if nonzero_positions(Rep_matrix A) = {} then 0 else Suc(Max ((image fst) (nonzero_positions (Rep_matrix A))))"
 | 
|
31  | 
  ncols :: "('a::zero) matrix \<Rightarrow> nat"
 | 
|
32  | 
  "ncols A == if nonzero_positions(Rep_matrix A) = {} then 0 else Suc(Max ((image snd) (nonzero_positions (Rep_matrix A))))"
 | 
|
33  | 
||
34  | 
lemma nrows:  | 
|
35  | 
assumes hyp: "nrows A \<le> m"  | 
|
36  | 
shows "(Rep_matrix A m n) = 0" (is ?concl)  | 
|
37  | 
proof cases  | 
|
38  | 
  assume "nonzero_positions(Rep_matrix A) = {}"
 | 
|
39  | 
then show "(Rep_matrix A m n) = 0" by (simp add: nonzero_positions_def)  | 
|
40  | 
next  | 
|
41  | 
  assume a: "nonzero_positions(Rep_matrix A) \<noteq> {}"
 | 
|
42  | 
let ?S = "fst`(nonzero_positions(Rep_matrix A))"  | 
|
43  | 
have c: "finite (?S)" by (simp add: finite_nonzero_positions)  | 
|
44  | 
from hyp have d: "Max (?S) < m" by (simp add: a nrows_def)  | 
|
45  | 
have "m \<notin> ?S"  | 
|
46  | 
proof -  | 
|
47  | 
have "m \<in> ?S \<Longrightarrow> m <= Max(?S)" by (simp add: Max_ge [OF c])  | 
|
48  | 
moreover from d have "~(m <= Max ?S)" by (simp)  | 
|
49  | 
ultimately show "m \<notin> ?S" by (auto)  | 
|
50  | 
qed  | 
|
51  | 
thus "Rep_matrix A m n = 0" by (simp add: nonzero_positions_def image_Collect)  | 
|
52  | 
qed  | 
|
53  | 
||
54  | 
constdefs  | 
|
55  | 
transpose_infmatrix :: "'a infmatrix \<Rightarrow> 'a infmatrix"  | 
|
56  | 
"transpose_infmatrix A j i == A i j"  | 
|
57  | 
  transpose_matrix :: "('a::zero) matrix \<Rightarrow> 'a matrix"
 | 
|
58  | 
"transpose_matrix == Abs_matrix o transpose_infmatrix o Rep_matrix"  | 
|
59  | 
||
60  | 
declare transpose_infmatrix_def[simp]  | 
|
61  | 
||
62  | 
lemma transpose_infmatrix_twice[simp]: "transpose_infmatrix (transpose_infmatrix A) = A"  | 
|
63  | 
by ((rule ext)+, simp)  | 
|
64  | 
||
65  | 
lemma transpose_infmatrix: "transpose_infmatrix (% j i. P j i) = (% j i. P i j)"  | 
|
66  | 
apply (rule ext)+  | 
|
67  | 
by (simp add: transpose_infmatrix_def)  | 
|
68  | 
||
69  | 
lemma transpose_infmatrix_closed[simp]: "Rep_matrix (Abs_matrix (transpose_infmatrix (Rep_matrix x))) = transpose_infmatrix (Rep_matrix x)"  | 
|
70  | 
apply (rule Abs_matrix_inverse)  | 
|
71  | 
apply (simp add: matrix_def nonzero_positions_def image_def)  | 
|
72  | 
proof -  | 
|
73  | 
  let ?A = "{pos. Rep_matrix x (snd pos) (fst pos) \<noteq> 0}"
 | 
|
74  | 
let ?swap = "% pos. (snd pos, fst pos)"  | 
|
75  | 
  let ?B = "{pos. Rep_matrix x (fst pos) (snd pos) \<noteq> 0}"
 | 
|
76  | 
have swap_image: "?swap`?A = ?B"  | 
|
77  | 
apply (simp add: image_def)  | 
|
78  | 
apply (rule set_ext)  | 
|
79  | 
apply (simp)  | 
|
80  | 
proof  | 
|
81  | 
fix y  | 
|
82  | 
assume hyp: "\<exists>a b. Rep_matrix x b a \<noteq> 0 \<and> y = (b, a)"  | 
|
83  | 
thus "Rep_matrix x (fst y) (snd y) \<noteq> 0"  | 
|
84  | 
proof -  | 
|
85  | 
from hyp obtain a b where "(Rep_matrix x b a \<noteq> 0 & y = (b,a))" by blast  | 
|
86  | 
then show "Rep_matrix x (fst y) (snd y) \<noteq> 0" by (simp)  | 
|
87  | 
qed  | 
|
88  | 
next  | 
|
89  | 
fix y  | 
|
90  | 
assume hyp: "Rep_matrix x (fst y) (snd y) \<noteq> 0"  | 
|
91  | 
show "\<exists> a b. (Rep_matrix x b a \<noteq> 0 & y = (b,a))"  | 
|
92  | 
by (rule exI[of _ "snd y"], rule exI[of _ "fst y"]) (simp add: hyp)  | 
|
93  | 
qed  | 
|
94  | 
then have "finite (?swap`?A)"  | 
|
95  | 
proof -  | 
|
96  | 
have "finite (nonzero_positions (Rep_matrix x))" by (simp add: finite_nonzero_positions)  | 
|
97  | 
then have "finite ?B" by (simp add: nonzero_positions_def)  | 
|
98  | 
with swap_image show "finite (?swap`?A)" by (simp)  | 
|
99  | 
qed  | 
|
100  | 
moreover  | 
|
101  | 
have "inj_on ?swap ?A" by (simp add: inj_on_def)  | 
|
102  | 
ultimately show "finite ?A"by (rule finite_imageD[of ?swap ?A])  | 
|
103  | 
qed  | 
|
104  | 
||
105  | 
lemma infmatrixforward: "(x::'a infmatrix) = y \<Longrightarrow> \<forall> a b. x a b = y a b" by auto  | 
|
106  | 
||
107  | 
lemma transpose_infmatrix_inject: "(transpose_infmatrix A = transpose_infmatrix B) = (A = B)"  | 
|
108  | 
apply (auto)  | 
|
109  | 
apply (rule ext)+  | 
|
110  | 
apply (simp add: transpose_infmatrix)  | 
|
111  | 
apply (drule infmatrixforward)  | 
|
112  | 
apply (simp)  | 
|
113  | 
done  | 
|
114  | 
||
115  | 
lemma transpose_matrix_inject: "(transpose_matrix A = transpose_matrix B) = (A = B)"  | 
|
116  | 
apply (simp add: transpose_matrix_def)  | 
|
117  | 
apply (subst Rep_matrix_inject[THEN sym])+  | 
|
118  | 
apply (simp only: transpose_infmatrix_closed transpose_infmatrix_inject)  | 
|
119  | 
done  | 
|
120  | 
||
121  | 
lemma transpose_matrix[simp]: "Rep_matrix(transpose_matrix A) j i = Rep_matrix A i j"  | 
|
122  | 
by (simp add: transpose_matrix_def)  | 
|
123  | 
||
124  | 
lemma transpose_transpose_id[simp]: "transpose_matrix (transpose_matrix A) = A"  | 
|
125  | 
by (simp add: transpose_matrix_def)  | 
|
126  | 
||
127  | 
lemma nrows_transpose[simp]: "nrows (transpose_matrix A) = ncols A"  | 
|
128  | 
by (simp add: nrows_def ncols_def nonzero_positions_def transpose_matrix_def image_def)  | 
|
129  | 
||
130  | 
lemma ncols_transpose[simp]: "ncols (transpose_matrix A) = nrows A"  | 
|
131  | 
by (simp add: nrows_def ncols_def nonzero_positions_def transpose_matrix_def image_def)  | 
|
132  | 
||
133  | 
lemma ncols: "ncols A <= n \<Longrightarrow> Rep_matrix A m n = 0"  | 
|
134  | 
proof -  | 
|
135  | 
assume "ncols A <= n"  | 
|
136  | 
then have "nrows (transpose_matrix A) <= n" by (simp)  | 
|
137  | 
then have "Rep_matrix (transpose_matrix A) n m = 0" by (rule nrows)  | 
|
138  | 
thus "Rep_matrix A m n = 0" by (simp add: transpose_matrix_def)  | 
|
139  | 
qed  | 
|
140  | 
||
141  | 
lemma ncols_le: "(ncols A <= n) = (! j i. n <= i \<longrightarrow> (Rep_matrix A j i) = 0)" (is "_ = ?st")  | 
|
142  | 
apply (auto)  | 
|
143  | 
apply (simp add: ncols)  | 
|
144  | 
proof (simp add: ncols_def, auto)  | 
|
145  | 
let ?P = "nonzero_positions (Rep_matrix A)"  | 
|
146  | 
let ?p = "snd`?P"  | 
|
147  | 
have a:"finite ?p" by (simp add: finite_nonzero_positions)  | 
|
148  | 
let ?m = "Max ?p"  | 
|
149  | 
assume "~(Suc (?m) <= n)"  | 
|
150  | 
then have b:"n <= ?m" by (simp)  | 
|
151  | 
fix a b  | 
|
152  | 
assume "(a,b) \<in> ?P"  | 
|
153  | 
  then have "?p \<noteq> {}" by (auto)
 | 
|
154  | 
with a have "?m \<in> ?p" by (simp)  | 
|
155  | 
moreover have "!x. (x \<in> ?p \<longrightarrow> (? y. (Rep_matrix A y x) \<noteq> 0))" by (simp add: nonzero_positions_def image_def)  | 
|
156  | 
ultimately have "? y. (Rep_matrix A y ?m) \<noteq> 0" by (simp)  | 
|
157  | 
moreover assume ?st  | 
|
158  | 
ultimately show "False" using b by (simp)  | 
|
159  | 
qed  | 
|
160  | 
||
161  | 
lemma less_ncols: "(n < ncols A) = (? j i. n <= i & (Rep_matrix A j i) \<noteq> 0)" (is ?concl)  | 
|
162  | 
proof -  | 
|
163  | 
have a: "!! (a::nat) b. (a < b) = (~(b <= a))" by arith  | 
|
164  | 
show ?concl by (simp add: a ncols_le)  | 
|
165  | 
qed  | 
|
166  | 
||
167  | 
lemma le_ncols: "(n <= ncols A) = (\<forall> m. (\<forall> j i. m <= i \<longrightarrow> (Rep_matrix A j i) = 0) \<longrightarrow> n <= m)" (is ?concl)  | 
|
168  | 
apply (auto)  | 
|
169  | 
apply (subgoal_tac "ncols A <= m")  | 
|
170  | 
apply (simp)  | 
|
171  | 
apply (simp add: ncols_le)  | 
|
172  | 
apply (drule_tac x="ncols A" in spec)  | 
|
173  | 
by (simp add: ncols)  | 
|
174  | 
||
175  | 
lemma nrows_le: "(nrows A <= n) = (! j i. n <= j \<longrightarrow> (Rep_matrix A j i) = 0)" (is ?s)  | 
|
176  | 
proof -  | 
|
177  | 
have "(nrows A <= n) = (ncols (transpose_matrix A) <= n)" by (simp)  | 
|
178  | 
also have "\<dots> = (! j i. n <= i \<longrightarrow> (Rep_matrix (transpose_matrix A) j i = 0))" by (rule ncols_le)  | 
|
179  | 
also have "\<dots> = (! j i. n <= i \<longrightarrow> (Rep_matrix A i j) = 0)" by (simp)  | 
|
180  | 
finally show "(nrows A <= n) = (! j i. n <= j \<longrightarrow> (Rep_matrix A j i) = 0)" by (auto)  | 
|
181  | 
qed  | 
|
182  | 
||
183  | 
lemma less_nrows: "(m < nrows A) = (? j i. m <= j & (Rep_matrix A j i) \<noteq> 0)" (is ?concl)  | 
|
184  | 
proof -  | 
|
185  | 
have a: "!! (a::nat) b. (a < b) = (~(b <= a))" by arith  | 
|
186  | 
show ?concl by (simp add: a nrows_le)  | 
|
187  | 
qed  | 
|
188  | 
||
189  | 
lemma le_nrows: "(n <= nrows A) = (\<forall> m. (\<forall> j i. m <= j \<longrightarrow> (Rep_matrix A j i) = 0) \<longrightarrow> n <= m)" (is ?concl)  | 
|
190  | 
apply (auto)  | 
|
191  | 
apply (subgoal_tac "nrows A <= m")  | 
|
192  | 
apply (simp)  | 
|
193  | 
apply (simp add: nrows_le)  | 
|
194  | 
apply (drule_tac x="nrows A" in spec)  | 
|
195  | 
by (simp add: nrows)  | 
|
196  | 
||
197  | 
lemma nrows_notzero: "Rep_matrix A m n \<noteq> 0 \<Longrightarrow> m < nrows A"  | 
|
198  | 
apply (case_tac "nrows A <= m")  | 
|
199  | 
apply (simp_all add: nrows)  | 
|
200  | 
done  | 
|
201  | 
||
202  | 
lemma ncols_notzero: "Rep_matrix A m n \<noteq> 0 \<Longrightarrow> n < ncols A"  | 
|
203  | 
apply (case_tac "ncols A <= n")  | 
|
204  | 
apply (simp_all add: ncols)  | 
|
205  | 
done  | 
|
206  | 
||
207  | 
lemma finite_natarray1: "finite {x. x < (n::nat)}"
 | 
|
208  | 
apply (induct n)  | 
|
209  | 
apply (simp)  | 
|
210  | 
proof -  | 
|
211  | 
fix n  | 
|
212  | 
  have "{x. x < Suc n} = insert n {x. x < n}"  by (rule set_ext, simp, arith)
 | 
|
213  | 
  moreover assume "finite {x. x < n}"
 | 
|
214  | 
  ultimately show "finite {x. x < Suc n}" by (simp)
 | 
|
215  | 
qed  | 
|
216  | 
||
217  | 
lemma finite_natarray2: "finite {pos. (fst pos) < (m::nat) & (snd pos) < (n::nat)}"
 | 
|
218  | 
apply (induct m)  | 
|
219  | 
apply (simp+)  | 
|
220  | 
proof -  | 
|
221  | 
fix m::nat  | 
|
222  | 
    let ?s0 = "{pos. fst pos < m & snd pos < n}"
 | 
|
223  | 
    let ?s1 = "{pos. fst pos < (Suc m) & snd pos < n}"
 | 
|
224  | 
    let ?sd = "{pos. fst pos = m & snd pos < n}"
 | 
|
225  | 
assume f0: "finite ?s0"  | 
|
226  | 
have f1: "finite ?sd"  | 
|
227  | 
proof -  | 
|
228  | 
let ?f = "% x. (m, x)"  | 
|
229  | 
      have "{pos. fst pos = m & snd pos < n} = ?f ` {x. x < n}" by (rule set_ext, simp add: image_def, auto)
 | 
|
230  | 
      moreover have "finite {x. x < n}" by (simp add: finite_natarray1)
 | 
|
231  | 
      ultimately show "finite {pos. fst pos = m & snd pos < n}" by (simp)
 | 
|
232  | 
qed  | 
|
233  | 
have su: "?s0 \<union> ?sd = ?s1" by (rule set_ext, simp, arith)  | 
|
234  | 
from f0 f1 have "finite (?s0 \<union> ?sd)" by (rule finite_UnI)  | 
|
235  | 
with su show "finite ?s1" by (simp)  | 
|
236  | 
qed  | 
|
237  | 
||
238  | 
lemma RepAbs_matrix:  | 
|
239  | 
assumes aem: "? m. ! j i. m <= j \<longrightarrow> x j i = 0" (is ?em) and aen:"? n. ! j i. (n <= i \<longrightarrow> x j i = 0)" (is ?en)  | 
|
240  | 
shows "(Rep_matrix (Abs_matrix x)) = x"  | 
|
241  | 
apply (rule Abs_matrix_inverse)  | 
|
242  | 
apply (simp add: matrix_def nonzero_positions_def)  | 
|
243  | 
proof -  | 
|
244  | 
from aem obtain m where a: "! j i. m <= j \<longrightarrow> x j i = 0" by (blast)  | 
|
245  | 
from aen obtain n where b: "! j i. n <= i \<longrightarrow> x j i = 0" by (blast)  | 
|
246  | 
  let ?u = "{pos. x (fst pos) (snd pos) \<noteq> 0}"
 | 
|
247  | 
  let ?v = "{pos. fst pos < m & snd pos < n}"
 | 
|
248  | 
have c: "!! (m::nat) a. ~(m <= a) \<Longrightarrow> a < m" by (arith)  | 
|
249  | 
  from a b have "(?u \<inter> (-?v)) = {}"
 | 
|
250  | 
apply (simp)  | 
|
251  | 
apply (rule set_ext)  | 
|
252  | 
apply (simp)  | 
|
253  | 
apply auto  | 
|
254  | 
by (rule c, auto)+  | 
|
255  | 
then have d: "?u \<subseteq> ?v" by blast  | 
|
256  | 
moreover have "finite ?v" by (simp add: finite_natarray2)  | 
|
257  | 
ultimately show "finite ?u" by (rule finite_subset)  | 
|
258  | 
qed  | 
|
259  | 
||
260  | 
constdefs  | 
|
261  | 
  apply_infmatrix :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a infmatrix \<Rightarrow> 'b infmatrix"
 | 
|
262  | 
"apply_infmatrix f == % A. (% j i. f (A j i))"  | 
|
263  | 
  apply_matrix :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a::zero) matrix \<Rightarrow> ('b::zero) matrix"
 | 
|
264  | 
"apply_matrix f == % A. Abs_matrix (apply_infmatrix f (Rep_matrix A))"  | 
|
265  | 
  combine_infmatrix :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a infmatrix \<Rightarrow> 'b infmatrix \<Rightarrow> 'c infmatrix"
 | 
|
266  | 
"combine_infmatrix f == % A B. (% j i. f (A j i) (B j i))"  | 
|
267  | 
  combine_matrix :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> ('a::zero) matrix \<Rightarrow> ('b::zero) matrix \<Rightarrow> ('c::zero) matrix"
 | 
|
268  | 
"combine_matrix f == % A B. Abs_matrix (combine_infmatrix f (Rep_matrix A) (Rep_matrix B))"  | 
|
269  | 
||
270  | 
lemma expand_apply_infmatrix[simp]: "apply_infmatrix f A j i = f (A j i)"  | 
|
271  | 
by (simp add: apply_infmatrix_def)  | 
|
272  | 
||
273  | 
lemma expand_combine_infmatrix[simp]: "combine_infmatrix f A B j i = f (A j i) (B j i)"  | 
|
274  | 
by (simp add: combine_infmatrix_def)  | 
|
275  | 
||
276  | 
constdefs  | 
|
277  | 
commutative :: "('a \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> bool"
 | 
|
278  | 
"commutative f == ! x y. f x y = f y x"  | 
|
279  | 
associative :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> bool"
 | 
|
280  | 
"associative f == ! x y z. f (f x y) z = f x (f y z)"  | 
|
281  | 
||
282  | 
text{*
 | 
|
283  | 
To reason about associativity and commutativity of operations on matrices,  | 
|
284  | 
let's take a step back and look at the general situtation: Assume that we have  | 
|
285  | 
sets $A$ and $B$ with $B \subset A$ and an abstraction $u: A \rightarrow B$. This abstraction has to fulfill $u(b) = b$ for all $b \in B$, but is arbitrary otherwise.  | 
|
286  | 
Each function $f: A \times A \rightarrow A$ now induces a function $f': B \times B \rightarrow B$ by $f' = u \circ f$.  | 
|
287  | 
It is obvious that commutativity of $f$ implies commutativity of $f'$: $f' x y = u (f x y) = u (f y x) = f' y x.$  | 
|
288  | 
*}  | 
|
289  | 
||
290  | 
lemma combine_infmatrix_commute:  | 
|
291  | 
"commutative f \<Longrightarrow> commutative (combine_infmatrix f)"  | 
|
292  | 
by (simp add: commutative_def combine_infmatrix_def)  | 
|
293  | 
||
294  | 
lemma combine_matrix_commute:  | 
|
295  | 
"commutative f \<Longrightarrow> commutative (combine_matrix f)"  | 
|
296  | 
by (simp add: combine_matrix_def commutative_def combine_infmatrix_def)  | 
|
297  | 
||
298  | 
text{*
 | 
|
299  | 
On the contrary, given an associative function $f$ we cannot expect $f'$ to be associative. A counterexample is given by $A=\ganz$, $B=\{-1, 0, 1\}$,
 | 
|
300  | 
as $f$ we take addition on $\ganz$, which is clearly associative. The abstraction is given by $u(a) = 0$ for $a \notin B$. Then we have  | 
|
301  | 
\[ f' (f' 1 1) -1 = u(f (u (f 1 1)) -1) = u(f (u 2) -1) = u (f 0 -1) = -1, \]  | 
|
302  | 
but on the other hand we have  | 
|
303  | 
\[ f' 1 (f' 1 -1) = u (f 1 (u (f 1 -1))) = u (f 1 0) = 1.\]  | 
|
304  | 
A way out of this problem is to assume that $f(A\times A)\subset A$ holds, and this is what we are going to do:  | 
|
305  | 
*}  | 
|
306  | 
||
307  | 
lemma nonzero_positions_combine_infmatrix[simp]: "f 0 0 = 0 \<Longrightarrow> nonzero_positions (combine_infmatrix f A B) \<subseteq> (nonzero_positions A) \<union> (nonzero_positions B)"  | 
|
308  | 
by (rule subsetI, simp add: nonzero_positions_def combine_infmatrix_def, auto)  | 
|
309  | 
||
310  | 
lemma finite_nonzero_positions_Rep[simp]: "finite (nonzero_positions (Rep_matrix A))"  | 
|
311  | 
by (insert Rep_matrix [of A], simp add: matrix_def)  | 
|
312  | 
||
313  | 
lemma combine_infmatrix_closed [simp]:  | 
|
314  | 
"f 0 0 = 0 \<Longrightarrow> Rep_matrix (Abs_matrix (combine_infmatrix f (Rep_matrix A) (Rep_matrix B))) = combine_infmatrix f (Rep_matrix A) (Rep_matrix B)"  | 
|
315  | 
apply (rule Abs_matrix_inverse)  | 
|
316  | 
apply (simp add: matrix_def)  | 
|
317  | 
apply (rule finite_subset[of _ "(nonzero_positions (Rep_matrix A)) \<union> (nonzero_positions (Rep_matrix B))"])  | 
|
318  | 
by (simp_all)  | 
|
319  | 
||
320  | 
text {* We need the next two lemmas only later, but it is analog to the above one, so we prove them now: *}
 | 
|
321  | 
lemma nonzero_positions_apply_infmatrix[simp]: "f 0 = 0 \<Longrightarrow> nonzero_positions (apply_infmatrix f A) \<subseteq> nonzero_positions A"  | 
|
322  | 
by (rule subsetI, simp add: nonzero_positions_def apply_infmatrix_def, auto)  | 
|
323  | 
||
324  | 
lemma apply_infmatrix_closed [simp]:  | 
|
325  | 
"f 0 = 0 \<Longrightarrow> Rep_matrix (Abs_matrix (apply_infmatrix f (Rep_matrix A))) = apply_infmatrix f (Rep_matrix A)"  | 
|
326  | 
apply (rule Abs_matrix_inverse)  | 
|
327  | 
apply (simp add: matrix_def)  | 
|
328  | 
apply (rule finite_subset[of _ "nonzero_positions (Rep_matrix A)"])  | 
|
329  | 
by (simp_all)  | 
|
330  | 
||
331  | 
lemma combine_infmatrix_assoc[simp]: "f 0 0 = 0 \<Longrightarrow> associative f \<Longrightarrow> associative (combine_infmatrix f)"  | 
|
332  | 
by (simp add: associative_def combine_infmatrix_def)  | 
|
333  | 
||
334  | 
lemma comb: "f = g \<Longrightarrow> x = y \<Longrightarrow> f x = g y"  | 
|
335  | 
by (auto)  | 
|
336  | 
||
337  | 
lemma combine_matrix_assoc: "f 0 0 = 0 \<Longrightarrow> associative f \<Longrightarrow> associative (combine_matrix f)"  | 
|
338  | 
apply (simp(no_asm) add: associative_def combine_matrix_def, auto)  | 
|
339  | 
apply (rule comb [of Abs_matrix Abs_matrix])  | 
|
340  | 
by (auto, insert combine_infmatrix_assoc[of f], simp add: associative_def)  | 
|
341  | 
||
342  | 
lemma Rep_apply_matrix[simp]: "f 0 = 0 \<Longrightarrow> Rep_matrix (apply_matrix f A) j i = f (Rep_matrix A j i)"  | 
|
343  | 
by (simp add: apply_matrix_def)  | 
|
344  | 
||
345  | 
lemma Rep_combine_matrix[simp]: "f 0 0 = 0 \<Longrightarrow> Rep_matrix (combine_matrix f A B) j i = f (Rep_matrix A j i) (Rep_matrix B j i)"  | 
|
346  | 
by(simp add: combine_matrix_def)  | 
|
347  | 
||
348  | 
lemma combine_nrows_max: "f 0 0 = 0 \<Longrightarrow> nrows (combine_matrix f A B) <= max (nrows A) (nrows B)"  | 
|
349  | 
by (simp add: nrows_le)  | 
|
350  | 
||
351  | 
lemma combine_ncols_max: "f 0 0 = 0 \<Longrightarrow> ncols (combine_matrix f A B) <= max (ncols A) (ncols B)"  | 
|
352  | 
by (simp add: ncols_le)  | 
|
353  | 
||
354  | 
lemma combine_nrows: "f 0 0 = 0 \<Longrightarrow> nrows A <= q \<Longrightarrow> nrows B <= q \<Longrightarrow> nrows(combine_matrix f A B) <= q"  | 
|
355  | 
by (simp add: nrows_le)  | 
|
356  | 
||
357  | 
lemma combine_ncols: "f 0 0 = 0 \<Longrightarrow> ncols A <= q \<Longrightarrow> ncols B <= q \<Longrightarrow> ncols(combine_matrix f A B) <= q"  | 
|
358  | 
by (simp add: ncols_le)  | 
|
359  | 
||
360  | 
constdefs  | 
|
361  | 
  zero_r_neutral :: "('a \<Rightarrow> 'b::zero \<Rightarrow> 'a) \<Rightarrow> bool"
 | 
|
362  | 
"zero_r_neutral f == ! a. f a 0 = a"  | 
|
363  | 
  zero_l_neutral :: "('a::zero \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> bool"
 | 
|
364  | 
"zero_l_neutral f == ! a. f 0 a = a"  | 
|
365  | 
  zero_closed :: "(('a::zero) \<Rightarrow> ('b::zero) \<Rightarrow> ('c::zero)) \<Rightarrow> bool"
 | 
|
366  | 
"zero_closed f == (!x. f x 0 = 0) & (!y. f 0 y = 0)"  | 
|
367  | 
||
368  | 
consts foldseq :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> (nat \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a"
 | 
|
369  | 
primrec  | 
|
370  | 
"foldseq f s 0 = s 0"  | 
|
371  | 
"foldseq f s (Suc n) = f (s 0) (foldseq f (% k. s(Suc k)) n)"  | 
|
372  | 
||
373  | 
consts foldseq_transposed ::  "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> (nat \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a"
 | 
|
374  | 
primrec  | 
|
375  | 
"foldseq_transposed f s 0 = s 0"  | 
|
376  | 
"foldseq_transposed f s (Suc n) = f (foldseq_transposed f s n) (s (Suc n))"  | 
|
377  | 
||
378  | 
lemma foldseq_assoc : "associative f \<Longrightarrow> foldseq f = foldseq_transposed f"  | 
|
379  | 
proof -  | 
|
380  | 
assume a:"associative f"  | 
|
381  | 
then have sublemma: "!! n. ! N s. N <= n \<longrightarrow> foldseq f s N = foldseq_transposed f s N"  | 
|
382  | 
proof -  | 
|
383  | 
fix n  | 
|
384  | 
show "!N s. N <= n \<longrightarrow> foldseq f s N = foldseq_transposed f s N"  | 
|
385  | 
proof (induct n)  | 
|
386  | 
show "!N s. N <= 0 \<longrightarrow> foldseq f s N = foldseq_transposed f s N" by simp  | 
|
387  | 
next  | 
|
388  | 
fix n  | 
|
389  | 
assume b:"! N s. N <= n \<longrightarrow> foldseq f s N = foldseq_transposed f s N"  | 
|
390  | 
have c:"!!N s. N <= n \<Longrightarrow> foldseq f s N = foldseq_transposed f s N" by (simp add: b)  | 
|
391  | 
show "! N t. N <= Suc n \<longrightarrow> foldseq f t N = foldseq_transposed f t N"  | 
|
392  | 
proof (auto)  | 
|
393  | 
fix N t  | 
|
394  | 
assume Nsuc: "N <= Suc n"  | 
|
395  | 
show "foldseq f t N = foldseq_transposed f t N"  | 
|
396  | 
proof cases  | 
|
397  | 
assume "N <= n"  | 
|
398  | 
then show "foldseq f t N = foldseq_transposed f t N" by (simp add: b)  | 
|
399  | 
next  | 
|
400  | 
assume "~(N <= n)"  | 
|
401  | 
with Nsuc have Nsuceq: "N = Suc n" by simp  | 
|
402  | 
have neqz: "n \<noteq> 0 \<Longrightarrow> ? m. n = Suc m & Suc m <= n" by arith  | 
|
403  | 
have assocf: "!! x y z. f x (f y z) = f (f x y) z" by (insert a, simp add: associative_def)  | 
|
404  | 
show "foldseq f t N = foldseq_transposed f t N"  | 
|
405  | 
apply (simp add: Nsuceq)  | 
|
406  | 
apply (subst c)  | 
|
407  | 
apply (simp)  | 
|
408  | 
apply (case_tac "n = 0")  | 
|
409  | 
apply (simp)  | 
|
410  | 
apply (drule neqz)  | 
|
411  | 
apply (erule exE)  | 
|
412  | 
apply (simp)  | 
|
413  | 
apply (subst assocf)  | 
|
414  | 
proof -  | 
|
415  | 
fix m  | 
|
416  | 
assume "n = Suc m & Suc m <= n"  | 
|
417  | 
then have mless: "Suc m <= n" by arith  | 
|
418  | 
then have step1: "foldseq_transposed f (% k. t (Suc k)) m = foldseq f (% k. t (Suc k)) m" (is "?T1 = ?T2")  | 
|
419  | 
apply (subst c)  | 
|
420  | 
by simp+  | 
|
421  | 
have step2: "f (t 0) ?T2 = foldseq f t (Suc m)" (is "_ = ?T3") by simp  | 
|
422  | 
have step3: "?T3 = foldseq_transposed f t (Suc m)" (is "_ = ?T4")  | 
|
423  | 
apply (subst c)  | 
|
424  | 
by (simp add: mless)+  | 
|
425  | 
have step4: "?T4 = f (foldseq_transposed f t m) (t (Suc m))" (is "_=?T5") by simp  | 
|
426  | 
from step1 step2 step3 step4 show sowhat: "f (f (t 0) ?T1) (t (Suc (Suc m))) = f ?T5 (t (Suc (Suc m)))" by simp  | 
|
427  | 
qed  | 
|
428  | 
qed  | 
|
429  | 
qed  | 
|
430  | 
qed  | 
|
431  | 
qed  | 
|
432  | 
show "foldseq f = foldseq_transposed f" by ((rule ext)+, insert sublemma, auto)  | 
|
433  | 
qed  | 
|
434  | 
||
435  | 
lemma foldseq_distr: "\<lbrakk>associative f; commutative f\<rbrakk> \<Longrightarrow> foldseq f (% k. f (u k) (v k)) n = f (foldseq f u n) (foldseq f v n)"  | 
|
436  | 
proof -  | 
|
437  | 
assume assoc: "associative f"  | 
|
438  | 
assume comm: "commutative f"  | 
|
439  | 
from assoc have a:"!! x y z. f (f x y) z = f x (f y z)" by (simp add: associative_def)  | 
|
440  | 
from comm have b: "!! x y. f x y = f y x" by (simp add: commutative_def)  | 
|
441  | 
from assoc comm have c: "!! x y z. f x (f y z) = f y (f x z)" by (simp add: commutative_def associative_def)  | 
|
442  | 
have "!! n. (! u v. foldseq f (%k. f (u k) (v k)) n = f (foldseq f u n) (foldseq f v n))"  | 
|
443  | 
apply (induct_tac n)  | 
|
444  | 
apply (simp+, auto)  | 
|
445  | 
by (simp add: a b c)  | 
|
446  | 
then show "foldseq f (% k. f (u k) (v k)) n = f (foldseq f u n) (foldseq f v n)" by simp  | 
|
447  | 
qed  | 
|
448  | 
||
449  | 
theorem "\<lbrakk>associative f; associative g; \<forall>a b c d. g (f a b) (f c d) = f (g a c) (g b d); ? x y. (f x) \<noteq> (f y); ? x y. (g x) \<noteq> (g y); f x x = x; g x x = x\<rbrakk> \<Longrightarrow> f=g | (! y. f y x = y) | (! y. g y x = y)"  | 
|
450  | 
oops  | 
|
451  | 
(* Model found  | 
|
452  | 
||
453  | 
Trying to find a model that refutes: \<lbrakk>associative f; associative g;  | 
|
454  | 
\<forall>a b c d. g (f a b) (f c d) = f (g a c) (g b d); \<exists>x y. f x \<noteq> f y;  | 
|
455  | 
\<exists>x y. g x \<noteq> g y; f x x = x; g x x = x\<rbrakk>  | 
|
456  | 
\<Longrightarrow> f = g \<or> (\<forall>y. f y x = y) \<or> (\<forall>y. g y x = y)  | 
|
457  | 
Searching for a model of size 1, translating term... invoking SAT solver... no model found.  | 
|
458  | 
Searching for a model of size 2, translating term... invoking SAT solver... no model found.  | 
|
459  | 
Searching for a model of size 3, translating term... invoking SAT solver...  | 
|
460  | 
Model found:  | 
|
461  | 
Size of types: 'a: 3  | 
|
462  | 
x: a1  | 
|
463  | 
g: (a0\<mapsto>(a0\<mapsto>a1, a1\<mapsto>a0, a2\<mapsto>a1), a1\<mapsto>(a0\<mapsto>a0, a1\<mapsto>a1, a2\<mapsto>a0), a2\<mapsto>(a0\<mapsto>a1, a1\<mapsto>a0, a2\<mapsto>a1))  | 
|
464  | 
f: (a0\<mapsto>(a0\<mapsto>a0, a1\<mapsto>a0, a2\<mapsto>a0), a1\<mapsto>(a0\<mapsto>a1, a1\<mapsto>a1, a2\<mapsto>a1), a2\<mapsto>(a0\<mapsto>a0, a1\<mapsto>a0, a2\<mapsto>a0))  | 
|
465  | 
*)  | 
|
466  | 
||
467  | 
lemma foldseq_zero:  | 
|
468  | 
assumes fz: "f 0 0 = 0" and sz: "! i. i <= n \<longrightarrow> s i = 0"  | 
|
469  | 
shows "foldseq f s n = 0"  | 
|
470  | 
proof -  | 
|
471  | 
have "!! n. ! s. (! i. i <= n \<longrightarrow> s i = 0) \<longrightarrow> foldseq f s n = 0"  | 
|
472  | 
apply (induct_tac n)  | 
|
473  | 
apply (simp)  | 
|
474  | 
by (simp add: fz)  | 
|
475  | 
then show "foldseq f s n = 0" by (simp add: sz)  | 
|
476  | 
qed  | 
|
477  | 
||
478  | 
lemma foldseq_significant_positions:  | 
|
479  | 
assumes p: "! i. i <= N \<longrightarrow> S i = T i"  | 
|
480  | 
shows "foldseq f S N = foldseq f T N" (is ?concl)  | 
|
481  | 
proof -  | 
|
482  | 
have "!! m . ! s t. (! i. i<=m \<longrightarrow> s i = t i) \<longrightarrow> foldseq f s m = foldseq f t m"  | 
|
483  | 
apply (induct_tac m)  | 
|
484  | 
apply (simp)  | 
|
485  | 
apply (simp)  | 
|
486  | 
apply (auto)  | 
|
487  | 
proof -  | 
|
488  | 
fix n  | 
|
489  | 
fix s::"nat\<Rightarrow>'a"  | 
|
490  | 
fix t::"nat\<Rightarrow>'a"  | 
|
491  | 
assume a: "\<forall>s t. (\<forall>i\<le>n. s i = t i) \<longrightarrow> foldseq f s n = foldseq f t n"  | 
|
492  | 
assume b: "\<forall>i\<le>Suc n. s i = t i"  | 
|
493  | 
have c:"!! a b. a = b \<Longrightarrow> f (t 0) a = f (t 0) b" by blast  | 
|
494  | 
have d:"!! s t. (\<forall>i\<le>n. s i = t i) \<Longrightarrow> foldseq f s n = foldseq f t n" by (simp add: a)  | 
|
495  | 
show "f (t 0) (foldseq f (\<lambda>k. s (Suc k)) n) = f (t 0) (foldseq f (\<lambda>k. t (Suc k)) n)" by (rule c, simp add: d b)  | 
|
496  | 
qed  | 
|
497  | 
with p show ?concl by simp  | 
|
498  | 
qed  | 
|
499  | 
||
500  | 
lemma foldseq_tail: "M <= N \<Longrightarrow> foldseq f S N = foldseq f (% k. (if k < M then (S k) else (foldseq f (% k. S(k+M)) (N-M)))) M" (is "?p \<Longrightarrow> ?concl")  | 
|
501  | 
proof -  | 
|
502  | 
have suc: "!! a b. \<lbrakk>a <= Suc b; a \<noteq> Suc b\<rbrakk> \<Longrightarrow> a <= b" by arith  | 
|
503  | 
have a:"!! a b c . a = b \<Longrightarrow> f c a = f c b" by blast  | 
|
504  | 
have "!! n. ! m s. m <= n \<longrightarrow> foldseq f s n = foldseq f (% k. (if k < m then (s k) else (foldseq f (% k. s(k+m)) (n-m)))) m"  | 
|
505  | 
apply (induct_tac n)  | 
|
506  | 
apply (simp)  | 
|
507  | 
apply (simp)  | 
|
508  | 
apply (auto)  | 
|
509  | 
apply (case_tac "m = Suc na")  | 
|
510  | 
apply (simp)  | 
|
511  | 
apply (rule a)  | 
|
512  | 
apply (rule foldseq_significant_positions)  | 
|
513  | 
apply (auto)  | 
|
514  | 
apply (drule suc, simp+)  | 
|
515  | 
proof -  | 
|
516  | 
fix na m s  | 
|
517  | 
assume suba:"\<forall>m\<le>na. \<forall>s. foldseq f s na = foldseq f (\<lambda>k. if k < m then s k else foldseq f (\<lambda>k. s (k + m)) (na - m))m"  | 
|
518  | 
assume subb:"m <= na"  | 
|
519  | 
from suba have subc:"!! m s. m <= na \<Longrightarrow>foldseq f s na = foldseq f (\<lambda>k. if k < m then s k else foldseq f (\<lambda>k. s (k + m)) (na - m))m" by simp  | 
|
520  | 
have subd: "foldseq f (\<lambda>k. if k < m then s (Suc k) else foldseq f (\<lambda>k. s (Suc (k + m))) (na - m)) m =  | 
|
521  | 
foldseq f (% k. s(Suc k)) na"  | 
|
522  | 
by (rule subc[of m "% k. s(Suc k)", THEN sym], simp add: subb)  | 
|
523  | 
from subb have sube: "m \<noteq> 0 \<Longrightarrow> ? mm. m = Suc mm & mm <= na" by arith  | 
|
524  | 
show "f (s 0) (foldseq f (\<lambda>k. if k < m then s (Suc k) else foldseq f (\<lambda>k. s (Suc (k + m))) (na - m)) m) =  | 
|
525  | 
foldseq f (\<lambda>k. if k < m then s k else foldseq f (\<lambda>k. s (k + m)) (Suc na - m)) m"  | 
|
526  | 
apply (simp add: subd)  | 
|
527  | 
apply (case_tac "m=0")  | 
|
528  | 
apply (simp)  | 
|
529  | 
apply (drule sube)  | 
|
530  | 
apply (auto)  | 
|
531  | 
apply (rule a)  | 
|
532  | 
by (simp add: subc if_def)  | 
|
533  | 
qed  | 
|
534  | 
then show "?p \<Longrightarrow> ?concl" by simp  | 
|
535  | 
qed  | 
|
536  | 
||
537  | 
lemma foldseq_zerotail:  | 
|
538  | 
assumes  | 
|
539  | 
fz: "f 0 0 = 0"  | 
|
540  | 
and sz: "! i. n <= i \<longrightarrow> s i = 0"  | 
|
541  | 
and nm: "n <= m"  | 
|
542  | 
shows  | 
|
543  | 
"foldseq f s n = foldseq f s m"  | 
|
544  | 
proof -  | 
|
545  | 
show "foldseq f s n = foldseq f s m"  | 
|
546  | 
apply (simp add: foldseq_tail[OF nm, of f s])  | 
|
547  | 
apply (rule foldseq_significant_positions)  | 
|
548  | 
apply (auto)  | 
|
549  | 
apply (subst foldseq_zero)  | 
|
550  | 
by (simp add: fz sz)+  | 
|
551  | 
qed  | 
|
552  | 
||
553  | 
lemma foldseq_zerotail2:  | 
|
554  | 
assumes "! x. f x 0 = x"  | 
|
555  | 
and "! i. n < i \<longrightarrow> s i = 0"  | 
|
556  | 
and nm: "n <= m"  | 
|
557  | 
shows  | 
|
558  | 
"foldseq f s n = foldseq f s m" (is ?concl)  | 
|
559  | 
proof -  | 
|
560  | 
have "f 0 0 = 0" by (simp add: prems)  | 
|
561  | 
have b:"!! m n. n <= m \<Longrightarrow> m \<noteq> n \<Longrightarrow> ? k. m-n = Suc k" by arith  | 
|
562  | 
have c: "0 <= m" by simp  | 
|
563  | 
have d: "!! k. k \<noteq> 0 \<Longrightarrow> ? l. k = Suc l" by arith  | 
|
564  | 
show ?concl  | 
|
565  | 
apply (subst foldseq_tail[OF nm])  | 
|
566  | 
apply (rule foldseq_significant_positions)  | 
|
567  | 
apply (auto)  | 
|
568  | 
apply (case_tac "m=n")  | 
|
569  | 
apply (simp+)  | 
|
570  | 
apply (drule b[OF nm])  | 
|
571  | 
apply (auto)  | 
|
572  | 
apply (case_tac "k=0")  | 
|
573  | 
apply (simp add: prems)  | 
|
574  | 
apply (drule d)  | 
|
575  | 
apply (auto)  | 
|
576  | 
by (simp add: prems foldseq_zero)  | 
|
577  | 
qed  | 
|
578  | 
||
579  | 
lemma foldseq_zerostart:  | 
|
580  | 
"! x. f 0 (f 0 x) = f 0 x \<Longrightarrow> ! i. i <= n \<longrightarrow> s i = 0 \<Longrightarrow> foldseq f s (Suc n) = f 0 (s (Suc n))"  | 
|
581  | 
proof -  | 
|
582  | 
assume f00x: "! x. f 0 (f 0 x) = f 0 x"  | 
|
583  | 
have "! s. (! i. i<=n \<longrightarrow> s i = 0) \<longrightarrow> foldseq f s (Suc n) = f 0 (s (Suc n))"  | 
|
584  | 
apply (induct n)  | 
|
585  | 
apply (simp)  | 
|
586  | 
apply (rule allI, rule impI)  | 
|
587  | 
proof -  | 
|
588  | 
fix n  | 
|
589  | 
fix s  | 
|
590  | 
have a:"foldseq f s (Suc (Suc n)) = f (s 0) (foldseq f (% k. s(Suc k)) (Suc n))" by simp  | 
|
591  | 
assume b: "! s. ((\<forall>i\<le>n. s i = 0) \<longrightarrow> foldseq f s (Suc n) = f 0 (s (Suc n)))"  | 
|
592  | 
from b have c:"!! s. (\<forall>i\<le>n. s i = 0) \<Longrightarrow> foldseq f s (Suc n) = f 0 (s (Suc n))" by simp  | 
|
593  | 
assume d: "! i. i <= Suc n \<longrightarrow> s i = 0"  | 
|
594  | 
show "foldseq f s (Suc (Suc n)) = f 0 (s (Suc (Suc n)))"  | 
|
595  | 
apply (subst a)  | 
|
596  | 
apply (subst c)  | 
|
597  | 
by (simp add: d f00x)+  | 
|
598  | 
qed  | 
|
599  | 
then show "! i. i <= n \<longrightarrow> s i = 0 \<Longrightarrow> foldseq f s (Suc n) = f 0 (s (Suc n))" by simp  | 
|
600  | 
qed  | 
|
601  | 
||
602  | 
lemma foldseq_zerostart2:  | 
|
603  | 
"! x. f 0 x = x \<Longrightarrow> ! i. i < n \<longrightarrow> s i = 0 \<Longrightarrow> foldseq f s n = s n"  | 
|
604  | 
proof -  | 
|
605  | 
assume a:"! i. i<n \<longrightarrow> s i = 0"  | 
|
606  | 
assume x:"! x. f 0 x = x"  | 
|
607  | 
from x have f00x: "! x. f 0 (f 0 x) = f 0 x" by blast  | 
|
608  | 
have b: "!! i l. i < Suc l = (i <= l)" by arith  | 
|
609  | 
have d: "!! k. k \<noteq> 0 \<Longrightarrow> ? l. k = Suc l" by arith  | 
|
610  | 
show "foldseq f s n = s n"  | 
|
611  | 
apply (case_tac "n=0")  | 
|
612  | 
apply (simp)  | 
|
613  | 
apply (insert a)  | 
|
614  | 
apply (drule d)  | 
|
615  | 
apply (auto)  | 
|
616  | 
apply (simp add: b)  | 
|
617  | 
apply (insert f00x)  | 
|
618  | 
apply (drule foldseq_zerostart)  | 
|
619  | 
by (simp add: x)+  | 
|
620  | 
qed  | 
|
621  | 
||
622  | 
lemma foldseq_almostzero:  | 
|
623  | 
assumes f0x:"! x. f 0 x = x" and fx0: "! x. f x 0 = x" and s0:"! i. i \<noteq> j \<longrightarrow> s i = 0"  | 
|
624  | 
shows "foldseq f s n = (if (j <= n) then (s j) else 0)" (is ?concl)  | 
|
625  | 
proof -  | 
|
626  | 
from s0 have a: "! i. i < j \<longrightarrow> s i = 0" by simp  | 
|
627  | 
from s0 have b: "! i. j < i \<longrightarrow> s i = 0" by simp  | 
|
628  | 
show ?concl  | 
|
629  | 
apply auto  | 
|
630  | 
apply (subst foldseq_zerotail2[of f, OF fx0, of j, OF b, of n, THEN sym])  | 
|
631  | 
apply simp  | 
|
632  | 
apply (subst foldseq_zerostart2)  | 
|
633  | 
apply (simp add: f0x a)+  | 
|
634  | 
apply (subst foldseq_zero)  | 
|
635  | 
by (simp add: s0 f0x)+  | 
|
636  | 
qed  | 
|
637  | 
||
638  | 
lemma foldseq_distr_unary:  | 
|
639  | 
assumes "!! a b. g (f a b) = f (g a) (g b)"  | 
|
640  | 
shows "g(foldseq f s n) = foldseq f (% x. g(s x)) n" (is ?concl)  | 
|
641  | 
proof -  | 
|
642  | 
have "! s. g(foldseq f s n) = foldseq f (% x. g(s x)) n"  | 
|
643  | 
apply (induct_tac n)  | 
|
644  | 
apply (simp)  | 
|
645  | 
apply (simp)  | 
|
646  | 
apply (auto)  | 
|
647  | 
apply (drule_tac x="% k. s (Suc k)" in spec)  | 
|
648  | 
by (simp add: prems)  | 
|
649  | 
then show ?concl by simp  | 
|
650  | 
qed  | 
|
651  | 
||
652  | 
constdefs  | 
|
653  | 
  mult_matrix_n :: "nat \<Rightarrow> (('a::zero) \<Rightarrow> ('b::zero) \<Rightarrow> ('c::zero)) \<Rightarrow> ('c \<Rightarrow> 'c \<Rightarrow> 'c) \<Rightarrow> 'a matrix \<Rightarrow> 'b matrix \<Rightarrow> 'c matrix"
 | 
|
654  | 
"mult_matrix_n n fmul fadd A B == Abs_matrix(% j i. foldseq fadd (% k. fmul (Rep_matrix A j k) (Rep_matrix B k i)) n)"  | 
|
655  | 
  mult_matrix :: "(('a::zero) \<Rightarrow> ('b::zero) \<Rightarrow> ('c::zero)) \<Rightarrow> ('c \<Rightarrow> 'c \<Rightarrow> 'c) \<Rightarrow> 'a matrix \<Rightarrow> 'b matrix \<Rightarrow> 'c matrix"
 | 
|
656  | 
"mult_matrix fmul fadd A B == mult_matrix_n (max (ncols A) (nrows B)) fmul fadd A B"  | 
|
657  | 
||
658  | 
lemma mult_matrix_n:  | 
|
659  | 
assumes prems: "ncols A \<le> n" (is ?An) "nrows B \<le> n" (is ?Bn) "fadd 0 0 = 0" "fmul 0 0 = 0"  | 
|
660  | 
shows c:"mult_matrix fmul fadd A B = mult_matrix_n n fmul fadd A B" (is ?concl)  | 
|
661  | 
proof -  | 
|
662  | 
show ?concl using prems  | 
|
663  | 
apply (simp add: mult_matrix_def mult_matrix_n_def)  | 
|
664  | 
apply (rule comb[of "Abs_matrix" "Abs_matrix"], simp, (rule ext)+)  | 
|
665  | 
by (rule foldseq_zerotail, simp_all add: nrows_le ncols_le prems)  | 
|
666  | 
qed  | 
|
667  | 
||
668  | 
lemma mult_matrix_nm:  | 
|
669  | 
assumes prems: "ncols A <= n" "nrows B <= n" "ncols A <= m" "nrows B <= m" "fadd 0 0 = 0" "fmul 0 0 = 0"  | 
|
670  | 
shows "mult_matrix_n n fmul fadd A B = mult_matrix_n m fmul fadd A B"  | 
|
671  | 
proof -  | 
|
672  | 
from prems have "mult_matrix_n n fmul fadd A B = mult_matrix fmul fadd A B" by (simp add: mult_matrix_n)  | 
|
673  | 
also from prems have "\<dots> = mult_matrix_n m fmul fadd A B" by (simp add: mult_matrix_n[THEN sym])  | 
|
674  | 
finally show "mult_matrix_n n fmul fadd A B = mult_matrix_n m fmul fadd A B" by simp  | 
|
675  | 
qed  | 
|
676  | 
||
677  | 
constdefs  | 
|
678  | 
  r_distributive :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> bool"
 | 
|
679  | 
"r_distributive fmul fadd == ! a u v. fmul a (fadd u v) = fadd (fmul a u) (fmul a v)"  | 
|
680  | 
  l_distributive :: "('a \<Rightarrow> 'b \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> bool"
 | 
|
681  | 
"l_distributive fmul fadd == ! a u v. fmul (fadd u v) a = fadd (fmul u a) (fmul v a)"  | 
|
682  | 
  distributive :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> bool"
 | 
|
683  | 
"distributive fmul fadd == l_distributive fmul fadd & r_distributive fmul fadd"  | 
|
684  | 
||
685  | 
lemma max1: "!! a x y. (a::nat) <= x \<Longrightarrow> a <= max x y" by (arith)  | 
|
686  | 
lemma max2: "!! b x y. (b::nat) <= y \<Longrightarrow> b <= max x y" by (arith)  | 
|
687  | 
||
688  | 
lemma r_distributive_matrix:  | 
|
689  | 
assumes prems:  | 
|
690  | 
"r_distributive fmul fadd"  | 
|
691  | 
"associative fadd"  | 
|
692  | 
"commutative fadd"  | 
|
693  | 
"fadd 0 0 = 0"  | 
|
694  | 
"! a. fmul a 0 = 0"  | 
|
695  | 
"! a. fmul 0 a = 0"  | 
|
696  | 
shows "r_distributive (mult_matrix fmul fadd) (combine_matrix fadd)" (is ?concl)  | 
|
697  | 
proof -  | 
|
698  | 
from prems show ?concl  | 
|
699  | 
apply (simp add: r_distributive_def mult_matrix_def, auto)  | 
|
700  | 
proof -  | 
|
701  | 
fix a::"'a matrix"  | 
|
702  | 
fix u::"'b matrix"  | 
|
703  | 
fix v::"'b matrix"  | 
|
704  | 
let ?mx = "max (ncols a) (max (nrows u) (nrows v))"  | 
|
705  | 
from prems show "mult_matrix_n (max (ncols a) (nrows (combine_matrix fadd u v))) fmul fadd a (combine_matrix fadd u v) =  | 
|
706  | 
combine_matrix fadd (mult_matrix_n (max (ncols a) (nrows u)) fmul fadd a u) (mult_matrix_n (max (ncols a) (nrows v)) fmul fadd a v)"  | 
|
707  | 
apply (subst mult_matrix_nm[of _ _ _ ?mx fadd fmul])  | 
|
708  | 
apply (simp add: max1 max2 combine_nrows combine_ncols)+  | 
|
709  | 
apply (subst mult_matrix_nm[of _ _ v ?mx fadd fmul])  | 
|
710  | 
apply (simp add: max1 max2 combine_nrows combine_ncols)+  | 
|
711  | 
apply (subst mult_matrix_nm[of _ _ u ?mx fadd fmul])  | 
|
712  | 
apply (simp add: max1 max2 combine_nrows combine_ncols)+  | 
|
713  | 
apply (simp add: mult_matrix_n_def r_distributive_def foldseq_distr[of fadd])  | 
|
714  | 
apply (simp add: combine_matrix_def combine_infmatrix_def)  | 
|
715  | 
apply (rule comb[of "Abs_matrix" "Abs_matrix"], simp, (rule ext)+)  | 
|
716  | 
apply (simplesubst RepAbs_matrix)  | 
|
717  | 
apply (simp, auto)  | 
|
718  | 
apply (rule exI[of _ "nrows a"], simp add: nrows_le foldseq_zero)  | 
|
719  | 
apply (rule exI[of _ "ncols v"], simp add: ncols_le foldseq_zero)  | 
|
720  | 
apply (subst RepAbs_matrix)  | 
|
721  | 
apply (simp, auto)  | 
|
722  | 
apply (rule exI[of _ "nrows a"], simp add: nrows_le foldseq_zero)  | 
|
723  | 
apply (rule exI[of _ "ncols u"], simp add: ncols_le foldseq_zero)  | 
|
724  | 
done  | 
|
725  | 
qed  | 
|
726  | 
qed  | 
|
727  | 
||
728  | 
lemma l_distributive_matrix:  | 
|
729  | 
assumes prems:  | 
|
730  | 
"l_distributive fmul fadd"  | 
|
731  | 
"associative fadd"  | 
|
732  | 
"commutative fadd"  | 
|
733  | 
"fadd 0 0 = 0"  | 
|
734  | 
"! a. fmul a 0 = 0"  | 
|
735  | 
"! a. fmul 0 a = 0"  | 
|
736  | 
shows "l_distributive (mult_matrix fmul fadd) (combine_matrix fadd)" (is ?concl)  | 
|
737  | 
proof -  | 
|
738  | 
from prems show ?concl  | 
|
739  | 
apply (simp add: l_distributive_def mult_matrix_def, auto)  | 
|
740  | 
proof -  | 
|
741  | 
fix a::"'b matrix"  | 
|
742  | 
fix u::"'a matrix"  | 
|
743  | 
fix v::"'a matrix"  | 
|
744  | 
let ?mx = "max (nrows a) (max (ncols u) (ncols v))"  | 
|
745  | 
from prems show "mult_matrix_n (max (ncols (combine_matrix fadd u v)) (nrows a)) fmul fadd (combine_matrix fadd u v) a =  | 
|
746  | 
combine_matrix fadd (mult_matrix_n (max (ncols u) (nrows a)) fmul fadd u a) (mult_matrix_n (max (ncols v) (nrows a)) fmul fadd v a)"  | 
|
747  | 
apply (subst mult_matrix_nm[of v _ _ ?mx fadd fmul])  | 
|
748  | 
apply (simp add: max1 max2 combine_nrows combine_ncols)+  | 
|
749  | 
apply (subst mult_matrix_nm[of u _ _ ?mx fadd fmul])  | 
|
750  | 
apply (simp add: max1 max2 combine_nrows combine_ncols)+  | 
|
751  | 
apply (subst mult_matrix_nm[of _ _ _ ?mx fadd fmul])  | 
|
752  | 
apply (simp add: max1 max2 combine_nrows combine_ncols)+  | 
|
753  | 
apply (simp add: mult_matrix_n_def l_distributive_def foldseq_distr[of fadd])  | 
|
754  | 
apply (simp add: combine_matrix_def combine_infmatrix_def)  | 
|
755  | 
apply (rule comb[of "Abs_matrix" "Abs_matrix"], simp, (rule ext)+)  | 
|
756  | 
apply (simplesubst RepAbs_matrix)  | 
|
757  | 
apply (simp, auto)  | 
|
758  | 
apply (rule exI[of _ "nrows v"], simp add: nrows_le foldseq_zero)  | 
|
759  | 
apply (rule exI[of _ "ncols a"], simp add: ncols_le foldseq_zero)  | 
|
760  | 
apply (subst RepAbs_matrix)  | 
|
761  | 
apply (simp, auto)  | 
|
762  | 
apply (rule exI[of _ "nrows u"], simp add: nrows_le foldseq_zero)  | 
|
763  | 
apply (rule exI[of _ "ncols a"], simp add: ncols_le foldseq_zero)  | 
|
764  | 
done  | 
|
765  | 
qed  | 
|
766  | 
qed  | 
|
767  | 
||
768  | 
instantiation matrix :: (zero) zero  | 
|
769  | 
begin  | 
|
770  | 
||
| 28562 | 771  | 
definition zero_matrix_def [code del]: "0 = Abs_matrix (\<lambda>j i. 0)"  | 
| 27484 | 772  | 
|
773  | 
instance ..  | 
|
774  | 
||
775  | 
end  | 
|
776  | 
||
777  | 
lemma Rep_zero_matrix_def[simp]: "Rep_matrix 0 j i = 0"  | 
|
778  | 
apply (simp add: zero_matrix_def)  | 
|
779  | 
apply (subst RepAbs_matrix)  | 
|
780  | 
by (auto)  | 
|
781  | 
||
782  | 
lemma zero_matrix_def_nrows[simp]: "nrows 0 = 0"  | 
|
783  | 
proof -  | 
|
784  | 
have a:"!! (x::nat). x <= 0 \<Longrightarrow> x = 0" by (arith)  | 
|
785  | 
show "nrows 0 = 0" by (rule a, subst nrows_le, simp)  | 
|
786  | 
qed  | 
|
787  | 
||
788  | 
lemma zero_matrix_def_ncols[simp]: "ncols 0 = 0"  | 
|
789  | 
proof -  | 
|
790  | 
have a:"!! (x::nat). x <= 0 \<Longrightarrow> x = 0" by (arith)  | 
|
791  | 
show "ncols 0 = 0" by (rule a, subst ncols_le, simp)  | 
|
792  | 
qed  | 
|
793  | 
||
794  | 
lemma combine_matrix_zero_l_neutral: "zero_l_neutral f \<Longrightarrow> zero_l_neutral (combine_matrix f)"  | 
|
795  | 
by (simp add: zero_l_neutral_def combine_matrix_def combine_infmatrix_def)  | 
|
796  | 
||
797  | 
lemma combine_matrix_zero_r_neutral: "zero_r_neutral f \<Longrightarrow> zero_r_neutral (combine_matrix f)"  | 
|
798  | 
by (simp add: zero_r_neutral_def combine_matrix_def combine_infmatrix_def)  | 
|
799  | 
||
800  | 
lemma mult_matrix_zero_closed: "\<lbrakk>fadd 0 0 = 0; zero_closed fmul\<rbrakk> \<Longrightarrow> zero_closed (mult_matrix fmul fadd)"  | 
|
801  | 
apply (simp add: zero_closed_def mult_matrix_def mult_matrix_n_def)  | 
|
802  | 
apply (auto)  | 
|
803  | 
by (subst foldseq_zero, (simp add: zero_matrix_def)+)+  | 
|
804  | 
||
805  | 
lemma mult_matrix_n_zero_right[simp]: "\<lbrakk>fadd 0 0 = 0; !a. fmul a 0 = 0\<rbrakk> \<Longrightarrow> mult_matrix_n n fmul fadd A 0 = 0"  | 
|
806  | 
apply (simp add: mult_matrix_n_def)  | 
|
807  | 
apply (subst foldseq_zero)  | 
|
808  | 
by (simp_all add: zero_matrix_def)  | 
|
809  | 
||
810  | 
lemma mult_matrix_n_zero_left[simp]: "\<lbrakk>fadd 0 0 = 0; !a. fmul 0 a = 0\<rbrakk> \<Longrightarrow> mult_matrix_n n fmul fadd 0 A = 0"  | 
|
811  | 
apply (simp add: mult_matrix_n_def)  | 
|
812  | 
apply (subst foldseq_zero)  | 
|
813  | 
by (simp_all add: zero_matrix_def)  | 
|
814  | 
||
815  | 
lemma mult_matrix_zero_left[simp]: "\<lbrakk>fadd 0 0 = 0; !a. fmul 0 a = 0\<rbrakk> \<Longrightarrow> mult_matrix fmul fadd 0 A = 0"  | 
|
816  | 
by (simp add: mult_matrix_def)  | 
|
817  | 
||
818  | 
lemma mult_matrix_zero_right[simp]: "\<lbrakk>fadd 0 0 = 0; !a. fmul a 0 = 0\<rbrakk> \<Longrightarrow> mult_matrix fmul fadd A 0 = 0"  | 
|
819  | 
by (simp add: mult_matrix_def)  | 
|
820  | 
||
821  | 
lemma apply_matrix_zero[simp]: "f 0 = 0 \<Longrightarrow> apply_matrix f 0 = 0"  | 
|
822  | 
apply (simp add: apply_matrix_def apply_infmatrix_def)  | 
|
823  | 
by (simp add: zero_matrix_def)  | 
|
824  | 
||
825  | 
lemma combine_matrix_zero: "f 0 0 = 0 \<Longrightarrow> combine_matrix f 0 0 = 0"  | 
|
826  | 
apply (simp add: combine_matrix_def combine_infmatrix_def)  | 
|
827  | 
by (simp add: zero_matrix_def)  | 
|
828  | 
||
829  | 
lemma transpose_matrix_zero[simp]: "transpose_matrix 0 = 0"  | 
|
830  | 
apply (simp add: transpose_matrix_def transpose_infmatrix_def zero_matrix_def RepAbs_matrix)  | 
|
831  | 
apply (subst Rep_matrix_inject[symmetric], (rule ext)+)  | 
|
832  | 
apply (simp add: RepAbs_matrix)  | 
|
833  | 
done  | 
|
834  | 
||
835  | 
lemma apply_zero_matrix_def[simp]: "apply_matrix (% x. 0) A = 0"  | 
|
836  | 
apply (simp add: apply_matrix_def apply_infmatrix_def)  | 
|
837  | 
by (simp add: zero_matrix_def)  | 
|
838  | 
||
839  | 
constdefs  | 
|
840  | 
  singleton_matrix :: "nat \<Rightarrow> nat \<Rightarrow> ('a::zero) \<Rightarrow> 'a matrix"
 | 
|
841  | 
"singleton_matrix j i a == Abs_matrix(% m n. if j = m & i = n then a else 0)"  | 
|
842  | 
  move_matrix :: "('a::zero) matrix \<Rightarrow> int \<Rightarrow> int \<Rightarrow> 'a matrix"
 | 
|
843  | 
"move_matrix A y x == Abs_matrix(% j i. if (neg ((int j)-y)) | (neg ((int i)-x)) then 0 else Rep_matrix A (nat ((int j)-y)) (nat ((int i)-x)))"  | 
|
844  | 
  take_rows :: "('a::zero) matrix \<Rightarrow> nat \<Rightarrow> 'a matrix"
 | 
|
845  | 
"take_rows A r == Abs_matrix(% j i. if (j < r) then (Rep_matrix A j i) else 0)"  | 
|
846  | 
  take_columns :: "('a::zero) matrix \<Rightarrow> nat \<Rightarrow> 'a matrix"
 | 
|
847  | 
"take_columns A c == Abs_matrix(% j i. if (i < c) then (Rep_matrix A j i) else 0)"  | 
|
848  | 
||
849  | 
constdefs  | 
|
850  | 
  column_of_matrix :: "('a::zero) matrix \<Rightarrow> nat \<Rightarrow> 'a matrix"
 | 
|
851  | 
"column_of_matrix A n == take_columns (move_matrix A 0 (- int n)) 1"  | 
|
852  | 
  row_of_matrix :: "('a::zero) matrix \<Rightarrow> nat \<Rightarrow> 'a matrix"
 | 
|
853  | 
"row_of_matrix A m == take_rows (move_matrix A (- int m) 0) 1"  | 
|
854  | 
||
855  | 
lemma Rep_singleton_matrix[simp]: "Rep_matrix (singleton_matrix j i e) m n = (if j = m & i = n then e else 0)"  | 
|
856  | 
apply (simp add: singleton_matrix_def)  | 
|
857  | 
apply (auto)  | 
|
858  | 
apply (subst RepAbs_matrix)  | 
|
859  | 
apply (rule exI[of _ "Suc m"], simp)  | 
|
860  | 
apply (rule exI[of _ "Suc n"], simp+)  | 
|
861  | 
by (subst RepAbs_matrix, rule exI[of _ "Suc j"], simp, rule exI[of _ "Suc i"], simp+)+  | 
|
862  | 
||
863  | 
lemma apply_singleton_matrix[simp]: "f 0 = 0 \<Longrightarrow> apply_matrix f (singleton_matrix j i x) = (singleton_matrix j i (f x))"  | 
|
864  | 
apply (subst Rep_matrix_inject[symmetric])  | 
|
865  | 
apply (rule ext)+  | 
|
866  | 
apply (simp)  | 
|
867  | 
done  | 
|
868  | 
||
869  | 
lemma singleton_matrix_zero[simp]: "singleton_matrix j i 0 = 0"  | 
|
870  | 
by (simp add: singleton_matrix_def zero_matrix_def)  | 
|
871  | 
||
872  | 
lemma nrows_singleton[simp]: "nrows(singleton_matrix j i e) = (if e = 0 then 0 else Suc j)"  | 
|
873  | 
proof-  | 
|
874  | 
have th: "\<not> (\<forall>m. m \<le> j)" "\<exists>n. \<not> n \<le> i" by arith+  | 
|
875  | 
from th show ?thesis  | 
|
876  | 
apply (auto)  | 
|
877  | 
apply (rule le_anti_sym)  | 
|
878  | 
apply (subst nrows_le)  | 
|
879  | 
apply (simp add: singleton_matrix_def, auto)  | 
|
880  | 
apply (subst RepAbs_matrix)  | 
|
881  | 
apply auto  | 
|
882  | 
apply (simp add: Suc_le_eq)  | 
|
883  | 
apply (rule not_leE)  | 
|
884  | 
apply (subst nrows_le)  | 
|
885  | 
by simp  | 
|
886  | 
qed  | 
|
887  | 
||
888  | 
lemma ncols_singleton[simp]: "ncols(singleton_matrix j i e) = (if e = 0 then 0 else Suc i)"  | 
|
889  | 
proof-  | 
|
890  | 
have th: "\<not> (\<forall>m. m \<le> j)" "\<exists>n. \<not> n \<le> i" by arith+  | 
|
891  | 
from th show ?thesis  | 
|
892  | 
apply (auto)  | 
|
893  | 
apply (rule le_anti_sym)  | 
|
894  | 
apply (subst ncols_le)  | 
|
895  | 
apply (simp add: singleton_matrix_def, auto)  | 
|
896  | 
apply (subst RepAbs_matrix)  | 
|
897  | 
apply auto  | 
|
898  | 
apply (simp add: Suc_le_eq)  | 
|
899  | 
apply (rule not_leE)  | 
|
900  | 
apply (subst ncols_le)  | 
|
901  | 
by simp  | 
|
902  | 
qed  | 
|
903  | 
||
904  | 
lemma combine_singleton: "f 0 0 = 0 \<Longrightarrow> combine_matrix f (singleton_matrix j i a) (singleton_matrix j i b) = singleton_matrix j i (f a b)"  | 
|
905  | 
apply (simp add: singleton_matrix_def combine_matrix_def combine_infmatrix_def)  | 
|
906  | 
apply (subst RepAbs_matrix)  | 
|
907  | 
apply (rule exI[of _ "Suc j"], simp)  | 
|
908  | 
apply (rule exI[of _ "Suc i"], simp)  | 
|
909  | 
apply (rule comb[of "Abs_matrix" "Abs_matrix"], simp, (rule ext)+)  | 
|
910  | 
apply (subst RepAbs_matrix)  | 
|
911  | 
apply (rule exI[of _ "Suc j"], simp)  | 
|
912  | 
apply (rule exI[of _ "Suc i"], simp)  | 
|
913  | 
by simp  | 
|
914  | 
||
915  | 
lemma transpose_singleton[simp]: "transpose_matrix (singleton_matrix j i a) = singleton_matrix i j a"  | 
|
916  | 
apply (subst Rep_matrix_inject[symmetric], (rule ext)+)  | 
|
917  | 
apply (simp)  | 
|
918  | 
done  | 
|
919  | 
||
920  | 
lemma Rep_move_matrix[simp]:  | 
|
921  | 
"Rep_matrix (move_matrix A y x) j i =  | 
|
922  | 
(if (neg ((int j)-y)) | (neg ((int i)-x)) then 0 else Rep_matrix A (nat((int j)-y)) (nat((int i)-x)))"  | 
|
923  | 
apply (simp add: move_matrix_def)  | 
|
924  | 
apply (auto)  | 
|
925  | 
by (subst RepAbs_matrix,  | 
|
926  | 
rule exI[of _ "(nrows A)+(nat (abs y))"], auto, rule nrows, arith,  | 
|
927  | 
rule exI[of _ "(ncols A)+(nat (abs x))"], auto, rule ncols, arith)+  | 
|
928  | 
||
929  | 
lemma move_matrix_0_0[simp]: "move_matrix A 0 0 = A"  | 
|
930  | 
by (simp add: move_matrix_def)  | 
|
931  | 
||
932  | 
lemma move_matrix_ortho: "move_matrix A j i = move_matrix (move_matrix A j 0) 0 i"  | 
|
933  | 
apply (subst Rep_matrix_inject[symmetric])  | 
|
934  | 
apply (rule ext)+  | 
|
935  | 
apply (simp)  | 
|
936  | 
done  | 
|
937  | 
||
938  | 
lemma transpose_move_matrix[simp]:  | 
|
939  | 
"transpose_matrix (move_matrix A x y) = move_matrix (transpose_matrix A) y x"  | 
|
940  | 
apply (subst Rep_matrix_inject[symmetric], (rule ext)+)  | 
|
941  | 
apply (simp)  | 
|
942  | 
done  | 
|
943  | 
||
944  | 
lemma move_matrix_singleton[simp]: "move_matrix (singleton_matrix u v x) j i =  | 
|
945  | 
(if (j + int u < 0) | (i + int v < 0) then 0 else (singleton_matrix (nat (j + int u)) (nat (i + int v)) x))"  | 
|
946  | 
apply (subst Rep_matrix_inject[symmetric])  | 
|
947  | 
apply (rule ext)+  | 
|
948  | 
apply (case_tac "j + int u < 0")  | 
|
949  | 
apply (simp, arith)  | 
|
950  | 
apply (case_tac "i + int v < 0")  | 
|
951  | 
apply (simp add: neg_def, arith)  | 
|
952  | 
apply (simp add: neg_def)  | 
|
953  | 
apply arith  | 
|
954  | 
done  | 
|
955  | 
||
956  | 
lemma Rep_take_columns[simp]:  | 
|
957  | 
"Rep_matrix (take_columns A c) j i =  | 
|
958  | 
(if i < c then (Rep_matrix A j i) else 0)"  | 
|
959  | 
apply (simp add: take_columns_def)  | 
|
960  | 
apply (simplesubst RepAbs_matrix)  | 
|
961  | 
apply (rule exI[of _ "nrows A"], auto, simp add: nrows_le)  | 
|
962  | 
apply (rule exI[of _ "ncols A"], auto, simp add: ncols_le)  | 
|
963  | 
done  | 
|
964  | 
||
965  | 
lemma Rep_take_rows[simp]:  | 
|
966  | 
"Rep_matrix (take_rows A r) j i =  | 
|
967  | 
(if j < r then (Rep_matrix A j i) else 0)"  | 
|
968  | 
apply (simp add: take_rows_def)  | 
|
969  | 
apply (simplesubst RepAbs_matrix)  | 
|
970  | 
apply (rule exI[of _ "nrows A"], auto, simp add: nrows_le)  | 
|
971  | 
apply (rule exI[of _ "ncols A"], auto, simp add: ncols_le)  | 
|
972  | 
done  | 
|
973  | 
||
974  | 
lemma Rep_column_of_matrix[simp]:  | 
|
975  | 
"Rep_matrix (column_of_matrix A c) j i = (if i = 0 then (Rep_matrix A j c) else 0)"  | 
|
976  | 
by (simp add: column_of_matrix_def)  | 
|
977  | 
||
978  | 
lemma Rep_row_of_matrix[simp]:  | 
|
979  | 
"Rep_matrix (row_of_matrix A r) j i = (if j = 0 then (Rep_matrix A r i) else 0)"  | 
|
980  | 
by (simp add: row_of_matrix_def)  | 
|
981  | 
||
982  | 
lemma column_of_matrix: "ncols A <= n \<Longrightarrow> column_of_matrix A n = 0"  | 
|
983  | 
apply (subst Rep_matrix_inject[THEN sym])  | 
|
984  | 
apply (rule ext)+  | 
|
985  | 
by (simp add: ncols)  | 
|
986  | 
||
987  | 
lemma row_of_matrix: "nrows A <= n \<Longrightarrow> row_of_matrix A n = 0"  | 
|
988  | 
apply (subst Rep_matrix_inject[THEN sym])  | 
|
989  | 
apply (rule ext)+  | 
|
990  | 
by (simp add: nrows)  | 
|
991  | 
||
992  | 
lemma mult_matrix_singleton_right[simp]:  | 
|
993  | 
assumes prems:  | 
|
994  | 
"! x. fmul x 0 = 0"  | 
|
995  | 
"! x. fmul 0 x = 0"  | 
|
996  | 
"! x. fadd 0 x = x"  | 
|
997  | 
"! x. fadd x 0 = x"  | 
|
998  | 
shows "(mult_matrix fmul fadd A (singleton_matrix j i e)) = apply_matrix (% x. fmul x e) (move_matrix (column_of_matrix A j) 0 (int i))"  | 
|
999  | 
apply (simp add: mult_matrix_def)  | 
|
1000  | 
apply (subst mult_matrix_nm[of _ _ _ "max (ncols A) (Suc j)"])  | 
|
1001  | 
apply (auto)  | 
|
1002  | 
apply (simp add: prems)+  | 
|
1003  | 
apply (simp add: mult_matrix_n_def apply_matrix_def apply_infmatrix_def)  | 
|
1004  | 
apply (rule comb[of "Abs_matrix" "Abs_matrix"], auto, (rule ext)+)  | 
|
1005  | 
apply (subst foldseq_almostzero[of _ j])  | 
|
1006  | 
apply (simp add: prems)+  | 
|
1007  | 
apply (auto)  | 
|
| 29700 | 1008  | 
apply (metis comm_monoid_add.mult_1 le_anti_sym le_diff_eq not_neg_nat zero_le_imp_of_nat zle_int)  | 
1009  | 
done  | 
|
| 27484 | 1010  | 
|
1011  | 
lemma mult_matrix_ext:  | 
|
1012  | 
assumes  | 
|
1013  | 
eprem:  | 
|
1014  | 
"? e. (! a b. a \<noteq> b \<longrightarrow> fmul a e \<noteq> fmul b e)"  | 
|
1015  | 
and fprems:  | 
|
1016  | 
"! a. fmul 0 a = 0"  | 
|
1017  | 
"! a. fmul a 0 = 0"  | 
|
1018  | 
"! a. fadd a 0 = a"  | 
|
1019  | 
"! a. fadd 0 a = a"  | 
|
1020  | 
and contraprems:  | 
|
1021  | 
"mult_matrix fmul fadd A = mult_matrix fmul fadd B"  | 
|
1022  | 
shows  | 
|
1023  | 
"A = B"  | 
|
1024  | 
proof(rule contrapos_np[of "False"], simp)  | 
|
1025  | 
assume a: "A \<noteq> B"  | 
|
1026  | 
have b: "!! f g. (! x y. f x y = g x y) \<Longrightarrow> f = g" by ((rule ext)+, auto)  | 
|
1027  | 
have "? j i. (Rep_matrix A j i) \<noteq> (Rep_matrix B j i)"  | 
|
1028  | 
apply (rule contrapos_np[of "False"], simp+)  | 
|
1029  | 
apply (insert b[of "Rep_matrix A" "Rep_matrix B"], simp)  | 
|
1030  | 
by (simp add: Rep_matrix_inject a)  | 
|
1031  | 
then obtain J I where c:"(Rep_matrix A J I) \<noteq> (Rep_matrix B J I)" by blast  | 
|
1032  | 
from eprem obtain e where eprops:"(! a b. a \<noteq> b \<longrightarrow> fmul a e \<noteq> fmul b e)" by blast  | 
|
1033  | 
let ?S = "singleton_matrix I 0 e"  | 
|
1034  | 
let ?comp = "mult_matrix fmul fadd"  | 
|
1035  | 
have d: "!!x f g. f = g \<Longrightarrow> f x = g x" by blast  | 
|
1036  | 
have e: "(% x. fmul x e) 0 = 0" by (simp add: prems)  | 
|
1037  | 
have "~(?comp A ?S = ?comp B ?S)"  | 
|
1038  | 
apply (rule notI)  | 
|
1039  | 
apply (simp add: fprems eprops)  | 
|
1040  | 
apply (simp add: Rep_matrix_inject[THEN sym])  | 
|
1041  | 
apply (drule d[of _ _ "J"], drule d[of _ _ "0"])  | 
|
1042  | 
by (simp add: e c eprops)  | 
|
1043  | 
with contraprems show "False" by simp  | 
|
1044  | 
qed  | 
|
1045  | 
||
1046  | 
constdefs  | 
|
1047  | 
  foldmatrix :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a infmatrix) \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'a"
 | 
|
1048  | 
"foldmatrix f g A m n == foldseq_transposed g (% j. foldseq f (A j) n) m"  | 
|
1049  | 
  foldmatrix_transposed :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a infmatrix) \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'a"
 | 
|
1050  | 
"foldmatrix_transposed f g A m n == foldseq g (% j. foldseq_transposed f (A j) n) m"  | 
|
1051  | 
||
1052  | 
lemma foldmatrix_transpose:  | 
|
1053  | 
assumes  | 
|
1054  | 
"! a b c d. g(f a b) (f c d) = f (g a c) (g b d)"  | 
|
1055  | 
shows  | 
|
1056  | 
"foldmatrix f g A m n = foldmatrix_transposed g f (transpose_infmatrix A) n m" (is ?concl)  | 
|
1057  | 
proof -  | 
|
1058  | 
have forall:"!! P x. (! x. P x) \<Longrightarrow> P x" by auto  | 
|
1059  | 
have tworows:"! A. foldmatrix f g A 1 n = foldmatrix_transposed g f (transpose_infmatrix A) n 1"  | 
|
1060  | 
apply (induct n)  | 
|
1061  | 
apply (simp add: foldmatrix_def foldmatrix_transposed_def prems)+  | 
|
1062  | 
apply (auto)  | 
|
1063  | 
by (drule_tac x="(% j i. A j (Suc i))" in forall, simp)  | 
|
1064  | 
show "foldmatrix f g A m n = foldmatrix_transposed g f (transpose_infmatrix A) n m"  | 
|
1065  | 
apply (simp add: foldmatrix_def foldmatrix_transposed_def)  | 
|
1066  | 
apply (induct m, simp)  | 
|
1067  | 
apply (simp)  | 
|
1068  | 
apply (insert tworows)  | 
|
1069  | 
apply (drule_tac x="% j i. (if j = 0 then (foldseq_transposed g (\<lambda>u. A u i) m) else (A (Suc m) i))" in spec)  | 
|
1070  | 
by (simp add: foldmatrix_def foldmatrix_transposed_def)  | 
|
1071  | 
qed  | 
|
1072  | 
||
1073  | 
lemma foldseq_foldseq:  | 
|
1074  | 
assumes  | 
|
1075  | 
"associative f"  | 
|
1076  | 
"associative g"  | 
|
1077  | 
"! a b c d. g(f a b) (f c d) = f (g a c) (g b d)"  | 
|
1078  | 
shows  | 
|
1079  | 
"foldseq g (% j. foldseq f (A j) n) m = foldseq f (% j. foldseq g ((transpose_infmatrix A) j) m) n"  | 
|
1080  | 
apply (insert foldmatrix_transpose[of g f A m n])  | 
|
1081  | 
by (simp add: foldmatrix_def foldmatrix_transposed_def foldseq_assoc[THEN sym] prems)  | 
|
1082  | 
||
1083  | 
lemma mult_n_nrows:  | 
|
1084  | 
assumes  | 
|
1085  | 
"! a. fmul 0 a = 0"  | 
|
1086  | 
"! a. fmul a 0 = 0"  | 
|
1087  | 
"fadd 0 0 = 0"  | 
|
1088  | 
shows "nrows (mult_matrix_n n fmul fadd A B) \<le> nrows A"  | 
|
1089  | 
apply (subst nrows_le)  | 
|
1090  | 
apply (simp add: mult_matrix_n_def)  | 
|
1091  | 
apply (subst RepAbs_matrix)  | 
|
1092  | 
apply (rule_tac x="nrows A" in exI)  | 
|
1093  | 
apply (simp add: nrows prems foldseq_zero)  | 
|
1094  | 
apply (rule_tac x="ncols B" in exI)  | 
|
1095  | 
apply (simp add: ncols prems foldseq_zero)  | 
|
1096  | 
by (simp add: nrows prems foldseq_zero)  | 
|
1097  | 
||
1098  | 
lemma mult_n_ncols:  | 
|
1099  | 
assumes  | 
|
1100  | 
"! a. fmul 0 a = 0"  | 
|
1101  | 
"! a. fmul a 0 = 0"  | 
|
1102  | 
"fadd 0 0 = 0"  | 
|
1103  | 
shows "ncols (mult_matrix_n n fmul fadd A B) \<le> ncols B"  | 
|
1104  | 
apply (subst ncols_le)  | 
|
1105  | 
apply (simp add: mult_matrix_n_def)  | 
|
1106  | 
apply (subst RepAbs_matrix)  | 
|
1107  | 
apply (rule_tac x="nrows A" in exI)  | 
|
1108  | 
apply (simp add: nrows prems foldseq_zero)  | 
|
1109  | 
apply (rule_tac x="ncols B" in exI)  | 
|
1110  | 
apply (simp add: ncols prems foldseq_zero)  | 
|
1111  | 
by (simp add: ncols prems foldseq_zero)  | 
|
1112  | 
||
1113  | 
lemma mult_nrows:  | 
|
1114  | 
assumes  | 
|
1115  | 
"! a. fmul 0 a = 0"  | 
|
1116  | 
"! a. fmul a 0 = 0"  | 
|
1117  | 
"fadd 0 0 = 0"  | 
|
1118  | 
shows "nrows (mult_matrix fmul fadd A B) \<le> nrows A"  | 
|
1119  | 
by (simp add: mult_matrix_def mult_n_nrows prems)  | 
|
1120  | 
||
1121  | 
lemma mult_ncols:  | 
|
1122  | 
assumes  | 
|
1123  | 
"! a. fmul 0 a = 0"  | 
|
1124  | 
"! a. fmul a 0 = 0"  | 
|
1125  | 
"fadd 0 0 = 0"  | 
|
1126  | 
shows "ncols (mult_matrix fmul fadd A B) \<le> ncols B"  | 
|
1127  | 
by (simp add: mult_matrix_def mult_n_ncols prems)  | 
|
1128  | 
||
1129  | 
lemma nrows_move_matrix_le: "nrows (move_matrix A j i) <= nat((int (nrows A)) + j)"  | 
|
1130  | 
apply (auto simp add: nrows_le)  | 
|
1131  | 
apply (rule nrows)  | 
|
1132  | 
apply (arith)  | 
|
1133  | 
done  | 
|
1134  | 
||
1135  | 
lemma ncols_move_matrix_le: "ncols (move_matrix A j i) <= nat((int (ncols A)) + i)"  | 
|
1136  | 
apply (auto simp add: ncols_le)  | 
|
1137  | 
apply (rule ncols)  | 
|
1138  | 
apply (arith)  | 
|
1139  | 
done  | 
|
1140  | 
||
1141  | 
lemma mult_matrix_assoc:  | 
|
1142  | 
assumes prems:  | 
|
1143  | 
"! a. fmul1 0 a = 0"  | 
|
1144  | 
"! a. fmul1 a 0 = 0"  | 
|
1145  | 
"! a. fmul2 0 a = 0"  | 
|
1146  | 
"! a. fmul2 a 0 = 0"  | 
|
1147  | 
"fadd1 0 0 = 0"  | 
|
1148  | 
"fadd2 0 0 = 0"  | 
|
1149  | 
"! a b c d. fadd2 (fadd1 a b) (fadd1 c d) = fadd1 (fadd2 a c) (fadd2 b d)"  | 
|
1150  | 
"associative fadd1"  | 
|
1151  | 
"associative fadd2"  | 
|
1152  | 
"! a b c. fmul2 (fmul1 a b) c = fmul1 a (fmul2 b c)"  | 
|
1153  | 
"! a b c. fmul2 (fadd1 a b) c = fadd1 (fmul2 a c) (fmul2 b c)"  | 
|
1154  | 
"! a b c. fmul1 c (fadd2 a b) = fadd2 (fmul1 c a) (fmul1 c b)"  | 
|
1155  | 
shows "mult_matrix fmul2 fadd2 (mult_matrix fmul1 fadd1 A B) C = mult_matrix fmul1 fadd1 A (mult_matrix fmul2 fadd2 B C)" (is ?concl)  | 
|
1156  | 
proof -  | 
|
1157  | 
have comb_left: "!! A B x y. A = B \<Longrightarrow> (Rep_matrix (Abs_matrix A)) x y = (Rep_matrix(Abs_matrix B)) x y" by blast  | 
|
1158  | 
have fmul2fadd1fold: "!! x s n. fmul2 (foldseq fadd1 s n) x = foldseq fadd1 (% k. fmul2 (s k) x) n"  | 
|
1159  | 
by (rule_tac g1 = "% y. fmul2 y x" in ssubst [OF foldseq_distr_unary], simp_all!)  | 
|
1160  | 
have fmul1fadd2fold: "!! x s n. fmul1 x (foldseq fadd2 s n) = foldseq fadd2 (% k. fmul1 x (s k)) n"  | 
|
1161  | 
by (rule_tac g1 = "% y. fmul1 x y" in ssubst [OF foldseq_distr_unary], simp_all!)  | 
|
1162  | 
let ?N = "max (ncols A) (max (ncols B) (max (nrows B) (nrows C)))"  | 
|
1163  | 
show ?concl  | 
|
1164  | 
apply (simp add: Rep_matrix_inject[THEN sym])  | 
|
1165  | 
apply (rule ext)+  | 
|
1166  | 
apply (simp add: mult_matrix_def)  | 
|
1167  | 
apply (simplesubst mult_matrix_nm[of _ "max (ncols (mult_matrix_n (max (ncols A) (nrows B)) fmul1 fadd1 A B)) (nrows C)" _ "max (ncols B) (nrows C)"])  | 
|
1168  | 
apply (simp add: max1 max2 mult_n_ncols mult_n_nrows prems)+  | 
|
1169  | 
apply (simplesubst mult_matrix_nm[of _ "max (ncols A) (nrows (mult_matrix_n (max (ncols B) (nrows C)) fmul2 fadd2 B C))" _ "max (ncols A) (nrows B)"]) apply (simp add: max1 max2 mult_n_ncols mult_n_nrows prems)+  | 
|
1170  | 
apply (simplesubst mult_matrix_nm[of _ _ _ "?N"])  | 
|
1171  | 
apply (simp add: max1 max2 mult_n_ncols mult_n_nrows prems)+  | 
|
1172  | 
apply (simplesubst mult_matrix_nm[of _ _ _ "?N"])  | 
|
1173  | 
apply (simp add: max1 max2 mult_n_ncols mult_n_nrows prems)+  | 
|
1174  | 
apply (simplesubst mult_matrix_nm[of _ _ _ "?N"])  | 
|
1175  | 
apply (simp add: max1 max2 mult_n_ncols mult_n_nrows prems)+  | 
|
1176  | 
apply (simplesubst mult_matrix_nm[of _ _ _ "?N"])  | 
|
1177  | 
apply (simp add: max1 max2 mult_n_ncols mult_n_nrows prems)+  | 
|
1178  | 
apply (simp add: mult_matrix_n_def)  | 
|
1179  | 
apply (rule comb_left)  | 
|
1180  | 
apply ((rule ext)+, simp)  | 
|
1181  | 
apply (simplesubst RepAbs_matrix)  | 
|
1182  | 
apply (rule exI[of _ "nrows B"])  | 
|
1183  | 
apply (simp add: nrows prems foldseq_zero)  | 
|
1184  | 
apply (rule exI[of _ "ncols C"])  | 
|
1185  | 
apply (simp add: prems ncols foldseq_zero)  | 
|
1186  | 
apply (subst RepAbs_matrix)  | 
|
1187  | 
apply (rule exI[of _ "nrows A"])  | 
|
1188  | 
apply (simp add: nrows prems foldseq_zero)  | 
|
1189  | 
apply (rule exI[of _ "ncols B"])  | 
|
1190  | 
apply (simp add: prems ncols foldseq_zero)  | 
|
1191  | 
apply (simp add: fmul2fadd1fold fmul1fadd2fold prems)  | 
|
1192  | 
apply (subst foldseq_foldseq)  | 
|
1193  | 
apply (simp add: prems)+  | 
|
1194  | 
by (simp add: transpose_infmatrix)  | 
|
1195  | 
qed  | 
|
1196  | 
||
1197  | 
lemma  | 
|
1198  | 
assumes prems:  | 
|
1199  | 
"! a. fmul1 0 a = 0"  | 
|
1200  | 
"! a. fmul1 a 0 = 0"  | 
|
1201  | 
"! a. fmul2 0 a = 0"  | 
|
1202  | 
"! a. fmul2 a 0 = 0"  | 
|
1203  | 
"fadd1 0 0 = 0"  | 
|
1204  | 
"fadd2 0 0 = 0"  | 
|
1205  | 
"! a b c d. fadd2 (fadd1 a b) (fadd1 c d) = fadd1 (fadd2 a c) (fadd2 b d)"  | 
|
1206  | 
"associative fadd1"  | 
|
1207  | 
"associative fadd2"  | 
|
1208  | 
"! a b c. fmul2 (fmul1 a b) c = fmul1 a (fmul2 b c)"  | 
|
1209  | 
"! a b c. fmul2 (fadd1 a b) c = fadd1 (fmul2 a c) (fmul2 b c)"  | 
|
1210  | 
"! a b c. fmul1 c (fadd2 a b) = fadd2 (fmul1 c a) (fmul1 c b)"  | 
|
1211  | 
shows  | 
|
1212  | 
"(mult_matrix fmul1 fadd1 A) o (mult_matrix fmul2 fadd2 B) = mult_matrix fmul2 fadd2 (mult_matrix fmul1 fadd1 A B)"  | 
|
1213  | 
apply (rule ext)+  | 
|
1214  | 
apply (simp add: comp_def )  | 
|
1215  | 
by (simp add: mult_matrix_assoc prems)  | 
|
1216  | 
||
1217  | 
lemma mult_matrix_assoc_simple:  | 
|
1218  | 
assumes prems:  | 
|
1219  | 
"! a. fmul 0 a = 0"  | 
|
1220  | 
"! a. fmul a 0 = 0"  | 
|
1221  | 
"fadd 0 0 = 0"  | 
|
1222  | 
"associative fadd"  | 
|
1223  | 
"commutative fadd"  | 
|
1224  | 
"associative fmul"  | 
|
1225  | 
"distributive fmul fadd"  | 
|
1226  | 
shows "mult_matrix fmul fadd (mult_matrix fmul fadd A B) C = mult_matrix fmul fadd A (mult_matrix fmul fadd B C)" (is ?concl)  | 
|
1227  | 
proof -  | 
|
1228  | 
have "!! a b c d. fadd (fadd a b) (fadd c d) = fadd (fadd a c) (fadd b d)"  | 
|
1229  | 
by (simp! add: associative_def commutative_def)  | 
|
1230  | 
then show ?concl  | 
|
1231  | 
apply (subst mult_matrix_assoc)  | 
|
1232  | 
apply (simp_all!)  | 
|
1233  | 
by (simp add: associative_def distributive_def l_distributive_def r_distributive_def)+  | 
|
1234  | 
qed  | 
|
1235  | 
||
1236  | 
lemma transpose_apply_matrix: "f 0 = 0 \<Longrightarrow> transpose_matrix (apply_matrix f A) = apply_matrix f (transpose_matrix A)"  | 
|
1237  | 
apply (simp add: Rep_matrix_inject[THEN sym])  | 
|
1238  | 
apply (rule ext)+  | 
|
1239  | 
by simp  | 
|
1240  | 
||
1241  | 
lemma transpose_combine_matrix: "f 0 0 = 0 \<Longrightarrow> transpose_matrix (combine_matrix f A B) = combine_matrix f (transpose_matrix A) (transpose_matrix B)"  | 
|
1242  | 
apply (simp add: Rep_matrix_inject[THEN sym])  | 
|
1243  | 
apply (rule ext)+  | 
|
1244  | 
by simp  | 
|
1245  | 
||
1246  | 
lemma Rep_mult_matrix:  | 
|
1247  | 
assumes  | 
|
1248  | 
"! a. fmul 0 a = 0"  | 
|
1249  | 
"! a. fmul a 0 = 0"  | 
|
1250  | 
"fadd 0 0 = 0"  | 
|
1251  | 
shows  | 
|
1252  | 
"Rep_matrix(mult_matrix fmul fadd A B) j i =  | 
|
1253  | 
foldseq fadd (% k. fmul (Rep_matrix A j k) (Rep_matrix B k i)) (max (ncols A) (nrows B))"  | 
|
1254  | 
apply (simp add: mult_matrix_def mult_matrix_n_def)  | 
|
1255  | 
apply (subst RepAbs_matrix)  | 
|
1256  | 
apply (rule exI[of _ "nrows A"], simp! add: nrows foldseq_zero)  | 
|
1257  | 
apply (rule exI[of _ "ncols B"], simp! add: ncols foldseq_zero)  | 
|
1258  | 
by simp  | 
|
1259  | 
||
1260  | 
lemma transpose_mult_matrix:  | 
|
1261  | 
assumes  | 
|
1262  | 
"! a. fmul 0 a = 0"  | 
|
1263  | 
"! a. fmul a 0 = 0"  | 
|
1264  | 
"fadd 0 0 = 0"  | 
|
1265  | 
"! x y. fmul y x = fmul x y"  | 
|
1266  | 
shows  | 
|
1267  | 
"transpose_matrix (mult_matrix fmul fadd A B) = mult_matrix fmul fadd (transpose_matrix B) (transpose_matrix A)"  | 
|
1268  | 
apply (simp add: Rep_matrix_inject[THEN sym])  | 
|
1269  | 
apply (rule ext)+  | 
|
1270  | 
by (simp! add: Rep_mult_matrix max_ac)  | 
|
1271  | 
||
1272  | 
lemma column_transpose_matrix: "column_of_matrix (transpose_matrix A) n = transpose_matrix (row_of_matrix A n)"  | 
|
1273  | 
apply (simp add: Rep_matrix_inject[THEN sym])  | 
|
1274  | 
apply (rule ext)+  | 
|
1275  | 
by simp  | 
|
1276  | 
||
1277  | 
lemma take_columns_transpose_matrix: "take_columns (transpose_matrix A) n = transpose_matrix (take_rows A n)"  | 
|
1278  | 
apply (simp add: Rep_matrix_inject[THEN sym])  | 
|
1279  | 
apply (rule ext)+  | 
|
1280  | 
by simp  | 
|
1281  | 
||
| 27580 | 1282  | 
instantiation matrix :: ("{zero, ord}") ord
 | 
| 27484 | 1283  | 
begin  | 
1284  | 
||
1285  | 
definition  | 
|
1286  | 
le_matrix_def: "A \<le> B \<longleftrightarrow> (\<forall>j i. Rep_matrix A j i \<le> Rep_matrix B j i)"  | 
|
1287  | 
||
1288  | 
definition  | 
|
| 28637 | 1289  | 
less_def: "A < (B\<Colon>'a matrix) \<longleftrightarrow> A \<le> B \<and> \<not> B \<le> A"  | 
| 27484 | 1290  | 
|
1291  | 
instance ..  | 
|
1292  | 
||
1293  | 
end  | 
|
1294  | 
||
| 27580 | 1295  | 
instance matrix :: ("{zero, order}") order
 | 
| 27484 | 1296  | 
apply intro_classes  | 
1297  | 
apply (simp_all add: le_matrix_def less_def)  | 
|
1298  | 
apply (auto)  | 
|
1299  | 
apply (drule_tac x=j in spec, drule_tac x=j in spec)  | 
|
1300  | 
apply (drule_tac x=i in spec, drule_tac x=i in spec)  | 
|
1301  | 
apply (simp)  | 
|
1302  | 
apply (simp add: Rep_matrix_inject[THEN sym])  | 
|
1303  | 
apply (rule ext)+  | 
|
1304  | 
apply (drule_tac x=xa in spec, drule_tac x=xa in spec)  | 
|
1305  | 
apply (drule_tac x=xb in spec, drule_tac x=xb in spec)  | 
|
| 28637 | 1306  | 
apply simp  | 
1307  | 
done  | 
|
| 27484 | 1308  | 
|
1309  | 
lemma le_apply_matrix:  | 
|
1310  | 
assumes  | 
|
1311  | 
"f 0 = 0"  | 
|
1312  | 
"! x y. x <= y \<longrightarrow> f x <= f y"  | 
|
1313  | 
  "(a::('a::{ord, zero}) matrix) <= b"
 | 
|
1314  | 
shows  | 
|
1315  | 
"apply_matrix f a <= apply_matrix f b"  | 
|
1316  | 
by (simp! add: le_matrix_def)  | 
|
1317  | 
||
1318  | 
lemma le_combine_matrix:  | 
|
1319  | 
assumes  | 
|
1320  | 
"f 0 0 = 0"  | 
|
1321  | 
"! a b c d. a <= b & c <= d \<longrightarrow> f a c <= f b d"  | 
|
1322  | 
"A <= B"  | 
|
1323  | 
"C <= D"  | 
|
1324  | 
shows  | 
|
1325  | 
"combine_matrix f A C <= combine_matrix f B D"  | 
|
1326  | 
by (simp! add: le_matrix_def)  | 
|
1327  | 
||
1328  | 
lemma le_left_combine_matrix:  | 
|
1329  | 
assumes  | 
|
1330  | 
"f 0 0 = 0"  | 
|
1331  | 
"! a b c. a <= b \<longrightarrow> f c a <= f c b"  | 
|
1332  | 
"A <= B"  | 
|
1333  | 
shows  | 
|
1334  | 
"combine_matrix f C A <= combine_matrix f C B"  | 
|
1335  | 
by (simp! add: le_matrix_def)  | 
|
1336  | 
||
1337  | 
lemma le_right_combine_matrix:  | 
|
1338  | 
assumes  | 
|
1339  | 
"f 0 0 = 0"  | 
|
1340  | 
"! a b c. a <= b \<longrightarrow> f a c <= f b c"  | 
|
1341  | 
"A <= B"  | 
|
1342  | 
shows  | 
|
1343  | 
"combine_matrix f A C <= combine_matrix f B C"  | 
|
1344  | 
by (simp! add: le_matrix_def)  | 
|
1345  | 
||
1346  | 
lemma le_transpose_matrix: "(A <= B) = (transpose_matrix A <= transpose_matrix B)"  | 
|
1347  | 
by (simp add: le_matrix_def, auto)  | 
|
1348  | 
||
1349  | 
lemma le_foldseq:  | 
|
1350  | 
assumes  | 
|
1351  | 
"! a b c d . a <= b & c <= d \<longrightarrow> f a c <= f b d"  | 
|
1352  | 
"! i. i <= n \<longrightarrow> s i <= t i"  | 
|
1353  | 
shows  | 
|
1354  | 
"foldseq f s n <= foldseq f t n"  | 
|
1355  | 
proof -  | 
|
1356  | 
have "! s t. (! i. i<=n \<longrightarrow> s i <= t i) \<longrightarrow> foldseq f s n <= foldseq f t n" by (induct_tac n, simp_all!)  | 
|
1357  | 
then show "foldseq f s n <= foldseq f t n" by (simp!)  | 
|
1358  | 
qed  | 
|
1359  | 
||
1360  | 
lemma le_left_mult:  | 
|
1361  | 
assumes  | 
|
1362  | 
"! a b c d. a <= b & c <= d \<longrightarrow> fadd a c <= fadd b d"  | 
|
1363  | 
"! c a b. 0 <= c & a <= b \<longrightarrow> fmul c a <= fmul c b"  | 
|
1364  | 
"! a. fmul 0 a = 0"  | 
|
1365  | 
"! a. fmul a 0 = 0"  | 
|
1366  | 
"fadd 0 0 = 0"  | 
|
1367  | 
"0 <= C"  | 
|
1368  | 
"A <= B"  | 
|
1369  | 
shows  | 
|
1370  | 
"mult_matrix fmul fadd C A <= mult_matrix fmul fadd C B"  | 
|
1371  | 
apply (simp! add: le_matrix_def Rep_mult_matrix)  | 
|
1372  | 
apply (auto)  | 
|
1373  | 
apply (simplesubst foldseq_zerotail[of _ _ _ "max (ncols C) (max (nrows A) (nrows B))"], simp_all add: nrows ncols max1 max2)+  | 
|
1374  | 
apply (rule le_foldseq)  | 
|
1375  | 
by (auto)  | 
|
1376  | 
||
1377  | 
lemma le_right_mult:  | 
|
1378  | 
assumes  | 
|
1379  | 
"! a b c d. a <= b & c <= d \<longrightarrow> fadd a c <= fadd b d"  | 
|
1380  | 
"! c a b. 0 <= c & a <= b \<longrightarrow> fmul a c <= fmul b c"  | 
|
1381  | 
"! a. fmul 0 a = 0"  | 
|
1382  | 
"! a. fmul a 0 = 0"  | 
|
1383  | 
"fadd 0 0 = 0"  | 
|
1384  | 
"0 <= C"  | 
|
1385  | 
"A <= B"  | 
|
1386  | 
shows  | 
|
1387  | 
"mult_matrix fmul fadd A C <= mult_matrix fmul fadd B C"  | 
|
1388  | 
apply (simp! add: le_matrix_def Rep_mult_matrix)  | 
|
1389  | 
apply (auto)  | 
|
1390  | 
apply (simplesubst foldseq_zerotail[of _ _ _ "max (nrows C) (max (ncols A) (ncols B))"], simp_all add: nrows ncols max1 max2)+  | 
|
1391  | 
apply (rule le_foldseq)  | 
|
1392  | 
by (auto)  | 
|
1393  | 
||
1394  | 
lemma spec2: "! j i. P j i \<Longrightarrow> P j i" by blast  | 
|
1395  | 
lemma neg_imp: "(\<not> Q \<longrightarrow> \<not> P) \<Longrightarrow> P \<longrightarrow> Q" by blast  | 
|
1396  | 
||
1397  | 
lemma singleton_matrix_le[simp]: "(singleton_matrix j i a <= singleton_matrix j i b) = (a <= (b::_::order))"  | 
|
1398  | 
by (auto simp add: le_matrix_def)  | 
|
1399  | 
||
1400  | 
lemma singleton_le_zero[simp]: "(singleton_matrix j i x <= 0) = (x <= (0::'a::{order,zero}))"
 | 
|
1401  | 
apply (auto)  | 
|
1402  | 
apply (simp add: le_matrix_def)  | 
|
1403  | 
apply (drule_tac j=j and i=i in spec2)  | 
|
1404  | 
apply (simp)  | 
|
1405  | 
apply (simp add: le_matrix_def)  | 
|
1406  | 
done  | 
|
1407  | 
||
1408  | 
lemma singleton_ge_zero[simp]: "(0 <= singleton_matrix j i x) = ((0::'a::{order,zero}) <= x)"
 | 
|
1409  | 
apply (auto)  | 
|
1410  | 
apply (simp add: le_matrix_def)  | 
|
1411  | 
apply (drule_tac j=j and i=i in spec2)  | 
|
1412  | 
apply (simp)  | 
|
1413  | 
apply (simp add: le_matrix_def)  | 
|
1414  | 
done  | 
|
1415  | 
||
1416  | 
lemma move_matrix_le_zero[simp]: "0 <= j \<Longrightarrow> 0 <= i \<Longrightarrow> (move_matrix A j i <= 0) = (A <= (0::('a::{order,zero}) matrix))"
 | 
|
1417  | 
apply (auto simp add: le_matrix_def neg_def)  | 
|
1418  | 
apply (drule_tac j="ja+(nat j)" and i="ia+(nat i)" in spec2)  | 
|
1419  | 
apply (auto)  | 
|
1420  | 
done  | 
|
1421  | 
||
1422  | 
lemma move_matrix_zero_le[simp]: "0 <= j \<Longrightarrow> 0 <= i \<Longrightarrow> (0 <= move_matrix A j i) = ((0::('a::{order,zero}) matrix) <= A)"
 | 
|
1423  | 
apply (auto simp add: le_matrix_def neg_def)  | 
|
1424  | 
apply (drule_tac j="ja+(nat j)" and i="ia+(nat i)" in spec2)  | 
|
1425  | 
apply (auto)  | 
|
1426  | 
done  | 
|
1427  | 
||
1428  | 
lemma move_matrix_le_move_matrix_iff[simp]: "0 <= j \<Longrightarrow> 0 <= i \<Longrightarrow> (move_matrix A j i <= move_matrix B j i) = (A <= (B::('a::{order,zero}) matrix))"
 | 
|
1429  | 
apply (auto simp add: le_matrix_def neg_def)  | 
|
1430  | 
apply (drule_tac j="ja+(nat j)" and i="ia+(nat i)" in spec2)  | 
|
1431  | 
apply (auto)  | 
|
1432  | 
done  | 
|
1433  | 
||
| 27580 | 1434  | 
instantiation matrix :: ("{lattice, zero}") lattice
 | 
| 25764 | 1435  | 
begin  | 
1436  | 
||
| 28562 | 1437  | 
definition [code del]: "inf = combine_matrix inf"  | 
| 25764 | 1438  | 
|
| 28562 | 1439  | 
definition [code del]: "sup = combine_matrix sup"  | 
| 25764 | 1440  | 
|
1441  | 
instance  | 
|
| 
22452
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
1442  | 
by default (auto simp add: inf_le1 inf_le2 le_infI le_matrix_def inf_matrix_def sup_matrix_def)  | 
| 
 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 
haftmann 
parents: 
22422 
diff
changeset
 | 
1443  | 
|
| 25764 | 1444  | 
end  | 
1445  | 
||
1446  | 
instantiation matrix :: ("{plus, zero}") plus
 | 
|
1447  | 
begin  | 
|
1448  | 
||
1449  | 
definition  | 
|
| 28562 | 1450  | 
plus_matrix_def [code del]: "A + B = combine_matrix (op +) A B"  | 
| 25764 | 1451  | 
|
1452  | 
instance ..  | 
|
1453  | 
||
1454  | 
end  | 
|
1455  | 
||
1456  | 
instantiation matrix :: ("{uminus, zero}") uminus
 | 
|
1457  | 
begin  | 
|
1458  | 
||
1459  | 
definition  | 
|
| 28562 | 1460  | 
minus_matrix_def [code del]: "- A = apply_matrix uminus A"  | 
| 25764 | 1461  | 
|
1462  | 
instance ..  | 
|
1463  | 
||
1464  | 
end  | 
|
1465  | 
||
1466  | 
instantiation matrix :: ("{minus, zero}") minus
 | 
|
1467  | 
begin  | 
|
| 14593 | 1468  | 
|
| 25764 | 1469  | 
definition  | 
| 28562 | 1470  | 
diff_matrix_def [code del]: "A - B = combine_matrix (op -) A B"  | 
| 25764 | 1471  | 
|
1472  | 
instance ..  | 
|
1473  | 
||
1474  | 
end  | 
|
1475  | 
||
1476  | 
instantiation matrix :: ("{plus, times, zero}") times
 | 
|
1477  | 
begin  | 
|
1478  | 
||
1479  | 
definition  | 
|
| 28562 | 1480  | 
times_matrix_def [code del]: "A * B = mult_matrix (op *) (op +) A B"  | 
| 14940 | 1481  | 
|
| 25764 | 1482  | 
instance ..  | 
1483  | 
||
1484  | 
end  | 
|
1485  | 
||
| 27653 | 1486  | 
instantiation matrix :: ("{lattice, uminus, zero}") abs
 | 
| 25764 | 1487  | 
begin  | 
| 14940 | 1488  | 
|
| 25764 | 1489  | 
definition  | 
| 28562 | 1490  | 
abs_matrix_def [code del]: "abs (A \<Colon> 'a matrix) = sup A (- A)"  | 
| 25764 | 1491  | 
|
1492  | 
instance ..  | 
|
1493  | 
||
1494  | 
end  | 
|
| 23879 | 1495  | 
|
| 27653 | 1496  | 
instance matrix :: (monoid_add) monoid_add  | 
1497  | 
proof  | 
|
1498  | 
fix A B C :: "'a matrix"  | 
|
| 14940 | 1499  | 
show "A + B + C = A + (B + C)"  | 
1500  | 
apply (simp add: plus_matrix_def)  | 
|
1501  | 
apply (rule combine_matrix_assoc[simplified associative_def, THEN spec, THEN spec, THEN spec])  | 
|
1502  | 
apply (simp_all add: add_assoc)  | 
|
1503  | 
done  | 
|
| 27653 | 1504  | 
show "0 + A = A"  | 
1505  | 
apply (simp add: plus_matrix_def)  | 
|
1506  | 
apply (rule combine_matrix_zero_l_neutral[simplified zero_l_neutral_def, THEN spec])  | 
|
1507  | 
apply (simp)  | 
|
1508  | 
done  | 
|
1509  | 
show "A + 0 = A"  | 
|
1510  | 
apply (simp add: plus_matrix_def)  | 
|
1511  | 
apply (rule combine_matrix_zero_r_neutral [simplified zero_r_neutral_def, THEN spec])  | 
|
1512  | 
apply (simp)  | 
|
1513  | 
done  | 
|
1514  | 
qed  | 
|
1515  | 
||
1516  | 
instance matrix :: (comm_monoid_add) comm_monoid_add  | 
|
1517  | 
proof  | 
|
1518  | 
fix A B :: "'a matrix"  | 
|
| 14940 | 1519  | 
show "A + B = B + A"  | 
1520  | 
apply (simp add: plus_matrix_def)  | 
|
1521  | 
apply (rule combine_matrix_commute[simplified commutative_def, THEN spec, THEN spec])  | 
|
1522  | 
apply (simp_all add: add_commute)  | 
|
1523  | 
done  | 
|
1524  | 
show "0 + A = A"  | 
|
1525  | 
apply (simp add: plus_matrix_def)  | 
|
1526  | 
apply (rule combine_matrix_zero_l_neutral[simplified zero_l_neutral_def, THEN spec])  | 
|
1527  | 
apply (simp)  | 
|
1528  | 
done  | 
|
| 27653 | 1529  | 
qed  | 
1530  | 
||
1531  | 
instance matrix :: (group_add) group_add  | 
|
1532  | 
proof  | 
|
1533  | 
fix A B :: "'a matrix"  | 
|
1534  | 
show "- A + A = 0"  | 
|
1535  | 
by (simp add: plus_matrix_def minus_matrix_def Rep_matrix_inject[symmetric] ext)  | 
|
1536  | 
show "A - B = A + - B"  | 
|
1537  | 
by (simp add: plus_matrix_def diff_matrix_def minus_matrix_def Rep_matrix_inject [symmetric] ext diff_minus)  | 
|
1538  | 
qed  | 
|
1539  | 
||
1540  | 
instance matrix :: (ab_group_add) ab_group_add  | 
|
1541  | 
proof  | 
|
1542  | 
fix A B :: "'a matrix"  | 
|
| 14940 | 1543  | 
show "- A + A = 0"  | 
1544  | 
by (simp add: plus_matrix_def minus_matrix_def Rep_matrix_inject[symmetric] ext)  | 
|
1545  | 
show "A - B = A + - B"  | 
|
1546  | 
by (simp add: plus_matrix_def diff_matrix_def minus_matrix_def Rep_matrix_inject[symmetric] ext)  | 
|
| 27653 | 1547  | 
qed  | 
1548  | 
||
1549  | 
instance matrix :: (pordered_ab_group_add) pordered_ab_group_add  | 
|
1550  | 
proof  | 
|
1551  | 
fix A B C :: "'a matrix"  | 
|
| 14940 | 1552  | 
assume "A <= B"  | 
1553  | 
then show "C + A <= C + B"  | 
|
1554  | 
apply (simp add: plus_matrix_def)  | 
|
1555  | 
apply (rule le_left_combine_matrix)  | 
|
1556  | 
apply (simp_all)  | 
|
1557  | 
done  | 
|
1558  | 
qed  | 
|
| 27653 | 1559  | 
|
1560  | 
instance matrix :: (lordered_ab_group_add) lordered_ab_group_add_meet ..  | 
|
1561  | 
instance matrix :: (lordered_ab_group_add) lordered_ab_group_add_join ..  | 
|
| 14593 | 1562  | 
|
| 27653 | 1563  | 
instance matrix :: (ring) ring  | 
| 14940 | 1564  | 
proof  | 
| 27653 | 1565  | 
fix A B C :: "'a matrix"  | 
| 14940 | 1566  | 
show "A * B * C = A * (B * C)"  | 
1567  | 
apply (simp add: times_matrix_def)  | 
|
1568  | 
apply (rule mult_matrix_assoc)  | 
|
| 29667 | 1569  | 
apply (simp_all add: associative_def algebra_simps)  | 
| 14940 | 1570  | 
done  | 
1571  | 
show "(A + B) * C = A * C + B * C"  | 
|
1572  | 
apply (simp add: times_matrix_def plus_matrix_def)  | 
|
1573  | 
apply (rule l_distributive_matrix[simplified l_distributive_def, THEN spec, THEN spec, THEN spec])  | 
|
| 29667 | 1574  | 
apply (simp_all add: associative_def commutative_def algebra_simps)  | 
| 14940 | 1575  | 
done  | 
1576  | 
show "A * (B + C) = A * B + A * C"  | 
|
1577  | 
apply (simp add: times_matrix_def plus_matrix_def)  | 
|
1578  | 
apply (rule r_distributive_matrix[simplified r_distributive_def, THEN spec, THEN spec, THEN spec])  | 
|
| 29667 | 1579  | 
apply (simp_all add: associative_def commutative_def algebra_simps)  | 
| 27653 | 1580  | 
done  | 
1581  | 
qed  | 
|
1582  | 
||
1583  | 
instance matrix :: (pordered_ring) pordered_ring  | 
|
1584  | 
proof  | 
|
1585  | 
fix A B C :: "'a matrix"  | 
|
| 14940 | 1586  | 
assume a: "A \<le> B"  | 
1587  | 
assume b: "0 \<le> C"  | 
|
1588  | 
from a b show "C * A \<le> C * B"  | 
|
1589  | 
apply (simp add: times_matrix_def)  | 
|
1590  | 
apply (rule le_left_mult)  | 
|
1591  | 
apply (simp_all add: add_mono mult_left_mono)  | 
|
1592  | 
done  | 
|
1593  | 
from a b show "A * C \<le> B * C"  | 
|
1594  | 
apply (simp add: times_matrix_def)  | 
|
1595  | 
apply (rule le_right_mult)  | 
|
1596  | 
apply (simp_all add: add_mono mult_right_mono)  | 
|
1597  | 
done  | 
|
| 27653 | 1598  | 
qed  | 
1599  | 
||
1600  | 
instance matrix :: (lordered_ring) lordered_ring  | 
|
1601  | 
proof  | 
|
1602  | 
  fix A B C :: "('a :: lordered_ring) matrix"
 | 
|
1603  | 
show "abs A = sup A (-A)"  | 
|
1604  | 
by (simp add: abs_matrix_def)  | 
|
1605  | 
qed  | 
|
| 14593 | 1606  | 
|
| 
25303
 
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
 
haftmann 
parents: 
23879 
diff
changeset
 | 
1607  | 
lemma Rep_matrix_add[simp]:  | 
| 27653 | 1608  | 
  "Rep_matrix ((a::('a::monoid_add)matrix)+b) j i  = (Rep_matrix a j i) + (Rep_matrix b j i)"
 | 
1609  | 
by (simp add: plus_matrix_def)  | 
|
| 14593 | 1610  | 
|
| 27653 | 1611  | 
lemma Rep_matrix_mult: "Rep_matrix ((a::('a::ring) matrix) * b) j i = 
 | 
| 14940 | 1612  | 
foldseq (op +) (% k. (Rep_matrix a j k) * (Rep_matrix b k i)) (max (ncols a) (nrows b))"  | 
1613  | 
apply (simp add: times_matrix_def)  | 
|
1614  | 
apply (simp add: Rep_mult_matrix)  | 
|
1615  | 
done  | 
|
| 14593 | 1616  | 
|
| 27653 | 1617  | 
lemma apply_matrix_add: "! x y. f (x+y) = (f x) + (f y) \<Longrightarrow> f 0 = (0::'a)  | 
1618  | 
  \<Longrightarrow> apply_matrix f ((a::('a::monoid_add) matrix) + b) = (apply_matrix f a) + (apply_matrix f b)"
 | 
|
| 14940 | 1619  | 
apply (subst Rep_matrix_inject[symmetric])  | 
| 14593 | 1620  | 
apply (rule ext)+  | 
| 14940 | 1621  | 
apply (simp)  | 
1622  | 
done  | 
|
| 14593 | 1623  | 
|
| 27653 | 1624  | 
lemma singleton_matrix_add: "singleton_matrix j i ((a::_::monoid_add)+b) = (singleton_matrix j i a) + (singleton_matrix j i b)"  | 
| 14940 | 1625  | 
apply (subst Rep_matrix_inject[symmetric])  | 
1626  | 
apply (rule ext)+  | 
|
1627  | 
apply (simp)  | 
|
1628  | 
done  | 
|
| 14593 | 1629  | 
|
| 27653 | 1630  | 
lemma nrows_mult: "nrows ((A::('a::ring) matrix) * B) <= nrows A"
 | 
| 14593 | 1631  | 
by (simp add: times_matrix_def mult_nrows)  | 
1632  | 
||
| 27653 | 1633  | 
lemma ncols_mult: "ncols ((A::('a::ring) matrix) * B) <= ncols B"
 | 
| 14593 | 1634  | 
by (simp add: times_matrix_def mult_ncols)  | 
1635  | 
||
| 
22422
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
21312 
diff
changeset
 | 
1636  | 
definition  | 
| 
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
21312 
diff
changeset
 | 
1637  | 
  one_matrix :: "nat \<Rightarrow> ('a::{zero,one}) matrix" where
 | 
| 
 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 
haftmann 
parents: 
21312 
diff
changeset
 | 
1638  | 
"one_matrix n = Abs_matrix (% j i. if j = i & j < n then 1 else 0)"  | 
| 14593 | 1639  | 
|
1640  | 
lemma Rep_one_matrix[simp]: "Rep_matrix (one_matrix n) j i = (if (j = i & j < n) then 1 else 0)"  | 
|
1641  | 
apply (simp add: one_matrix_def)  | 
|
| 15481 | 1642  | 
apply (simplesubst RepAbs_matrix)  | 
| 14593 | 1643  | 
apply (rule exI[of _ n], simp add: split_if)+  | 
| 
16733
 
236dfafbeb63
linear arithmetic now takes "&" in assumptions apart.
 
nipkow 
parents: 
15481 
diff
changeset
 | 
1644  | 
by (simp add: split_if)  | 
| 14593 | 1645  | 
|
| 20633 | 1646  | 
lemma nrows_one_matrix[simp]: "nrows ((one_matrix n) :: ('a::zero_neq_one)matrix) = n" (is "?r = _")
 | 
| 14593 | 1647  | 
proof -  | 
1648  | 
have "?r <= n" by (simp add: nrows_le)  | 
|
| 14940 | 1649  | 
moreover have "n <= ?r" by (simp add:le_nrows, arith)  | 
| 14593 | 1650  | 
ultimately show "?r = n" by simp  | 
1651  | 
qed  | 
|
1652  | 
||
| 20633 | 1653  | 
lemma ncols_one_matrix[simp]: "ncols ((one_matrix n) :: ('a::zero_neq_one)matrix) = n" (is "?r = _")
 | 
| 14593 | 1654  | 
proof -  | 
1655  | 
have "?r <= n" by (simp add: ncols_le)  | 
|
1656  | 
moreover have "n <= ?r" by (simp add: le_ncols, arith)  | 
|
1657  | 
ultimately show "?r = n" by simp  | 
|
1658  | 
qed  | 
|
1659  | 
||
| 27653 | 1660  | 
lemma one_matrix_mult_right[simp]: "ncols A <= n \<Longrightarrow> (A::('a::{ring_1}) matrix) * (one_matrix n) = A"
 | 
| 14593 | 1661  | 
apply (subst Rep_matrix_inject[THEN sym])  | 
1662  | 
apply (rule ext)+  | 
|
1663  | 
apply (simp add: times_matrix_def Rep_mult_matrix)  | 
|
1664  | 
apply (rule_tac j1="xa" in ssubst[OF foldseq_almostzero])  | 
|
1665  | 
apply (simp_all)  | 
|
1666  | 
by (simp add: max_def ncols)  | 
|
1667  | 
||
| 27653 | 1668  | 
lemma one_matrix_mult_left[simp]: "nrows A <= n \<Longrightarrow> (one_matrix n) * A = (A::('a::ring_1) matrix)"
 | 
| 14593 | 1669  | 
apply (subst Rep_matrix_inject[THEN sym])  | 
1670  | 
apply (rule ext)+  | 
|
1671  | 
apply (simp add: times_matrix_def Rep_mult_matrix)  | 
|
1672  | 
apply (rule_tac j1="x" in ssubst[OF foldseq_almostzero])  | 
|
1673  | 
apply (simp_all)  | 
|
1674  | 
by (simp add: max_def nrows)  | 
|
1675  | 
||
| 27653 | 1676  | 
lemma transpose_matrix_mult: "transpose_matrix ((A::('a::comm_ring) matrix)*B) = (transpose_matrix B) * (transpose_matrix A)"
 | 
| 14940 | 1677  | 
apply (simp add: times_matrix_def)  | 
1678  | 
apply (subst transpose_mult_matrix)  | 
|
1679  | 
apply (simp_all add: mult_commute)  | 
|
1680  | 
done  | 
|
1681  | 
||
| 27653 | 1682  | 
lemma transpose_matrix_add: "transpose_matrix ((A::('a::monoid_add) matrix)+B) = transpose_matrix A + transpose_matrix B"
 | 
| 14940 | 1683  | 
by (simp add: plus_matrix_def transpose_combine_matrix)  | 
1684  | 
||
| 27653 | 1685  | 
lemma transpose_matrix_diff: "transpose_matrix ((A::('a::group_add) matrix)-B) = transpose_matrix A - transpose_matrix B"
 | 
| 14940 | 1686  | 
by (simp add: diff_matrix_def transpose_combine_matrix)  | 
1687  | 
||
| 27653 | 1688  | 
lemma transpose_matrix_minus: "transpose_matrix (-(A::('a::group_add) matrix)) = - transpose_matrix (A::'a matrix)"
 | 
| 14940 | 1689  | 
by (simp add: minus_matrix_def transpose_apply_matrix)  | 
1690  | 
||
1691  | 
constdefs  | 
|
| 27653 | 1692  | 
  right_inverse_matrix :: "('a::{ring_1}) matrix \<Rightarrow> 'a matrix \<Rightarrow> bool"
 | 
| 14940 | 1693  | 
"right_inverse_matrix A X == (A * X = one_matrix (max (nrows A) (ncols X))) \<and> nrows X \<le> ncols A"  | 
| 27653 | 1694  | 
  left_inverse_matrix :: "('a::{ring_1}) matrix \<Rightarrow> 'a matrix \<Rightarrow> bool"
 | 
| 14940 | 1695  | 
"left_inverse_matrix A X == (X * A = one_matrix (max(nrows X) (ncols A))) \<and> ncols X \<le> nrows A"  | 
| 27653 | 1696  | 
  inverse_matrix :: "('a::{ring_1}) matrix \<Rightarrow> 'a matrix \<Rightarrow> bool"
 | 
| 14940 | 1697  | 
"inverse_matrix A X == (right_inverse_matrix A X) \<and> (left_inverse_matrix A X)"  | 
| 14593 | 1698  | 
|
1699  | 
lemma right_inverse_matrix_dim: "right_inverse_matrix A X \<Longrightarrow> nrows A = ncols X"  | 
|
1700  | 
apply (insert ncols_mult[of A X], insert nrows_mult[of A X])  | 
|
1701  | 
by (simp add: right_inverse_matrix_def)  | 
|
1702  | 
||
| 14940 | 1703  | 
lemma left_inverse_matrix_dim: "left_inverse_matrix A Y \<Longrightarrow> ncols A = nrows Y"  | 
1704  | 
apply (insert ncols_mult[of Y A], insert nrows_mult[of Y A])  | 
|
1705  | 
by (simp add: left_inverse_matrix_def)  | 
|
1706  | 
||
1707  | 
lemma left_right_inverse_matrix_unique:  | 
|
1708  | 
assumes "left_inverse_matrix A Y" "right_inverse_matrix A X"  | 
|
1709  | 
shows "X = Y"  | 
|
1710  | 
proof -  | 
|
1711  | 
have "Y = Y * one_matrix (nrows A)"  | 
|
1712  | 
apply (subst one_matrix_mult_right)  | 
|
1713  | 
apply (insert prems)  | 
|
1714  | 
by (simp_all add: left_inverse_matrix_def)  | 
|
1715  | 
also have "\<dots> = Y * (A * X)"  | 
|
1716  | 
apply (insert prems)  | 
|
1717  | 
apply (frule right_inverse_matrix_dim)  | 
|
1718  | 
by (simp add: right_inverse_matrix_def)  | 
|
1719  | 
also have "\<dots> = (Y * A) * X" by (simp add: mult_assoc)  | 
|
1720  | 
also have "\<dots> = X"  | 
|
1721  | 
apply (insert prems)  | 
|
1722  | 
apply (frule left_inverse_matrix_dim)  | 
|
1723  | 
apply (simp_all add: left_inverse_matrix_def right_inverse_matrix_def one_matrix_mult_left)  | 
|
1724  | 
done  | 
|
1725  | 
ultimately show "X = Y" by (simp)  | 
|
1726  | 
qed  | 
|
1727  | 
||
1728  | 
lemma inverse_matrix_inject: "\<lbrakk> inverse_matrix A X; inverse_matrix A Y \<rbrakk> \<Longrightarrow> X = Y"  | 
|
1729  | 
by (auto simp add: inverse_matrix_def left_right_inverse_matrix_unique)  | 
|
1730  | 
||
1731  | 
lemma one_matrix_inverse: "inverse_matrix (one_matrix n) (one_matrix n)"  | 
|
1732  | 
by (simp add: inverse_matrix_def left_inverse_matrix_def right_inverse_matrix_def)  | 
|
1733  | 
||
1734  | 
lemma zero_imp_mult_zero: "(a::'a::ring) = 0 | b = 0 \<Longrightarrow> a * b = 0"  | 
|
1735  | 
by auto  | 
|
1736  | 
||
1737  | 
lemma Rep_matrix_zero_imp_mult_zero:  | 
|
1738  | 
  "! j i k. (Rep_matrix A j k = 0) | (Rep_matrix B k i) = 0  \<Longrightarrow> A * B = (0::('a::lordered_ring) matrix)"
 | 
|
1739  | 
apply (subst Rep_matrix_inject[symmetric])  | 
|
1740  | 
apply (rule ext)+  | 
|
1741  | 
apply (auto simp add: Rep_matrix_mult foldseq_zero zero_imp_mult_zero)  | 
|
1742  | 
done  | 
|
1743  | 
||
| 27653 | 1744  | 
lemma add_nrows: "nrows (A::('a::monoid_add) matrix) <= u \<Longrightarrow> nrows B <= u \<Longrightarrow> nrows (A + B) <= u"
 | 
| 14940 | 1745  | 
apply (simp add: plus_matrix_def)  | 
1746  | 
apply (rule combine_nrows)  | 
|
1747  | 
apply (simp_all)  | 
|
1748  | 
done  | 
|
1749  | 
||
| 27653 | 1750  | 
lemma move_matrix_row_mult: "move_matrix ((A::('a::ring) matrix) * B) j 0 = (move_matrix A j 0) * B"
 | 
| 14940 | 1751  | 
apply (subst Rep_matrix_inject[symmetric])  | 
1752  | 
apply (rule ext)+  | 
|
1753  | 
apply (auto simp add: Rep_matrix_mult foldseq_zero)  | 
|
1754  | 
apply (rule_tac foldseq_zerotail[symmetric])  | 
|
1755  | 
apply (auto simp add: nrows zero_imp_mult_zero max2)  | 
|
1756  | 
apply (rule order_trans)  | 
|
1757  | 
apply (rule ncols_move_matrix_le)  | 
|
1758  | 
apply (simp add: max1)  | 
|
1759  | 
done  | 
|
1760  | 
||
| 27653 | 1761  | 
lemma move_matrix_col_mult: "move_matrix ((A::('a::ring) matrix) * B) 0 i = A * (move_matrix B 0 i)"
 | 
| 14940 | 1762  | 
apply (subst Rep_matrix_inject[symmetric])  | 
1763  | 
apply (rule ext)+  | 
|
1764  | 
apply (auto simp add: Rep_matrix_mult foldseq_zero)  | 
|
1765  | 
apply (rule_tac foldseq_zerotail[symmetric])  | 
|
1766  | 
apply (auto simp add: ncols zero_imp_mult_zero max1)  | 
|
1767  | 
apply (rule order_trans)  | 
|
1768  | 
apply (rule nrows_move_matrix_le)  | 
|
1769  | 
apply (simp add: max2)  | 
|
1770  | 
done  | 
|
1771  | 
||
| 27653 | 1772  | 
lemma move_matrix_add: "((move_matrix (A + B) j i)::(('a::monoid_add) matrix)) = (move_matrix A j i) + (move_matrix B j i)" 
 | 
| 14940 | 1773  | 
apply (subst Rep_matrix_inject[symmetric])  | 
1774  | 
apply (rule ext)+  | 
|
1775  | 
apply (simp)  | 
|
1776  | 
done  | 
|
1777  | 
||
| 27653 | 1778  | 
lemma move_matrix_mult: "move_matrix ((A::('a::ring) matrix)*B) j i = (move_matrix A j 0) * (move_matrix B 0 i)"
 | 
| 14940 | 1779  | 
by (simp add: move_matrix_ortho[of "A*B"] move_matrix_col_mult move_matrix_row_mult)  | 
1780  | 
||
1781  | 
constdefs  | 
|
| 27653 | 1782  | 
  scalar_mult :: "('a::ring) \<Rightarrow> 'a matrix \<Rightarrow> 'a matrix"
 | 
| 14940 | 1783  | 
"scalar_mult a m == apply_matrix (op * a) m"  | 
1784  | 
||
1785  | 
lemma scalar_mult_zero[simp]: "scalar_mult y 0 = 0"  | 
|
| 
23477
 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 
nipkow 
parents: 
22452 
diff
changeset
 | 
1786  | 
by (simp add: scalar_mult_def)  | 
| 14940 | 1787  | 
|
1788  | 
lemma scalar_mult_add: "scalar_mult y (a+b) = (scalar_mult y a) + (scalar_mult y b)"  | 
|
| 29667 | 1789  | 
by (simp add: scalar_mult_def apply_matrix_add algebra_simps)  | 
| 14940 | 1790  | 
|
1791  | 
lemma Rep_scalar_mult[simp]: "Rep_matrix (scalar_mult y a) j i = y * (Rep_matrix a j i)"  | 
|
| 
23477
 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 
nipkow 
parents: 
22452 
diff
changeset
 | 
1792  | 
by (simp add: scalar_mult_def)  | 
| 14940 | 1793  | 
|
1794  | 
lemma scalar_mult_singleton[simp]: "scalar_mult y (singleton_matrix j i x) = singleton_matrix j i (y * x)"  | 
|
| 
23477
 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 
nipkow 
parents: 
22452 
diff
changeset
 | 
1795  | 
apply (subst Rep_matrix_inject[symmetric])  | 
| 
 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 
nipkow 
parents: 
22452 
diff
changeset
 | 
1796  | 
apply (rule ext)+  | 
| 
 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 
nipkow 
parents: 
22452 
diff
changeset
 | 
1797  | 
apply (auto)  | 
| 
 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 
nipkow 
parents: 
22452 
diff
changeset
 | 
1798  | 
done  | 
| 14940 | 1799  | 
|
| 27653 | 1800  | 
lemma Rep_minus[simp]: "Rep_matrix (-(A::_::group_add)) x y = - (Rep_matrix A x y)"  | 
| 
23477
 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 
nipkow 
parents: 
22452 
diff
changeset
 | 
1801  | 
by (simp add: minus_matrix_def)  | 
| 14940 | 1802  | 
|
| 27653 | 1803  | 
lemma Rep_abs[simp]: "Rep_matrix (abs (A::_::lordered_ab_group_add)) x y = abs (Rep_matrix A x y)"  | 
| 
23477
 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 
nipkow 
parents: 
22452 
diff
changeset
 | 
1804  | 
by (simp add: abs_lattice sup_matrix_def)  | 
| 14940 | 1805  | 
|
| 14593 | 1806  | 
end  |