| author | wenzelm | 
| Sat, 17 Sep 2005 19:17:34 +0200 | |
| changeset 17475 | d008d04068a1 | 
| parent 16933 | 91ded127f5f7 | 
| child 20217 | 25b068a99d2b | 
| permissions | -rw-r--r-- | 
| 14593 | 1 | (* Title: HOL/Matrix/MatrixGeneral.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Steven Obua | |
| 4 | *) | |
| 5 | ||
| 16417 | 6 | theory MatrixGeneral imports Main begin | 
| 14593 | 7 | |
| 8 | types 'a infmatrix = "[nat, nat] \<Rightarrow> 'a" | |
| 9 | ||
| 10 | constdefs | |
| 11 |   nonzero_positions :: "('a::zero) infmatrix \<Rightarrow> (nat*nat) set"
 | |
| 14662 | 12 |   "nonzero_positions A == {pos. A (fst pos) (snd pos) ~= 0}"
 | 
| 14593 | 13 | |
| 14 | typedef 'a matrix = "{(f::(('a::zero) infmatrix)). finite (nonzero_positions f)}"
 | |
| 15 | apply (rule_tac x="(% j i. 0)" in exI) | |
| 16 | by (simp add: nonzero_positions_def) | |
| 17 | ||
| 18 | declare Rep_matrix_inverse[simp] | |
| 19 | ||
| 20 | lemma finite_nonzero_positions : "finite (nonzero_positions (Rep_matrix A))" | |
| 14662 | 21 | apply (rule Abs_matrix_induct) | 
| 14593 | 22 | by (simp add: Abs_matrix_inverse matrix_def) | 
| 23 | ||
| 24 | constdefs | |
| 25 |   nrows :: "('a::zero) matrix \<Rightarrow> nat"
 | |
| 26 |   "nrows A == if nonzero_positions(Rep_matrix A) = {} then 0 else Suc(Max ((image fst) (nonzero_positions (Rep_matrix A))))"
 | |
| 27 |   ncols :: "('a::zero) matrix \<Rightarrow> nat"
 | |
| 14662 | 28 |   "ncols A == if nonzero_positions(Rep_matrix A) = {} then 0 else Suc(Max ((image snd) (nonzero_positions (Rep_matrix A))))"
 | 
| 14593 | 29 | |
| 14662 | 30 | lemma nrows: | 
| 31 | assumes hyp: "nrows A \<le> m" | |
| 14593 | 32 | shows "(Rep_matrix A m n) = 0" (is ?concl) | 
| 33 | proof cases | |
| 14662 | 34 |   assume "nonzero_positions(Rep_matrix A) = {}"
 | 
| 14593 | 35 | then show "(Rep_matrix A m n) = 0" by (simp add: nonzero_positions_def) | 
| 36 | next | |
| 37 |   assume a: "nonzero_positions(Rep_matrix A) \<noteq> {}"
 | |
| 14662 | 38 | let ?S = "fst`(nonzero_positions(Rep_matrix A))" | 
| 14593 | 39 |   from a have b: "?S \<noteq> {}" by (simp)
 | 
| 40 | have c: "finite (?S)" by (simp add: finite_nonzero_positions) | |
| 41 | from hyp have d: "Max (?S) < m" by (simp add: a nrows_def) | |
| 14662 | 42 | have "m \<notin> ?S" | 
| 43 | proof - | |
| 15580 | 44 | have "m \<in> ?S \<Longrightarrow> m <= Max(?S)" by (simp add: Max_ge[OF c b]) | 
| 14593 | 45 | moreover from d have "~(m <= Max ?S)" by (simp) | 
| 46 | ultimately show "m \<notin> ?S" by (auto) | |
| 47 | qed | |
| 48 | thus "Rep_matrix A m n = 0" by (simp add: nonzero_positions_def image_Collect) | |
| 49 | qed | |
| 14662 | 50 | |
| 14593 | 51 | constdefs | 
| 52 | transpose_infmatrix :: "'a infmatrix \<Rightarrow> 'a infmatrix" | |
| 53 | "transpose_infmatrix A j i == A i j" | |
| 54 |   transpose_matrix :: "('a::zero) matrix \<Rightarrow> 'a matrix"
 | |
| 55 | "transpose_matrix == Abs_matrix o transpose_infmatrix o Rep_matrix" | |
| 56 | ||
| 57 | declare transpose_infmatrix_def[simp] | |
| 58 | ||
| 59 | lemma transpose_infmatrix_twice[simp]: "transpose_infmatrix (transpose_infmatrix A) = A" | |
| 60 | by ((rule ext)+, simp) | |
| 61 | ||
| 62 | lemma transpose_infmatrix: "transpose_infmatrix (% j i. P j i) = (% j i. P i j)" | |
| 63 | apply (rule ext)+ | |
| 64 | by (simp add: transpose_infmatrix_def) | |
| 65 | ||
| 14662 | 66 | lemma transpose_infmatrix_closed[simp]: "Rep_matrix (Abs_matrix (transpose_infmatrix (Rep_matrix x))) = transpose_infmatrix (Rep_matrix x)" | 
| 14593 | 67 | apply (rule Abs_matrix_inverse) | 
| 68 | apply (simp add: matrix_def nonzero_positions_def image_def) | |
| 69 | proof - | |
| 70 |   let ?A = "{pos. Rep_matrix x (snd pos) (fst pos) \<noteq> 0}"
 | |
| 71 | let ?swap = "% pos. (snd pos, fst pos)" | |
| 72 |   let ?B = "{pos. Rep_matrix x (fst pos) (snd pos) \<noteq> 0}"
 | |
| 14662 | 73 | have swap_image: "?swap`?A = ?B" | 
| 14593 | 74 | apply (simp add: image_def) | 
| 75 | apply (rule set_ext) | |
| 76 | apply (simp) | |
| 77 | proof | |
| 78 | fix y | |
| 14662 | 79 | assume hyp: "\<exists>a b. Rep_matrix x b a \<noteq> 0 \<and> y = (b, a)" | 
| 80 | thus "Rep_matrix x (fst y) (snd y) \<noteq> 0" | |
| 81 | proof - | |
| 82 | from hyp obtain a b where "(Rep_matrix x b a \<noteq> 0 & y = (b,a))" by blast | |
| 83 | then show "Rep_matrix x (fst y) (snd y) \<noteq> 0" by (simp) | |
| 84 | qed | |
| 14593 | 85 | next | 
| 86 | fix y | |
| 87 | assume hyp: "Rep_matrix x (fst y) (snd y) \<noteq> 0" | |
| 14662 | 88 | show "\<exists> a b. (Rep_matrix x b a \<noteq> 0 & y = (b,a))" by (rule exI[of _ "snd y"], rule exI[of _ "fst y"], simp) | 
| 14593 | 89 | qed | 
| 14662 | 90 | then have "finite (?swap`?A)" | 
| 14593 | 91 | proof - | 
| 92 | have "finite (nonzero_positions (Rep_matrix x))" by (simp add: finite_nonzero_positions) | |
| 93 | then have "finite ?B" by (simp add: nonzero_positions_def) | |
| 94 | with swap_image show "finite (?swap`?A)" by (simp) | |
| 14662 | 95 | qed | 
| 14593 | 96 | moreover | 
| 97 | have "inj_on ?swap ?A" by (simp add: inj_on_def) | |
| 14662 | 98 | ultimately show "finite ?A"by (rule finite_imageD[of ?swap ?A]) | 
| 99 | qed | |
| 14593 | 100 | |
| 14940 | 101 | lemma infmatrixforward: "(x::'a infmatrix) = y \<Longrightarrow> \<forall> a b. x a b = y a b" by auto | 
| 102 | ||
| 103 | lemma transpose_infmatrix_inject: "(transpose_infmatrix A = transpose_infmatrix B) = (A = B)" | |
| 104 | apply (auto) | |
| 105 | apply (rule ext)+ | |
| 106 | apply (simp add: transpose_infmatrix) | |
| 107 | apply (drule infmatrixforward) | |
| 108 | apply (simp) | |
| 109 | done | |
| 110 | ||
| 111 | lemma transpose_matrix_inject: "(transpose_matrix A = transpose_matrix B) = (A = B)" | |
| 112 | apply (simp add: transpose_matrix_def) | |
| 113 | apply (subst Rep_matrix_inject[THEN sym])+ | |
| 114 | apply (simp only: transpose_infmatrix_closed transpose_infmatrix_inject) | |
| 115 | done | |
| 116 | ||
| 14593 | 117 | lemma transpose_matrix[simp]: "Rep_matrix(transpose_matrix A) j i = Rep_matrix A i j" | 
| 118 | by (simp add: transpose_matrix_def) | |
| 119 | ||
| 120 | lemma transpose_transpose_id[simp]: "transpose_matrix (transpose_matrix A) = A" | |
| 121 | by (simp add: transpose_matrix_def) | |
| 122 | ||
| 14662 | 123 | lemma nrows_transpose[simp]: "nrows (transpose_matrix A) = ncols A" | 
| 14593 | 124 | by (simp add: nrows_def ncols_def nonzero_positions_def transpose_matrix_def image_def) | 
| 125 | ||
| 14662 | 126 | lemma ncols_transpose[simp]: "ncols (transpose_matrix A) = nrows A" | 
| 14593 | 127 | by (simp add: nrows_def ncols_def nonzero_positions_def transpose_matrix_def image_def) | 
| 128 | ||
| 14662 | 129 | lemma ncols: "ncols A <= n \<Longrightarrow> Rep_matrix A m n = 0" | 
| 14593 | 130 | proof - | 
| 131 | assume "ncols A <= n" | |
| 14662 | 132 | then have "nrows (transpose_matrix A) <= n" by (simp) | 
| 14593 | 133 | then have "Rep_matrix (transpose_matrix A) n m = 0" by (rule nrows) | 
| 134 | thus "Rep_matrix A m n = 0" by (simp add: transpose_matrix_def) | |
| 135 | qed | |
| 136 | ||
| 137 | lemma ncols_le: "(ncols A <= n) = (! j i. n <= i \<longrightarrow> (Rep_matrix A j i) = 0)" (is "_ = ?st") | |
| 138 | apply (auto) | |
| 139 | apply (simp add: ncols) | |
| 140 | proof (simp add: ncols_def, auto) | |
| 141 | let ?P = "nonzero_positions (Rep_matrix A)" | |
| 142 | let ?p = "snd`?P" | |
| 143 | have a:"finite ?p" by (simp add: finite_nonzero_positions) | |
| 14662 | 144 | let ?m = "Max ?p" | 
| 14593 | 145 | assume "~(Suc (?m) <= n)" | 
| 146 | then have b:"n <= ?m" by (simp) | |
| 147 | fix a b | |
| 148 | assume "(a,b) \<in> ?P" | |
| 149 |   then have "?p \<noteq> {}" by (auto)
 | |
| 14662 | 150 | with a have "?m \<in> ?p" by (simp) | 
| 14593 | 151 | moreover have "!x. (x \<in> ?p \<longrightarrow> (? y. (Rep_matrix A y x) \<noteq> 0))" by (simp add: nonzero_positions_def image_def) | 
| 152 | ultimately have "? y. (Rep_matrix A y ?m) \<noteq> 0" by (simp) | |
| 153 | moreover assume ?st | |
| 154 | ultimately show "False" using b by (simp) | |
| 155 | qed | |
| 156 | ||
| 157 | lemma less_ncols: "(n < ncols A) = (? j i. n <= i & (Rep_matrix A j i) \<noteq> 0)" (is ?concl) | |
| 158 | proof - | |
| 159 | have a: "!! (a::nat) b. (a < b) = (~(b <= a))" by arith | |
| 160 | show ?concl by (simp add: a ncols_le) | |
| 161 | qed | |
| 162 | ||
| 163 | 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) | |
| 164 | apply (auto) | |
| 165 | apply (subgoal_tac "ncols A <= m") | |
| 166 | apply (simp) | |
| 167 | apply (simp add: ncols_le) | |
| 168 | apply (drule_tac x="ncols A" in spec) | |
| 169 | by (simp add: ncols) | |
| 170 | ||
| 171 | lemma nrows_le: "(nrows A <= n) = (! j i. n <= j \<longrightarrow> (Rep_matrix A j i) = 0)" (is ?s) | |
| 172 | proof - | |
| 173 | have "(nrows A <= n) = (ncols (transpose_matrix A) <= n)" by (simp) | |
| 174 | also have "\<dots> = (! j i. n <= i \<longrightarrow> (Rep_matrix (transpose_matrix A) j i = 0))" by (rule ncols_le) | |
| 175 | also have "\<dots> = (! j i. n <= i \<longrightarrow> (Rep_matrix A i j) = 0)" by (simp) | |
| 176 | finally show "(nrows A <= n) = (! j i. n <= j \<longrightarrow> (Rep_matrix A j i) = 0)" by (auto) | |
| 177 | qed | |
| 178 | ||
| 179 | lemma less_nrows: "(m < nrows A) = (? j i. m <= j & (Rep_matrix A j i) \<noteq> 0)" (is ?concl) | |
| 180 | proof - | |
| 181 | have a: "!! (a::nat) b. (a < b) = (~(b <= a))" by arith | |
| 182 | show ?concl by (simp add: a nrows_le) | |
| 183 | qed | |
| 184 | ||
| 185 | 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) | |
| 186 | apply (auto) | |
| 187 | apply (subgoal_tac "nrows A <= m") | |
| 188 | apply (simp) | |
| 189 | apply (simp add: nrows_le) | |
| 190 | apply (drule_tac x="nrows A" in spec) | |
| 191 | by (simp add: nrows) | |
| 192 | ||
| 14940 | 193 | lemma nrows_notzero: "Rep_matrix A m n \<noteq> 0 \<Longrightarrow> m < nrows A" | 
| 194 | apply (case_tac "nrows A <= m") | |
| 195 | apply (simp_all add: nrows) | |
| 196 | done | |
| 197 | ||
| 198 | lemma ncols_notzero: "Rep_matrix A m n \<noteq> 0 \<Longrightarrow> n < ncols A" | |
| 199 | apply (case_tac "ncols A <= n") | |
| 200 | apply (simp_all add: ncols) | |
| 201 | done | |
| 202 | ||
| 14593 | 203 | lemma finite_natarray1: "finite {x. x < (n::nat)}"
 | 
| 204 | apply (induct n) | |
| 205 | apply (simp) | |
| 206 | proof - | |
| 207 | fix n | |
| 208 |   have "{x. x < Suc n} = insert n {x. x < n}"  by (rule set_ext, simp, arith)
 | |
| 209 |   moreover assume "finite {x. x < n}"
 | |
| 210 |   ultimately show "finite {x. x < Suc n}" by (simp)
 | |
| 211 | qed | |
| 212 | ||
| 213 | lemma finite_natarray2: "finite {pos. (fst pos) < (m::nat) & (snd pos) < (n::nat)}"
 | |
| 214 | apply (induct m) | |
| 215 | apply (simp+) | |
| 216 | proof - | |
| 217 | fix m::nat | |
| 218 |     let ?s0 = "{pos. fst pos < m & snd pos < n}"
 | |
| 219 |     let ?s1 = "{pos. fst pos < (Suc m) & snd pos < n}"
 | |
| 220 |     let ?sd = "{pos. fst pos = m & snd pos < n}"
 | |
| 221 | assume f0: "finite ?s0" | |
| 14662 | 222 | have f1: "finite ?sd" | 
| 14593 | 223 | proof - | 
| 14662 | 224 | let ?f = "% x. (m, x)" | 
| 14593 | 225 |       have "{pos. fst pos = m & snd pos < n} = ?f ` {x. x < n}" by (rule set_ext, simp add: image_def, auto)
 | 
| 226 |       moreover have "finite {x. x < n}" by (simp add: finite_natarray1)
 | |
| 227 |       ultimately show "finite {pos. fst pos = m & snd pos < n}" by (simp)
 | |
| 228 | qed | |
| 229 | have su: "?s0 \<union> ?sd = ?s1" by (rule set_ext, simp, arith) | |
| 230 | from f0 f1 have "finite (?s0 \<union> ?sd)" by (rule finite_UnI) | |
| 231 | with su show "finite ?s1" by (simp) | |
| 232 | qed | |
| 14662 | 233 | |
| 234 | lemma RepAbs_matrix: | |
| 14593 | 235 | 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) | 
| 14662 | 236 | shows "(Rep_matrix (Abs_matrix x)) = x" | 
| 14593 | 237 | apply (rule Abs_matrix_inverse) | 
| 238 | apply (simp add: matrix_def nonzero_positions_def) | |
| 14662 | 239 | proof - | 
| 14593 | 240 | from aem obtain m where a: "! j i. m <= j \<longrightarrow> x j i = 0" by (blast) | 
| 14662 | 241 | from aen obtain n where b: "! j i. n <= i \<longrightarrow> x j i = 0" by (blast) | 
| 14593 | 242 |   let ?u = "{pos. x (fst pos) (snd pos) \<noteq> 0}"
 | 
| 243 |   let ?v = "{pos. fst pos < m & snd pos < n}"
 | |
| 14662 | 244 | have c: "!! (m::nat) a. ~(m <= a) \<Longrightarrow> a < m" by (arith) | 
| 14593 | 245 |   from a b have "(?u \<inter> (-?v)) = {}"
 | 
| 246 | apply (simp) | |
| 247 | apply (rule set_ext) | |
| 248 | apply (simp) | |
| 249 | apply auto | |
| 250 | by (rule c, auto)+ | |
| 251 | then have d: "?u \<subseteq> ?v" by blast | |
| 252 | moreover have "finite ?v" by (simp add: finite_natarray2) | |
| 253 | ultimately show "finite ?u" by (rule finite_subset) | |
| 254 | qed | |
| 255 | ||
| 14662 | 256 | constdefs | 
| 14593 | 257 |   apply_infmatrix :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a infmatrix \<Rightarrow> 'b infmatrix"
 | 
| 258 | "apply_infmatrix f == % A. (% j i. f (A j i))" | |
| 259 |   apply_matrix :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a::zero) matrix \<Rightarrow> ('b::zero) matrix"
 | |
| 14662 | 260 | "apply_matrix f == % A. Abs_matrix (apply_infmatrix f (Rep_matrix A))" | 
| 14593 | 261 |   combine_infmatrix :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a infmatrix \<Rightarrow> 'b infmatrix \<Rightarrow> 'c infmatrix"
 | 
| 14662 | 262 | "combine_infmatrix f == % A B. (% j i. f (A j i) (B j i))" | 
| 14593 | 263 |   combine_matrix :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> ('a::zero) matrix \<Rightarrow> ('b::zero) matrix \<Rightarrow> ('c::zero) matrix"
 | 
| 264 | "combine_matrix f == % A B. Abs_matrix (combine_infmatrix f (Rep_matrix A) (Rep_matrix B))" | |
| 265 | ||
| 266 | lemma expand_apply_infmatrix[simp]: "apply_infmatrix f A j i = f (A j i)" | |
| 267 | by (simp add: apply_infmatrix_def) | |
| 268 | ||
| 14662 | 269 | lemma expand_combine_infmatrix[simp]: "combine_infmatrix f A B j i = f (A j i) (B j i)" | 
| 270 | by (simp add: combine_infmatrix_def) | |
| 14593 | 271 | |
| 272 | constdefs | |
| 273 | commutative :: "('a \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> bool"
 | |
| 274 | "commutative f == ! x y. f x y = f y x" | |
| 275 | associative :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> bool"
 | |
| 276 | "associative f == ! x y z. f (f x y) z = f x (f y z)" | |
| 277 | ||
| 14662 | 278 | text{*
 | 
| 14593 | 279 | To reason about associativity and commutativity of operations on matrices, | 
| 280 | let's take a step back and look at the general situtation: Assume that we have | |
| 281 | 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. | |
| 282 | Each function $f: A \times A \rightarrow A$ now induces a function $f': B \times B \rightarrow B$ by $f' = u \circ f$. | |
| 14662 | 283 | 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.$ | 
| 14593 | 284 | *} | 
| 285 | ||
| 14662 | 286 | lemma combine_infmatrix_commute: | 
| 14593 | 287 | "commutative f \<Longrightarrow> commutative (combine_infmatrix f)" | 
| 288 | by (simp add: commutative_def combine_infmatrix_def) | |
| 289 | ||
| 290 | lemma combine_matrix_commute: | |
| 291 | "commutative f \<Longrightarrow> commutative (combine_matrix f)" | |
| 292 | by (simp add: combine_matrix_def commutative_def combine_infmatrix_def) | |
| 293 | ||
| 294 | text{*
 | |
| 14662 | 295 | 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\}$,
 | 
| 14593 | 296 | 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 | 
| 297 | \[ f' (f' 1 1) -1 = u(f (u (f 1 1)) -1) = u(f (u 2) -1) = u (f 0 -1) = -1, \] | |
| 298 | but on the other hand we have | |
| 299 | \[ f' 1 (f' 1 -1) = u (f 1 (u (f 1 -1))) = u (f 1 0) = 1.\] | |
| 300 | 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: | |
| 301 | *} | |
| 302 | ||
| 303 | 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)" | |
| 304 | by (rule subsetI, simp add: nonzero_positions_def combine_infmatrix_def, auto) | |
| 305 | ||
| 306 | lemma finite_nonzero_positions_Rep[simp]: "finite (nonzero_positions (Rep_matrix A))" | |
| 307 | by (insert Rep_matrix [of A], simp add: matrix_def) | |
| 308 | ||
| 14662 | 309 | lemma combine_infmatrix_closed [simp]: | 
| 14593 | 310 | "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)" | 
| 311 | apply (rule Abs_matrix_inverse) | |
| 312 | apply (simp add: matrix_def) | |
| 313 | apply (rule finite_subset[of _ "(nonzero_positions (Rep_matrix A)) \<union> (nonzero_positions (Rep_matrix B))"]) | |
| 314 | by (simp_all) | |
| 315 | ||
| 316 | text {* We need the next two lemmas only later, but it is analog to the above one, so we prove them now: *}
 | |
| 317 | lemma nonzero_positions_apply_infmatrix[simp]: "f 0 = 0 \<Longrightarrow> nonzero_positions (apply_infmatrix f A) \<subseteq> nonzero_positions A" | |
| 318 | by (rule subsetI, simp add: nonzero_positions_def apply_infmatrix_def, auto) | |
| 319 | ||
| 320 | lemma apply_infmatrix_closed [simp]: | |
| 321 | "f 0 = 0 \<Longrightarrow> Rep_matrix (Abs_matrix (apply_infmatrix f (Rep_matrix A))) = apply_infmatrix f (Rep_matrix A)" | |
| 322 | apply (rule Abs_matrix_inverse) | |
| 323 | apply (simp add: matrix_def) | |
| 324 | apply (rule finite_subset[of _ "nonzero_positions (Rep_matrix A)"]) | |
| 325 | by (simp_all) | |
| 326 | ||
| 327 | lemma combine_infmatrix_assoc[simp]: "f 0 0 = 0 \<Longrightarrow> associative f \<Longrightarrow> associative (combine_infmatrix f)" | |
| 328 | by (simp add: associative_def combine_infmatrix_def) | |
| 329 | ||
| 330 | lemma comb: "f = g \<Longrightarrow> x = y \<Longrightarrow> f x = g y" | |
| 331 | by (auto) | |
| 332 | ||
| 333 | lemma combine_matrix_assoc: "f 0 0 = 0 \<Longrightarrow> associative f \<Longrightarrow> associative (combine_matrix f)" | |
| 334 | apply (simp(no_asm) add: associative_def combine_matrix_def, auto) | |
| 335 | apply (rule comb [of Abs_matrix Abs_matrix]) | |
| 336 | by (auto, insert combine_infmatrix_assoc[of f], simp add: associative_def) | |
| 337 | ||
| 338 | lemma Rep_apply_matrix[simp]: "f 0 = 0 \<Longrightarrow> Rep_matrix (apply_matrix f A) j i = f (Rep_matrix A j i)" | |
| 339 | by (simp add: apply_matrix_def) | |
| 340 | ||
| 341 | 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)" | |
| 14662 | 342 | by(simp add: combine_matrix_def) | 
| 14593 | 343 | |
| 344 | lemma combine_nrows: "f 0 0 = 0 \<Longrightarrow> nrows (combine_matrix f A B) <= max (nrows A) (nrows B)" | |
| 345 | by (simp add: nrows_le) | |
| 346 | ||
| 347 | lemma combine_ncols: "f 0 0 = 0 \<Longrightarrow> ncols (combine_matrix f A B) <= max (ncols A) (ncols B)" | |
| 348 | by (simp add: ncols_le) | |
| 349 | ||
| 350 | lemma combine_nrows: "f 0 0 = 0 \<Longrightarrow> nrows A <= q \<Longrightarrow> nrows B <= q \<Longrightarrow> nrows(combine_matrix f A B) <= q" | |
| 351 | by (simp add: nrows_le) | |
| 352 | ||
| 353 | lemma combine_ncols: "f 0 0 = 0 \<Longrightarrow> ncols A <= q \<Longrightarrow> ncols B <= q \<Longrightarrow> ncols(combine_matrix f A B) <= q" | |
| 354 | by (simp add: ncols_le) | |
| 355 | ||
| 356 | constdefs | |
| 14654 | 357 |   zero_r_neutral :: "('a \<Rightarrow> 'b::zero \<Rightarrow> 'a) \<Rightarrow> bool"
 | 
| 14593 | 358 | "zero_r_neutral f == ! a. f a 0 = a" | 
| 14654 | 359 |   zero_l_neutral :: "('a::zero \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> bool"
 | 
| 14593 | 360 | "zero_l_neutral f == ! a. f 0 a = a" | 
| 361 |   zero_closed :: "(('a::zero) \<Rightarrow> ('b::zero) \<Rightarrow> ('c::zero)) \<Rightarrow> bool"
 | |
| 362 | "zero_closed f == (!x. f x 0 = 0) & (!y. f 0 y = 0)" | |
| 363 | ||
| 364 | consts foldseq :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> (nat \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a"
 | |
| 14662 | 365 | primrec | 
| 14593 | 366 | "foldseq f s 0 = s 0" | 
| 367 | "foldseq f s (Suc n) = f (s 0) (foldseq f (% k. s(Suc k)) n)" | |
| 368 | ||
| 369 | consts foldseq_transposed ::  "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> (nat \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a"
 | |
| 370 | primrec | |
| 371 | "foldseq_transposed f s 0 = s 0" | |
| 372 | "foldseq_transposed f s (Suc n) = f (foldseq_transposed f s n) (s (Suc n))" | |
| 373 | ||
| 374 | lemma foldseq_assoc : "associative f \<Longrightarrow> foldseq f = foldseq_transposed f" | |
| 375 | proof - | |
| 376 | assume a:"associative f" | |
| 14662 | 377 | then have sublemma: "!! n. ! N s. N <= n \<longrightarrow> foldseq f s N = foldseq_transposed f s N" | 
| 14593 | 378 | proof - | 
| 379 | fix n | |
| 14662 | 380 | show "!N s. N <= n \<longrightarrow> foldseq f s N = foldseq_transposed f s N" | 
| 381 | proof (induct n) | |
| 14593 | 382 | show "!N s. N <= 0 \<longrightarrow> foldseq f s N = foldseq_transposed f s N" by simp | 
| 383 | next | |
| 384 | fix n | |
| 385 | assume b:"! N s. N <= n \<longrightarrow> foldseq f s N = foldseq_transposed f s N" | |
| 386 | have c:"!!N s. N <= n \<Longrightarrow> foldseq f s N = foldseq_transposed f s N" by (simp add: b) | |
| 387 | show "! N t. N <= Suc n \<longrightarrow> foldseq f t N = foldseq_transposed f t N" | |
| 388 | proof (auto) | |
| 14662 | 389 | fix N t | 
| 390 | assume Nsuc: "N <= Suc n" | |
| 391 | show "foldseq f t N = foldseq_transposed f t N" | |
| 392 | proof cases | |
| 393 | assume "N <= n" | |
| 394 | then show "foldseq f t N = foldseq_transposed f t N" by (simp add: b) | |
| 395 | next | |
| 396 | assume "~(N <= n)" | |
| 397 | with Nsuc have Nsuceq: "N = Suc n" by simp | |
| 398 | have neqz: "n \<noteq> 0 \<Longrightarrow> ? m. n = Suc m & Suc m <= n" by arith | |
| 399 | have assocf: "!! x y z. f x (f y z) = f (f x y) z" by (insert a, simp add: associative_def) | |
| 400 | show "foldseq f t N = foldseq_transposed f t N" | |
| 401 | apply (simp add: Nsuceq) | |
| 402 | apply (subst c) | |
| 403 | apply (simp) | |
| 404 | apply (case_tac "n = 0") | |
| 405 | apply (simp) | |
| 406 | apply (drule neqz) | |
| 407 | apply (erule exE) | |
| 408 | apply (simp) | |
| 409 | apply (subst assocf) | |
| 410 | proof - | |
| 411 | fix m | |
| 412 | assume "n = Suc m & Suc m <= n" | |
| 413 | then have mless: "Suc m <= n" by arith | |
| 414 | then have step1: "foldseq_transposed f (% k. t (Suc k)) m = foldseq f (% k. t (Suc k)) m" (is "?T1 = ?T2") | |
| 415 | apply (subst c) | |
| 416 | by simp+ | |
| 417 | have step2: "f (t 0) ?T2 = foldseq f t (Suc m)" (is "_ = ?T3") by simp | |
| 418 | have step3: "?T3 = foldseq_transposed f t (Suc m)" (is "_ = ?T4") | |
| 419 | apply (subst c) | |
| 420 | by (simp add: mless)+ | |
| 421 | have step4: "?T4 = f (foldseq_transposed f t m) (t (Suc m))" (is "_=?T5") by simp | |
| 422 | from step1 step2 step3 step4 show sowhat: "f (f (t 0) ?T1) (t (Suc (Suc m))) = f ?T5 (t (Suc (Suc m)))" by simp | |
| 423 | qed | |
| 424 | qed | |
| 425 | qed | |
| 14593 | 426 | qed | 
| 427 | qed | |
| 428 | show "foldseq f = foldseq_transposed f" by ((rule ext)+, insert sublemma, auto) | |
| 429 | qed | |
| 430 | ||
| 431 | 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)" | |
| 432 | proof - | |
| 433 | assume assoc: "associative f" | |
| 434 | assume comm: "commutative f" | |
| 435 | from assoc have a:"!! x y z. f (f x y) z = f x (f y z)" by (simp add: associative_def) | |
| 436 | from comm have b: "!! x y. f x y = f y x" by (simp add: commutative_def) | |
| 437 | 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) | |
| 438 | have "!! n. (! u v. foldseq f (%k. f (u k) (v k)) n = f (foldseq f u n) (foldseq f v n))" | |
| 439 | apply (induct_tac n) | |
| 440 | apply (simp+, auto) | |
| 441 | by (simp add: a b c) | |
| 442 | then show "foldseq f (% k. f (u k) (v k)) n = f (foldseq f u n) (foldseq f v n)" by simp | |
| 443 | qed | |
| 444 | ||
| 445 | 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)" | |
| 446 | oops | |
| 14662 | 447 | (* Model found | 
| 14593 | 448 | |
| 449 | Trying to find a model that refutes: \<lbrakk>associative f; associative g; | |
| 450 | \<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; | |
| 451 | \<exists>x y. g x \<noteq> g y; f x x = x; g x x = x\<rbrakk> | |
| 452 | \<Longrightarrow> f = g \<or> (\<forall>y. f y x = y) \<or> (\<forall>y. g y x = y) | |
| 453 | Searching for a model of size 1, translating term... invoking SAT solver... no model found. | |
| 454 | Searching for a model of size 2, translating term... invoking SAT solver... no model found. | |
| 455 | Searching for a model of size 3, translating term... invoking SAT solver... | |
| 456 | Model found: | |
| 457 | Size of types: 'a: 3 | |
| 458 | x: a1 | |
| 459 | 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)) | |
| 460 | 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)) | |
| 461 | *) | |
| 462 | ||
| 14662 | 463 | lemma foldseq_zero: | 
| 14593 | 464 | assumes fz: "f 0 0 = 0" and sz: "! i. i <= n \<longrightarrow> s i = 0" | 
| 465 | shows "foldseq f s n = 0" | |
| 466 | proof - | |
| 467 | have "!! n. ! s. (! i. i <= n \<longrightarrow> s i = 0) \<longrightarrow> foldseq f s n = 0" | |
| 468 | apply (induct_tac n) | |
| 469 | apply (simp) | |
| 470 | by (simp add: fz) | |
| 471 | then show "foldseq f s n = 0" by (simp add: sz) | |
| 472 | qed | |
| 473 | ||
| 14662 | 474 | lemma foldseq_significant_positions: | 
| 14593 | 475 | assumes p: "! i. i <= N \<longrightarrow> S i = T i" | 
| 476 | shows "foldseq f S N = foldseq f T N" (is ?concl) | |
| 477 | proof - | |
| 478 | have "!! m . ! s t. (! i. i<=m \<longrightarrow> s i = t i) \<longrightarrow> foldseq f s m = foldseq f t m" | |
| 479 | apply (induct_tac m) | |
| 480 | apply (simp) | |
| 481 | apply (simp) | |
| 482 | apply (auto) | |
| 483 | proof - | |
| 14662 | 484 | fix n | 
| 485 | fix s::"nat\<Rightarrow>'a" | |
| 14593 | 486 | fix t::"nat\<Rightarrow>'a" | 
| 487 | assume a: "\<forall>s t. (\<forall>i\<le>n. s i = t i) \<longrightarrow> foldseq f s n = foldseq f t n" | |
| 488 | assume b: "\<forall>i\<le>Suc n. s i = t i" | |
| 489 | have c:"!! a b. a = b \<Longrightarrow> f (t 0) a = f (t 0) b" by blast | |
| 490 | 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) | |
| 491 | 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) | |
| 492 | qed | |
| 493 | with p show ?concl by simp | |
| 494 | qed | |
| 14662 | 495 | |
| 14593 | 496 | 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") | 
| 497 | proof - | |
| 498 | have suc: "!! a b. \<lbrakk>a <= Suc b; a \<noteq> Suc b\<rbrakk> \<Longrightarrow> a <= b" by arith | |
| 499 | have a:"!! a b c . a = b \<Longrightarrow> f c a = f c b" by blast | |
| 14662 | 500 | 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" | 
| 14593 | 501 | apply (induct_tac n) | 
| 502 | apply (simp) | |
| 503 | apply (simp) | |
| 504 | apply (auto) | |
| 505 | apply (case_tac "m = Suc na") | |
| 14662 | 506 | apply (simp) | 
| 14593 | 507 | apply (rule a) | 
| 508 | apply (rule foldseq_significant_positions) | |
| 15197 | 509 | apply (auto) | 
| 14593 | 510 | apply (drule suc, simp+) | 
| 511 | proof - | |
| 512 | fix na m s | |
| 513 | 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" | |
| 514 | assume subb:"m <= na" | |
| 14662 | 515 | 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 | 
| 516 | 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 = | |
| 517 | foldseq f (% k. s(Suc k)) na" | |
| 518 | by (rule subc[of m "% k. s(Suc k)", THEN sym], simp add: subb) | |
| 14593 | 519 | from subb have sube: "m \<noteq> 0 \<Longrightarrow> ? mm. m = Suc mm & mm <= na" by arith | 
| 520 | 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) = | |
| 521 | foldseq f (\<lambda>k. if k < m then s k else foldseq f (\<lambda>k. s (k + m)) (Suc na - m)) m" | |
| 14662 | 522 | apply (simp add: subd) | 
| 523 | apply (case_tac "m=0") | |
| 524 | apply (simp) | |
| 525 | apply (drule sube) | |
| 526 | apply (auto) | |
| 527 | apply (rule a) | |
| 528 | by (simp add: subc if_def) | |
| 14593 | 529 | qed | 
| 14662 | 530 | then show "?p \<Longrightarrow> ?concl" by simp | 
| 14593 | 531 | qed | 
| 14662 | 532 | |
| 14593 | 533 | lemma foldseq_zerotail: | 
| 534 | assumes | |
| 14662 | 535 | fz: "f 0 0 = 0" | 
| 536 | and sz: "! i. n <= i \<longrightarrow> s i = 0" | |
| 537 | and nm: "n <= m" | |
| 14593 | 538 | shows | 
| 14662 | 539 | "foldseq f s n = foldseq f s m" | 
| 15580 | 540 | proof - | 
| 541 | show "foldseq f s n = foldseq f s m" | |
| 14593 | 542 | apply (simp add: foldseq_tail[OF nm, of f s]) | 
| 543 | apply (rule foldseq_significant_positions) | |
| 544 | apply (auto) | |
| 545 | apply (subst foldseq_zero) | |
| 546 | by (simp add: fz sz)+ | |
| 15580 | 547 | qed | 
| 14593 | 548 | |
| 549 | lemma foldseq_zerotail2: | |
| 550 | assumes "! x. f x 0 = x" | |
| 551 | and "! i. n < i \<longrightarrow> s i = 0" | |
| 552 | and nm: "n <= m" | |
| 553 | shows | |
| 554 | "foldseq f s n = foldseq f s m" (is ?concl) | |
| 555 | proof - | |
| 556 | have "f 0 0 = 0" by (simp add: prems) | |
| 557 | have b:"!! m n. n <= m \<Longrightarrow> m \<noteq> n \<Longrightarrow> ? k. m-n = Suc k" by arith | |
| 558 | have c: "0 <= m" by simp | |
| 559 | have d: "!! k. k \<noteq> 0 \<Longrightarrow> ? l. k = Suc l" by arith | |
| 14662 | 560 | show ?concl | 
| 14593 | 561 | apply (subst foldseq_tail[OF nm]) | 
| 562 | apply (rule foldseq_significant_positions) | |
| 563 | apply (auto) | |
| 564 | apply (case_tac "m=n") | |
| 15197 | 565 | apply (simp+) | 
| 14593 | 566 | apply (drule b[OF nm]) | 
| 567 | apply (auto) | |
| 568 | apply (case_tac "k=0") | |
| 569 | apply (simp add: prems) | |
| 570 | apply (drule d) | |
| 571 | apply (auto) | |
| 572 | by (simp add: prems foldseq_zero) | |
| 573 | qed | |
| 14662 | 574 | |
| 14593 | 575 | lemma foldseq_zerostart: | 
| 14662 | 576 | "! 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))" | 
| 14593 | 577 | proof - | 
| 578 | assume f00x: "! x. f 0 (f 0 x) = f 0 x" | |
| 579 | have "! s. (! i. i<=n \<longrightarrow> s i = 0) \<longrightarrow> foldseq f s (Suc n) = f 0 (s (Suc n))" | |
| 580 | apply (induct n) | |
| 581 | apply (simp) | |
| 582 | apply (rule allI, rule impI) | |
| 583 | proof - | |
| 584 | fix n | |
| 585 | fix s | |
| 586 | have a:"foldseq f s (Suc (Suc n)) = f (s 0) (foldseq f (% k. s(Suc k)) (Suc n))" by simp | |
| 587 | assume b: "! s. ((\<forall>i\<le>n. s i = 0) \<longrightarrow> foldseq f s (Suc n) = f 0 (s (Suc n)))" | |
| 588 | 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 | |
| 589 | assume d: "! i. i <= Suc n \<longrightarrow> s i = 0" | |
| 590 | show "foldseq f s (Suc (Suc n)) = f 0 (s (Suc (Suc n)))" | |
| 14662 | 591 | apply (subst a) | 
| 592 | apply (subst c) | |
| 593 | by (simp add: d f00x)+ | |
| 14593 | 594 | qed | 
| 595 | then show "! i. i <= n \<longrightarrow> s i = 0 \<Longrightarrow> foldseq f s (Suc n) = f 0 (s (Suc n))" by simp | |
| 596 | qed | |
| 597 | ||
| 598 | lemma foldseq_zerostart2: | |
| 14662 | 599 | "! x. f 0 x = x \<Longrightarrow> ! i. i < n \<longrightarrow> s i = 0 \<Longrightarrow> foldseq f s n = s n" | 
| 14593 | 600 | proof - | 
| 601 | assume a:"! i. i<n \<longrightarrow> s i = 0" | |
| 602 | assume x:"! x. f 0 x = x" | |
| 14662 | 603 | from x have f00x: "! x. f 0 (f 0 x) = f 0 x" by blast | 
| 14593 | 604 | have b: "!! i l. i < Suc l = (i <= l)" by arith | 
| 605 | have d: "!! k. k \<noteq> 0 \<Longrightarrow> ? l. k = Suc l" by arith | |
| 14662 | 606 | show "foldseq f s n = s n" | 
| 14593 | 607 | apply (case_tac "n=0") | 
| 608 | apply (simp) | |
| 609 | apply (insert a) | |
| 610 | apply (drule d) | |
| 611 | apply (auto) | |
| 612 | apply (simp add: b) | |
| 613 | apply (insert f00x) | |
| 614 | apply (drule foldseq_zerostart) | |
| 615 | by (simp add: x)+ | |
| 616 | qed | |
| 14662 | 617 | |
| 618 | lemma foldseq_almostzero: | |
| 14593 | 619 | assumes f0x:"! x. f 0 x = x" and fx0: "! x. f x 0 = x" and s0:"! i. i \<noteq> j \<longrightarrow> s i = 0" | 
| 620 | shows "foldseq f s n = (if (j <= n) then (s j) else 0)" (is ?concl) | |
| 621 | proof - | |
| 622 | from s0 have a: "! i. i < j \<longrightarrow> s i = 0" by simp | |
| 14662 | 623 | from s0 have b: "! i. j < i \<longrightarrow> s i = 0" by simp | 
| 14593 | 624 | show ?concl | 
| 625 | apply auto | |
| 626 | apply (subst foldseq_zerotail2[of f, OF fx0, of j, OF b, of n, THEN sym]) | |
| 627 | apply simp | |
| 628 | apply (subst foldseq_zerostart2) | |
| 629 | apply (simp add: f0x a)+ | |
| 630 | apply (subst foldseq_zero) | |
| 631 | by (simp add: s0 f0x)+ | |
| 632 | qed | |
| 633 | ||
| 634 | lemma foldseq_distr_unary: | |
| 635 | assumes "!! a b. g (f a b) = f (g a) (g b)" | |
| 636 | shows "g(foldseq f s n) = foldseq f (% x. g(s x)) n" (is ?concl) | |
| 637 | proof - | |
| 638 | have "! s. g(foldseq f s n) = foldseq f (% x. g(s x)) n" | |
| 639 | apply (induct_tac n) | |
| 640 | apply (simp) | |
| 641 | apply (simp) | |
| 642 | apply (auto) | |
| 643 | apply (drule_tac x="% k. s (Suc k)" in spec) | |
| 644 | by (simp add: prems) | |
| 645 | then show ?concl by simp | |
| 646 | qed | |
| 647 | ||
| 14662 | 648 | constdefs | 
| 14593 | 649 |   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"
 | 
| 14662 | 650 | "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)" | 
| 14593 | 651 |   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"
 | 
| 652 | "mult_matrix fmul fadd A B == mult_matrix_n (max (ncols A) (nrows B)) fmul fadd A B" | |
| 653 | ||
| 14662 | 654 | lemma mult_matrix_n: | 
| 14593 | 655 | assumes prems: "ncols A \<le> n" (is ?An) "nrows B \<le> n" (is ?Bn) "fadd 0 0 = 0" "fmul 0 0 = 0" | 
| 656 | shows c:"mult_matrix fmul fadd A B = mult_matrix_n n fmul fadd A B" (is ?concl) | |
| 657 | proof - | |
| 658 | show ?concl using prems | |
| 659 | apply (simp add: mult_matrix_def mult_matrix_n_def) | |
| 660 | apply (rule comb[of "Abs_matrix" "Abs_matrix"], simp, (rule ext)+) | |
| 14662 | 661 | by (rule foldseq_zerotail, simp_all add: nrows_le ncols_le prems) | 
| 14593 | 662 | qed | 
| 663 | ||
| 664 | lemma mult_matrix_nm: | |
| 665 | assumes prems: "ncols A <= n" "nrows B <= n" "ncols A <= m" "nrows B <= m" "fadd 0 0 = 0" "fmul 0 0 = 0" | |
| 666 | shows "mult_matrix_n n fmul fadd A B = mult_matrix_n m fmul fadd A B" | |
| 667 | proof - | |
| 15580 | 668 | from prems have "mult_matrix_n n fmul fadd A B = mult_matrix fmul fadd A B" by (simp add: mult_matrix_n) | 
| 14593 | 669 | also from prems have "\<dots> = mult_matrix_n m fmul fadd A B" by (simp add: mult_matrix_n[THEN sym]) | 
| 670 | finally show "mult_matrix_n n fmul fadd A B = mult_matrix_n m fmul fadd A B" by simp | |
| 671 | qed | |
| 14662 | 672 | |
| 673 | constdefs | |
| 14593 | 674 |   r_distributive :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> bool"
 | 
| 675 | "r_distributive fmul fadd == ! a u v. fmul a (fadd u v) = fadd (fmul a u) (fmul a v)" | |
| 676 |   l_distributive :: "('a \<Rightarrow> 'b \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> bool"
 | |
| 677 | "l_distributive fmul fadd == ! a u v. fmul (fadd u v) a = fadd (fmul u a) (fmul v a)" | |
| 678 |   distributive :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> bool"
 | |
| 14662 | 679 | "distributive fmul fadd == l_distributive fmul fadd & r_distributive fmul fadd" | 
| 14593 | 680 | |
| 681 | lemma max1: "!! a x y. (a::nat) <= x \<Longrightarrow> a <= max x y" by (arith) | |
| 682 | lemma max2: "!! b x y. (b::nat) <= y \<Longrightarrow> b <= max x y" by (arith) | |
| 683 | ||
| 684 | lemma r_distributive_matrix: | |
| 14662 | 685 | assumes prems: | 
| 686 | "r_distributive fmul fadd" | |
| 687 | "associative fadd" | |
| 688 | "commutative fadd" | |
| 689 | "fadd 0 0 = 0" | |
| 690 | "! a. fmul a 0 = 0" | |
| 14593 | 691 | "! a. fmul 0 a = 0" | 
| 692 | shows "r_distributive (mult_matrix fmul fadd) (combine_matrix fadd)" (is ?concl) | |
| 693 | proof - | |
| 14662 | 694 | from prems show ?concl | 
| 14593 | 695 | apply (simp add: r_distributive_def mult_matrix_def, auto) | 
| 696 | proof - | |
| 697 | fix a::"'a matrix" | |
| 698 | fix u::"'b matrix" | |
| 699 | fix v::"'b matrix" | |
| 700 | let ?mx = "max (ncols a) (max (nrows u) (nrows v))" | |
| 701 | from prems show "mult_matrix_n (max (ncols a) (nrows (combine_matrix fadd u v))) fmul fadd a (combine_matrix fadd u v) = | |
| 702 | 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)" | |
| 15488 | 703 | apply (subst mult_matrix_nm[of _ _ _ ?mx fadd fmul]) | 
| 14662 | 704 | apply (simp add: max1 max2 combine_nrows combine_ncols)+ | 
| 15488 | 705 | apply (subst mult_matrix_nm[of _ _ v ?mx fadd fmul]) | 
| 14662 | 706 | apply (simp add: max1 max2 combine_nrows combine_ncols)+ | 
| 15488 | 707 | apply (subst mult_matrix_nm[of _ _ u ?mx fadd fmul]) | 
| 14662 | 708 | apply (simp add: max1 max2 combine_nrows combine_ncols)+ | 
| 709 | apply (simp add: mult_matrix_n_def r_distributive_def foldseq_distr[of fadd]) | |
| 710 | apply (simp add: combine_matrix_def combine_infmatrix_def) | |
| 711 | apply (rule comb[of "Abs_matrix" "Abs_matrix"], simp, (rule ext)+) | |
| 15481 | 712 | apply (simplesubst RepAbs_matrix) | 
| 14662 | 713 | apply (simp, auto) | 
| 714 | apply (rule exI[of _ "nrows a"], simp add: nrows_le foldseq_zero) | |
| 715 | apply (rule exI[of _ "ncols v"], simp add: ncols_le foldseq_zero) | |
| 716 | apply (subst 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 u"], simp add: ncols_le foldseq_zero) | |
| 720 | done | |
| 14593 | 721 | qed | 
| 722 | qed | |
| 723 | ||
| 724 | lemma l_distributive_matrix: | |
| 14662 | 725 | assumes prems: | 
| 726 | "l_distributive fmul fadd" | |
| 727 | "associative fadd" | |
| 728 | "commutative fadd" | |
| 729 | "fadd 0 0 = 0" | |
| 730 | "! a. fmul a 0 = 0" | |
| 14593 | 731 | "! a. fmul 0 a = 0" | 
| 732 | shows "l_distributive (mult_matrix fmul fadd) (combine_matrix fadd)" (is ?concl) | |
| 733 | proof - | |
| 14662 | 734 | from prems show ?concl | 
| 14593 | 735 | apply (simp add: l_distributive_def mult_matrix_def, auto) | 
| 736 | proof - | |
| 737 | fix a::"'b matrix" | |
| 738 | fix u::"'a matrix" | |
| 739 | fix v::"'a matrix" | |
| 740 | let ?mx = "max (nrows a) (max (ncols u) (ncols v))" | |
| 741 | from prems show "mult_matrix_n (max (ncols (combine_matrix fadd u v)) (nrows a)) fmul fadd (combine_matrix fadd u v) a = | |
| 742 | 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)" | |
| 15488 | 743 | apply (subst mult_matrix_nm[of v _ _ ?mx fadd fmul]) | 
| 14662 | 744 | apply (simp add: max1 max2 combine_nrows combine_ncols)+ | 
| 15488 | 745 | apply (subst mult_matrix_nm[of u _ _ ?mx fadd fmul]) | 
| 14662 | 746 | apply (simp add: max1 max2 combine_nrows combine_ncols)+ | 
| 747 | apply (subst mult_matrix_nm[of _ _ _ ?mx fadd fmul]) | |
| 748 | apply (simp add: max1 max2 combine_nrows combine_ncols)+ | |
| 749 | apply (simp add: mult_matrix_n_def l_distributive_def foldseq_distr[of fadd]) | |
| 750 | apply (simp add: combine_matrix_def combine_infmatrix_def) | |
| 751 | apply (rule comb[of "Abs_matrix" "Abs_matrix"], simp, (rule ext)+) | |
| 15481 | 752 | apply (simplesubst RepAbs_matrix) | 
| 14662 | 753 | apply (simp, auto) | 
| 754 | apply (rule exI[of _ "nrows v"], simp add: nrows_le foldseq_zero) | |
| 755 | apply (rule exI[of _ "ncols a"], simp add: ncols_le foldseq_zero) | |
| 756 | apply (subst RepAbs_matrix) | |
| 757 | apply (simp, auto) | |
| 758 | apply (rule exI[of _ "nrows u"], simp add: nrows_le foldseq_zero) | |
| 759 | apply (rule exI[of _ "ncols a"], simp add: ncols_le foldseq_zero) | |
| 760 | done | |
| 14593 | 761 | qed | 
| 762 | qed | |
| 763 | ||
| 14691 | 764 | instance matrix :: (zero) zero .. | 
| 14593 | 765 | |
| 766 | defs(overloaded) | |
| 14662 | 767 |   zero_matrix_def: "(0::('a::zero) matrix) == Abs_matrix(% j i. 0)"
 | 
| 14593 | 768 | |
| 769 | lemma Rep_zero_matrix_def[simp]: "Rep_matrix 0 j i = 0" | |
| 770 | apply (simp add: zero_matrix_def) | |
| 771 | apply (subst RepAbs_matrix) | |
| 772 | by (auto) | |
| 773 | ||
| 774 | lemma zero_matrix_def_nrows[simp]: "nrows 0 = 0" | |
| 775 | proof - | |
| 776 | have a:"!! (x::nat). x <= 0 \<Longrightarrow> x = 0" by (arith) | |
| 777 | show "nrows 0 = 0" by (rule a, subst nrows_le, simp) | |
| 778 | qed | |
| 779 | ||
| 780 | lemma zero_matrix_def_ncols[simp]: "ncols 0 = 0" | |
| 781 | proof - | |
| 782 | have a:"!! (x::nat). x <= 0 \<Longrightarrow> x = 0" by (arith) | |
| 783 | show "ncols 0 = 0" by (rule a, subst ncols_le, simp) | |
| 784 | qed | |
| 14662 | 785 | |
| 14940 | 786 | lemma combine_matrix_zero_l_neutral: "zero_l_neutral f \<Longrightarrow> zero_l_neutral (combine_matrix f)" | 
| 787 | by (simp add: zero_l_neutral_def combine_matrix_def combine_infmatrix_def) | |
| 788 | ||
| 14593 | 789 | lemma combine_matrix_zero_r_neutral: "zero_r_neutral f \<Longrightarrow> zero_r_neutral (combine_matrix f)" | 
| 790 | by (simp add: zero_r_neutral_def combine_matrix_def combine_infmatrix_def) | |
| 791 | ||
| 792 | lemma mult_matrix_zero_closed: "\<lbrakk>fadd 0 0 = 0; zero_closed fmul\<rbrakk> \<Longrightarrow> zero_closed (mult_matrix fmul fadd)" | |
| 793 | apply (simp add: zero_closed_def mult_matrix_def mult_matrix_n_def) | |
| 15580 | 794 | apply (auto) | 
| 795 | by (subst foldseq_zero, (simp add: zero_matrix_def)+)+ | |
| 14593 | 796 | |
| 797 | 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" | |
| 798 | apply (simp add: mult_matrix_n_def) | |
| 15580 | 799 | apply (subst foldseq_zero) | 
| 800 | by (simp_all add: zero_matrix_def) | |
| 14593 | 801 | |
| 802 | 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" | |
| 803 | apply (simp add: mult_matrix_n_def) | |
| 15580 | 804 | apply (subst foldseq_zero) | 
| 805 | by (simp_all add: zero_matrix_def) | |
| 14593 | 806 | |
| 807 | 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" | |
| 808 | by (simp add: mult_matrix_def) | |
| 809 | ||
| 810 | 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" | |
| 811 | by (simp add: mult_matrix_def) | |
| 812 | ||
| 813 | lemma apply_matrix_zero[simp]: "f 0 = 0 \<Longrightarrow> apply_matrix f 0 = 0" | |
| 814 | apply (simp add: apply_matrix_def apply_infmatrix_def) | |
| 815 | by (simp add: zero_matrix_def) | |
| 816 | ||
| 817 | lemma combine_matrix_zero: "f 0 0 = 0 \<Longrightarrow> combine_matrix f 0 0 = 0" | |
| 818 | apply (simp add: combine_matrix_def combine_infmatrix_def) | |
| 819 | by (simp add: zero_matrix_def) | |
| 820 | ||
| 14940 | 821 | lemma transpose_matrix_zero[simp]: "transpose_matrix 0 = 0" | 
| 822 | apply (simp add: transpose_matrix_def transpose_infmatrix_def zero_matrix_def RepAbs_matrix) | |
| 823 | apply (subst Rep_matrix_inject[symmetric], (rule ext)+) | |
| 824 | apply (simp add: RepAbs_matrix) | |
| 825 | done | |
| 826 | ||
| 14593 | 827 | lemma apply_zero_matrix_def[simp]: "apply_matrix (% x. 0) A = 0" | 
| 828 | apply (simp add: apply_matrix_def apply_infmatrix_def) | |
| 829 | by (simp add: zero_matrix_def) | |
| 830 | ||
| 831 | constdefs | |
| 832 |   singleton_matrix :: "nat \<Rightarrow> nat \<Rightarrow> ('a::zero) \<Rightarrow> 'a matrix"
 | |
| 833 | "singleton_matrix j i a == Abs_matrix(% m n. if j = m & i = n then a else 0)" | |
| 834 |   move_matrix :: "('a::zero) matrix \<Rightarrow> int \<Rightarrow> int \<Rightarrow> 'a matrix"
 | |
| 835 | "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)))" | |
| 836 |   take_rows :: "('a::zero) matrix \<Rightarrow> nat \<Rightarrow> 'a matrix"
 | |
| 837 | "take_rows A r == Abs_matrix(% j i. if (j < r) then (Rep_matrix A j i) else 0)" | |
| 838 |   take_columns :: "('a::zero) matrix \<Rightarrow> nat \<Rightarrow> 'a matrix"
 | |
| 839 | "take_columns A c == Abs_matrix(% j i. if (i < c) then (Rep_matrix A j i) else 0)" | |
| 840 | ||
| 841 | constdefs | |
| 842 |   column_of_matrix :: "('a::zero) matrix \<Rightarrow> nat \<Rightarrow> 'a matrix"
 | |
| 14662 | 843 | "column_of_matrix A n == take_columns (move_matrix A 0 (- int n)) 1" | 
| 14593 | 844 |   row_of_matrix :: "('a::zero) matrix \<Rightarrow> nat \<Rightarrow> 'a matrix"
 | 
| 845 | "row_of_matrix A m == take_rows (move_matrix A (- int m) 0) 1" | |
| 846 | ||
| 847 | lemma Rep_singleton_matrix[simp]: "Rep_matrix (singleton_matrix j i e) m n = (if j = m & i = n then e else 0)" | |
| 848 | apply (simp add: singleton_matrix_def) | |
| 849 | apply (auto) | |
| 850 | apply (subst RepAbs_matrix) | |
| 851 | apply (rule exI[of _ "Suc m"], simp) | |
| 852 | apply (rule exI[of _ "Suc n"], simp+) | |
| 15580 | 853 | by (subst RepAbs_matrix, rule exI[of _ "Suc j"], simp, rule exI[of _ "Suc i"], simp+)+ | 
| 14593 | 854 | |
| 14940 | 855 | lemma apply_singleton_matrix[simp]: "f 0 = 0 \<Longrightarrow> apply_matrix f (singleton_matrix j i x) = (singleton_matrix j i (f x))" | 
| 856 | apply (subst Rep_matrix_inject[symmetric]) | |
| 857 | apply (rule ext)+ | |
| 858 | apply (simp) | |
| 859 | done | |
| 860 | ||
| 14593 | 861 | lemma singleton_matrix_zero[simp]: "singleton_matrix j i 0 = 0" | 
| 862 | by (simp add: singleton_matrix_def zero_matrix_def) | |
| 863 | ||
| 864 | lemma nrows_singleton[simp]: "nrows(singleton_matrix j i e) = (if e = 0 then 0 else Suc j)" | |
| 865 | apply (auto) | |
| 866 | apply (rule le_anti_sym) | |
| 867 | apply (subst nrows_le) | |
| 868 | apply (simp add: singleton_matrix_def, auto) | |
| 869 | apply (subst RepAbs_matrix) | |
| 870 | apply (simp, arith) | |
| 871 | apply (simp, arith) | |
| 872 | apply (simp) | |
| 873 | apply (simp add: Suc_le_eq) | |
| 874 | apply (rule not_leE) | |
| 875 | apply (subst nrows_le) | |
| 876 | by simp | |
| 877 | ||
| 14662 | 878 | lemma ncols_singleton[simp]: "ncols(singleton_matrix j i e) = (if e = 0 then 0 else Suc i)" | 
| 14593 | 879 | apply (auto) | 
| 880 | apply (rule le_anti_sym) | |
| 881 | apply (subst ncols_le) | |
| 882 | apply (simp add: singleton_matrix_def, auto) | |
| 883 | apply (subst RepAbs_matrix) | |
| 884 | apply (simp, arith) | |
| 885 | apply (simp, arith) | |
| 886 | apply (simp) | |
| 887 | apply (simp add: Suc_le_eq) | |
| 888 | apply (rule not_leE) | |
| 889 | apply (subst ncols_le) | |
| 890 | by simp | |
| 14662 | 891 | |
| 14593 | 892 | 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)" | 
| 893 | apply (simp add: singleton_matrix_def combine_matrix_def combine_infmatrix_def) | |
| 15580 | 894 | apply (subst RepAbs_matrix) | 
| 14593 | 895 | apply (rule exI[of _ "Suc j"], simp) | 
| 896 | apply (rule exI[of _ "Suc i"], simp) | |
| 897 | apply (rule comb[of "Abs_matrix" "Abs_matrix"], simp, (rule ext)+) | |
| 898 | apply (subst RepAbs_matrix) | |
| 899 | apply (rule exI[of _ "Suc j"], simp) | |
| 900 | apply (rule exI[of _ "Suc i"], simp) | |
| 901 | by simp | |
| 902 | ||
| 14940 | 903 | lemma transpose_singleton[simp]: "transpose_matrix (singleton_matrix j i a) = singleton_matrix i j a" | 
| 904 | apply (subst Rep_matrix_inject[symmetric], (rule ext)+) | |
| 905 | apply (simp) | |
| 906 | done | |
| 907 | ||
| 14662 | 908 | lemma Rep_move_matrix[simp]: | 
| 909 | "Rep_matrix (move_matrix A y x) j i = | |
| 14593 | 910 | (if (neg ((int j)-y)) | (neg ((int i)-x)) then 0 else Rep_matrix A (nat((int j)-y)) (nat((int i)-x)))" | 
| 911 | apply (simp add: move_matrix_def) | |
| 912 | apply (auto) | |
| 15580 | 913 | by (subst RepAbs_matrix, | 
| 14593 | 914 | rule exI[of _ "(nrows A)+(nat (abs y))"], auto, rule nrows, arith, | 
| 915 | rule exI[of _ "(ncols A)+(nat (abs x))"], auto, rule ncols, arith)+ | |
| 916 | ||
| 14940 | 917 | lemma move_matrix_0_0[simp]: "move_matrix A 0 0 = A" | 
| 918 | by (simp add: move_matrix_def) | |
| 919 | ||
| 920 | lemma move_matrix_ortho: "move_matrix A j i = move_matrix (move_matrix A j 0) 0 i" | |
| 921 | apply (subst Rep_matrix_inject[symmetric]) | |
| 922 | apply (rule ext)+ | |
| 923 | apply (simp) | |
| 924 | done | |
| 925 | ||
| 926 | lemma transpose_move_matrix[simp]: | |
| 927 | "transpose_matrix (move_matrix A x y) = move_matrix (transpose_matrix A) y x" | |
| 928 | apply (subst Rep_matrix_inject[symmetric], (rule ext)+) | |
| 929 | apply (simp) | |
| 930 | done | |
| 931 | ||
| 932 | lemma move_matrix_singleton[simp]: "move_matrix (singleton_matrix u v x) j i = | |
| 933 | (if (j + int u < 0) | (i + int v < 0) then 0 else (singleton_matrix (nat (j + int u)) (nat (i + int v)) x))" | |
| 934 | apply (subst Rep_matrix_inject[symmetric]) | |
| 935 | apply (rule ext)+ | |
| 936 | apply (case_tac "j + int u < 0") | |
| 937 | apply (simp, arith) | |
| 938 | apply (case_tac "i + int v < 0") | |
| 939 | apply (simp add: neg_def, arith) | |
| 940 | apply (simp add: neg_def) | |
| 941 | apply arith | |
| 942 | done | |
| 943 | ||
| 14593 | 944 | lemma Rep_take_columns[simp]: | 
| 945 | "Rep_matrix (take_columns A c) j i = | |
| 14662 | 946 | (if i < c then (Rep_matrix A j i) else 0)" | 
| 14593 | 947 | apply (simp add: take_columns_def) | 
| 15481 | 948 | apply (simplesubst RepAbs_matrix) | 
| 14593 | 949 | apply (rule exI[of _ "nrows A"], auto, simp add: nrows_le) | 
| 950 | apply (rule exI[of _ "ncols A"], auto, simp add: ncols_le) | |
| 951 | done | |
| 952 | ||
| 953 | lemma Rep_take_rows[simp]: | |
| 954 | "Rep_matrix (take_rows A r) j i = | |
| 14662 | 955 | (if j < r then (Rep_matrix A j i) else 0)" | 
| 14593 | 956 | apply (simp add: take_rows_def) | 
| 15481 | 957 | apply (simplesubst RepAbs_matrix) | 
| 14593 | 958 | apply (rule exI[of _ "nrows A"], auto, simp add: nrows_le) | 
| 959 | apply (rule exI[of _ "ncols A"], auto, simp add: ncols_le) | |
| 960 | done | |
| 961 | ||
| 962 | lemma Rep_column_of_matrix[simp]: | |
| 963 | "Rep_matrix (column_of_matrix A c) j i = (if i = 0 then (Rep_matrix A j c) else 0)" | |
| 964 | by (simp add: column_of_matrix_def) | |
| 965 | ||
| 966 | lemma Rep_row_of_matrix[simp]: | |
| 967 | "Rep_matrix (row_of_matrix A r) j i = (if j = 0 then (Rep_matrix A r i) else 0)" | |
| 968 | by (simp add: row_of_matrix_def) | |
| 969 | ||
| 14940 | 970 | lemma column_of_matrix: "ncols A <= n \<Longrightarrow> column_of_matrix A n = 0" | 
| 971 | apply (subst Rep_matrix_inject[THEN sym]) | |
| 972 | apply (rule ext)+ | |
| 973 | by (simp add: ncols) | |
| 974 | ||
| 975 | lemma row_of_matrix: "nrows A <= n \<Longrightarrow> row_of_matrix A n = 0" | |
| 976 | apply (subst Rep_matrix_inject[THEN sym]) | |
| 977 | apply (rule ext)+ | |
| 978 | by (simp add: nrows) | |
| 979 | ||
| 14593 | 980 | lemma mult_matrix_singleton_right[simp]: | 
| 981 | assumes prems: | |
| 982 | "! x. fmul x 0 = 0" | |
| 983 | "! x. fmul 0 x = 0" | |
| 984 | "! x. fadd 0 x = x" | |
| 985 | "! x. fadd x 0 = x" | |
| 986 | 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))" | |
| 987 | apply (simp add: mult_matrix_def) | |
| 988 | apply (subst mult_matrix_nm[of _ _ _ "max (ncols A) (Suc j)"]) | |
| 989 | apply (simp add: max_def)+ | |
| 990 | apply (auto) | |
| 991 | apply (simp add: prems)+ | |
| 992 | apply (simp add: mult_matrix_n_def apply_matrix_def apply_infmatrix_def) | |
| 993 | apply (rule comb[of "Abs_matrix" "Abs_matrix"], auto, (rule ext)+) | |
| 994 | apply (subst foldseq_almostzero[of _ j]) | |
| 995 | apply (simp add: prems)+ | |
| 996 | apply (auto) | |
| 997 | apply (insert ncols_le[of A j]) | |
| 998 | apply (arith) | |
| 999 | proof - | |
| 1000 | fix k | |
| 1001 | fix l | |
| 1002 | assume a:"~neg(int l - int i)" | |
| 1003 | assume b:"nat (int l - int i) = 0" | |
| 1004 | from a b have a: "l = i" by(insert not_neg_nat[of "int l - int i"], simp add: a b) | |
| 1005 | assume c: "i \<noteq> l" | |
| 1006 | from c a show "fmul (Rep_matrix A k j) e = 0" by blast | |
| 1007 | qed | |
| 1008 | ||
| 1009 | lemma mult_matrix_ext: | |
| 1010 | assumes | |
| 1011 | eprem: | |
| 1012 | "? e. (! a b. a \<noteq> b \<longrightarrow> fmul a e \<noteq> fmul b e)" | |
| 1013 | and fprems: | |
| 1014 | "! a. fmul 0 a = 0" | |
| 1015 | "! a. fmul a 0 = 0" | |
| 14662 | 1016 | "! a. fadd a 0 = a" | 
| 14593 | 1017 | "! a. fadd 0 a = a" | 
| 1018 | and contraprems: | |
| 1019 | "mult_matrix fmul fadd A = mult_matrix fmul fadd B" | |
| 1020 | shows | |
| 1021 | "A = B" | |
| 1022 | proof(rule contrapos_np[of "False"], simp) | |
| 1023 | assume a: "A \<noteq> B" | |
| 1024 | have b: "!! f g. (! x y. f x y = g x y) \<Longrightarrow> f = g" by ((rule ext)+, auto) | |
| 1025 | have "? j i. (Rep_matrix A j i) \<noteq> (Rep_matrix B j i)" | |
| 14662 | 1026 | apply (rule contrapos_np[of "False"], simp+) | 
| 14593 | 1027 | apply (insert b[of "Rep_matrix A" "Rep_matrix B"], simp) | 
| 1028 | by (simp add: Rep_matrix_inject a) | |
| 1029 | then obtain J I where c:"(Rep_matrix A J I) \<noteq> (Rep_matrix B J I)" by blast | |
| 1030 | from eprem obtain e where eprops:"(! a b. a \<noteq> b \<longrightarrow> fmul a e \<noteq> fmul b e)" by blast | |
| 1031 | let ?S = "singleton_matrix I 0 e" | |
| 1032 | let ?comp = "mult_matrix fmul fadd" | |
| 1033 | have d: "!!x f g. f = g \<Longrightarrow> f x = g x" by blast | |
| 1034 | have e: "(% x. fmul x e) 0 = 0" by (simp add: prems) | |
| 1035 | have "~(?comp A ?S = ?comp B ?S)" | |
| 1036 | apply (rule notI) | |
| 1037 | apply (simp add: fprems eprops) | |
| 1038 | apply (simp add: Rep_matrix_inject[THEN sym]) | |
| 1039 | apply (drule d[of _ _ "J"], drule d[of _ _ "0"]) | |
| 1040 | by (simp add: e c eprops) | |
| 1041 | with contraprems show "False" by simp | |
| 1042 | qed | |
| 1043 | ||
| 1044 | constdefs | |
| 1045 |   foldmatrix :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a infmatrix) \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'a"
 | |
| 1046 | "foldmatrix f g A m n == foldseq_transposed g (% j. foldseq f (A j) n) m" | |
| 1047 |   foldmatrix_transposed :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a infmatrix) \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'a"
 | |
| 14662 | 1048 | "foldmatrix_transposed f g A m n == foldseq g (% j. foldseq_transposed f (A j) n) m" | 
| 14593 | 1049 | |
| 1050 | lemma foldmatrix_transpose: | |
| 1051 | assumes | |
| 1052 | "! a b c d. g(f a b) (f c d) = f (g a c) (g b d)" | |
| 1053 | shows | |
| 1054 | "foldmatrix f g A m n = foldmatrix_transposed g f (transpose_infmatrix A) n m" (is ?concl) | |
| 1055 | proof - | |
| 1056 | have forall:"!! P x. (! x. P x) \<Longrightarrow> P x" by auto | |
| 1057 | have tworows:"! A. foldmatrix f g A 1 n = foldmatrix_transposed g f (transpose_infmatrix A) n 1" | |
| 14662 | 1058 | apply (induct n) | 
| 14593 | 1059 | apply (simp add: foldmatrix_def foldmatrix_transposed_def prems)+ | 
| 1060 | apply (auto) | |
| 16933 | 1061 | by (drule_tac x="(% j i. A j (Suc i))" in forall, simp) | 
| 14593 | 1062 | show "foldmatrix f g A m n = foldmatrix_transposed g f (transpose_infmatrix A) n m" | 
| 1063 | apply (simp add: foldmatrix_def foldmatrix_transposed_def) | |
| 1064 | apply (induct m, simp) | |
| 14662 | 1065 | apply (simp) | 
| 14593 | 1066 | apply (insert tworows) | 
| 15236 
f289e8ba2bb3
Proofs needed to be updated because induction now preserves name of
 nipkow parents: 
15197diff
changeset | 1067 | 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) | 
| 14593 | 1068 | by (simp add: foldmatrix_def foldmatrix_transposed_def) | 
| 1069 | qed | |
| 1070 | ||
| 1071 | lemma foldseq_foldseq: | |
| 14662 | 1072 | assumes | 
| 14593 | 1073 | "associative f" | 
| 1074 | "associative g" | |
| 14662 | 1075 | "! a b c d. g(f a b) (f c d) = f (g a c) (g b d)" | 
| 1076 | shows | |
| 14593 | 1077 | "foldseq g (% j. foldseq f (A j) n) m = foldseq f (% j. foldseq g ((transpose_infmatrix A) j) m) n" | 
| 1078 | apply (insert foldmatrix_transpose[of g f A m n]) | |
| 1079 | by (simp add: foldmatrix_def foldmatrix_transposed_def foldseq_assoc[THEN sym] prems) | |
| 14662 | 1080 | |
| 1081 | lemma mult_n_nrows: | |
| 1082 | assumes | |
| 14593 | 1083 | "! a. fmul 0 a = 0" | 
| 1084 | "! a. fmul a 0 = 0" | |
| 1085 | "fadd 0 0 = 0" | |
| 1086 | shows "nrows (mult_matrix_n n fmul fadd A B) \<le> nrows A" | |
| 1087 | apply (subst nrows_le) | |
| 1088 | apply (simp add: mult_matrix_n_def) | |
| 1089 | apply (subst RepAbs_matrix) | |
| 1090 | apply (rule_tac x="nrows A" in exI) | |
| 1091 | apply (simp add: nrows prems foldseq_zero) | |
| 1092 | apply (rule_tac x="ncols B" in exI) | |
| 1093 | apply (simp add: ncols prems foldseq_zero) | |
| 1094 | by (simp add: nrows prems foldseq_zero) | |
| 1095 | ||
| 14662 | 1096 | lemma mult_n_ncols: | 
| 1097 | assumes | |
| 14593 | 1098 | "! a. fmul 0 a = 0" | 
| 1099 | "! a. fmul a 0 = 0" | |
| 1100 | "fadd 0 0 = 0" | |
| 1101 | shows "ncols (mult_matrix_n n fmul fadd A B) \<le> ncols B" | |
| 1102 | apply (subst ncols_le) | |
| 1103 | apply (simp add: mult_matrix_n_def) | |
| 1104 | apply (subst RepAbs_matrix) | |
| 1105 | apply (rule_tac x="nrows A" in exI) | |
| 1106 | apply (simp add: nrows prems foldseq_zero) | |
| 1107 | apply (rule_tac x="ncols B" in exI) | |
| 1108 | apply (simp add: ncols prems foldseq_zero) | |
| 1109 | by (simp add: ncols prems foldseq_zero) | |
| 1110 | ||
| 14662 | 1111 | lemma mult_nrows: | 
| 1112 | assumes | |
| 14593 | 1113 | "! a. fmul 0 a = 0" | 
| 1114 | "! a. fmul a 0 = 0" | |
| 1115 | "fadd 0 0 = 0" | |
| 1116 | shows "nrows (mult_matrix fmul fadd A B) \<le> nrows A" | |
| 1117 | by (simp add: mult_matrix_def mult_n_nrows prems) | |
| 1118 | ||
| 1119 | lemma mult_ncols: | |
| 14662 | 1120 | assumes | 
| 14593 | 1121 | "! a. fmul 0 a = 0" | 
| 1122 | "! a. fmul a 0 = 0" | |
| 1123 | "fadd 0 0 = 0" | |
| 1124 | shows "ncols (mult_matrix fmul fadd A B) \<le> ncols B" | |
| 1125 | by (simp add: mult_matrix_def mult_n_ncols prems) | |
| 1126 | ||
| 14940 | 1127 | lemma nrows_move_matrix_le: "nrows (move_matrix A j i) <= nat((int (nrows A)) + j)" | 
| 1128 | apply (auto simp add: nrows_le) | |
| 1129 | apply (rule nrows) | |
| 1130 | apply (arith) | |
| 1131 | done | |
| 1132 | ||
| 1133 | lemma ncols_move_matrix_le: "ncols (move_matrix A j i) <= nat((int (ncols A)) + i)" | |
| 1134 | apply (auto simp add: ncols_le) | |
| 1135 | apply (rule ncols) | |
| 1136 | apply (arith) | |
| 1137 | done | |
| 1138 | ||
| 14593 | 1139 | lemma mult_matrix_assoc: | 
| 1140 | assumes prems: | |
| 1141 | "! a. fmul1 0 a = 0" | |
| 1142 | "! a. fmul1 a 0 = 0" | |
| 1143 | "! a. fmul2 0 a = 0" | |
| 1144 | "! a. fmul2 a 0 = 0" | |
| 1145 | "fadd1 0 0 = 0" | |
| 1146 | "fadd2 0 0 = 0" | |
| 1147 | "! a b c d. fadd2 (fadd1 a b) (fadd1 c d) = fadd1 (fadd2 a c) (fadd2 b d)" | |
| 1148 | "associative fadd1" | |
| 1149 | "associative fadd2" | |
| 1150 | "! a b c. fmul2 (fmul1 a b) c = fmul1 a (fmul2 b c)" | |
| 1151 | "! a b c. fmul2 (fadd1 a b) c = fadd1 (fmul2 a c) (fmul2 b c)" | |
| 1152 | "! a b c. fmul1 c (fadd2 a b) = fadd2 (fmul1 c a) (fmul1 c b)" | |
| 1153 | 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) | |
| 1154 | proof - | |
| 1155 | 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 | |
| 1156 | have fmul2fadd1fold: "!! x s n. fmul2 (foldseq fadd1 s n) x = foldseq fadd1 (% k. fmul2 (s k) x) n" | |
| 1157 | by (rule_tac g1 = "% y. fmul2 y x" in ssubst [OF foldseq_distr_unary], simp_all!) | |
| 1158 | have fmul1fadd2fold: "!! x s n. fmul1 x (foldseq fadd2 s n) = foldseq fadd2 (% k. fmul1 x (s k)) n" | |
| 1159 | by (rule_tac g1 = "% y. fmul1 x y" in ssubst [OF foldseq_distr_unary], simp_all!) | |
| 1160 | let ?N = "max (ncols A) (max (ncols B) (max (nrows B) (nrows C)))" | |
| 1161 | show ?concl | |
| 1162 | apply (simp add: Rep_matrix_inject[THEN sym]) | |
| 1163 | apply (rule ext)+ | |
| 1164 | apply (simp add: mult_matrix_def) | |
| 15481 | 1165 | 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)"]) | 
| 14662 | 1166 | apply (simp add: max1 max2 mult_n_ncols mult_n_nrows prems)+ | 
| 15481 | 1167 | 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)+ | 
| 1168 | apply (simplesubst mult_matrix_nm[of _ _ _ "?N"]) | |
| 14593 | 1169 | apply (simp add: max1 max2 mult_n_ncols mult_n_nrows prems)+ | 
| 15481 | 1170 | apply (simplesubst mult_matrix_nm[of _ _ _ "?N"]) | 
| 14593 | 1171 | apply (simp add: max1 max2 mult_n_ncols mult_n_nrows prems)+ | 
| 15481 | 1172 | apply (simplesubst mult_matrix_nm[of _ _ _ "?N"]) | 
| 14593 | 1173 | apply (simp add: max1 max2 mult_n_ncols mult_n_nrows prems)+ | 
| 15481 | 1174 | apply (simplesubst mult_matrix_nm[of _ _ _ "?N"]) | 
| 14593 | 1175 | apply (simp add: max1 max2 mult_n_ncols mult_n_nrows prems)+ | 
| 1176 | apply (simp add: mult_matrix_n_def) | |
| 1177 | apply (rule comb_left) | |
| 1178 | apply ((rule ext)+, simp) | |
| 15481 | 1179 | apply (simplesubst RepAbs_matrix) | 
| 14593 | 1180 | apply (rule exI[of _ "nrows B"]) | 
| 1181 | apply (simp add: nrows prems foldseq_zero) | |
| 1182 | apply (rule exI[of _ "ncols C"]) | |
| 1183 | apply (simp add: prems ncols foldseq_zero) | |
| 1184 | apply (subst RepAbs_matrix) | |
| 1185 | apply (rule exI[of _ "nrows A"]) | |
| 1186 | apply (simp add: nrows prems foldseq_zero) | |
| 1187 | apply (rule exI[of _ "ncols B"]) | |
| 1188 | apply (simp add: prems ncols foldseq_zero) | |
| 1189 | apply (simp add: fmul2fadd1fold fmul1fadd2fold prems) | |
| 1190 | apply (subst foldseq_foldseq) | |
| 1191 | apply (simp add: prems)+ | |
| 1192 | by (simp add: transpose_infmatrix) | |
| 1193 | qed | |
| 14662 | 1194 | |
| 1195 | lemma | |
| 14593 | 1196 | assumes prems: | 
| 1197 | "! a. fmul1 0 a = 0" | |
| 1198 | "! a. fmul1 a 0 = 0" | |
| 1199 | "! a. fmul2 0 a = 0" | |
| 1200 | "! a. fmul2 a 0 = 0" | |
| 1201 | "fadd1 0 0 = 0" | |
| 1202 | "fadd2 0 0 = 0" | |
| 1203 | "! a b c d. fadd2 (fadd1 a b) (fadd1 c d) = fadd1 (fadd2 a c) (fadd2 b d)" | |
| 1204 | "associative fadd1" | |
| 1205 | "associative fadd2" | |
| 1206 | "! a b c. fmul2 (fmul1 a b) c = fmul1 a (fmul2 b c)" | |
| 1207 | "! a b c. fmul2 (fadd1 a b) c = fadd1 (fmul2 a c) (fmul2 b c)" | |
| 1208 | "! a b c. fmul1 c (fadd2 a b) = fadd2 (fmul1 c a) (fmul1 c b)" | |
| 1209 | shows | |
| 14662 | 1210 | "(mult_matrix fmul1 fadd1 A) o (mult_matrix fmul2 fadd2 B) = mult_matrix fmul2 fadd2 (mult_matrix fmul1 fadd1 A B)" | 
| 14593 | 1211 | apply (rule ext)+ | 
| 1212 | apply (simp add: comp_def ) | |
| 1213 | by (simp add: mult_matrix_assoc prems) | |
| 1214 | ||
| 1215 | lemma mult_matrix_assoc_simple: | |
| 1216 | assumes prems: | |
| 1217 | "! a. fmul 0 a = 0" | |
| 1218 | "! a. fmul a 0 = 0" | |
| 1219 | "fadd 0 0 = 0" | |
| 1220 | "associative fadd" | |
| 1221 | "commutative fadd" | |
| 1222 | "associative fmul" | |
| 1223 | "distributive fmul fadd" | |
| 1224 | 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) | |
| 1225 | proof - | |
| 1226 | have "!! a b c d. fadd (fadd a b) (fadd c d) = fadd (fadd a c) (fadd b d)" | |
| 1227 | by (simp! add: associative_def commutative_def) | |
| 14662 | 1228 | then show ?concl | 
| 15488 | 1229 | apply (subst mult_matrix_assoc) | 
| 14593 | 1230 | apply (simp_all!) | 
| 1231 | by (simp add: associative_def distributive_def l_distributive_def r_distributive_def)+ | |
| 1232 | qed | |
| 1233 | ||
| 1234 | lemma transpose_apply_matrix: "f 0 = 0 \<Longrightarrow> transpose_matrix (apply_matrix f A) = apply_matrix f (transpose_matrix A)" | |
| 1235 | apply (simp add: Rep_matrix_inject[THEN sym]) | |
| 1236 | apply (rule ext)+ | |
| 1237 | by simp | |
| 1238 | ||
| 1239 | 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)" | |
| 1240 | apply (simp add: Rep_matrix_inject[THEN sym]) | |
| 1241 | apply (rule ext)+ | |
| 1242 | by simp | |
| 1243 | ||
| 14662 | 1244 | lemma Rep_mult_matrix: | 
| 1245 | assumes | |
| 1246 | "! a. fmul 0 a = 0" | |
| 14593 | 1247 | "! a. fmul a 0 = 0" | 
| 1248 | "fadd 0 0 = 0" | |
| 1249 | shows | |
| 14662 | 1250 | "Rep_matrix(mult_matrix fmul fadd A B) j i = | 
| 14593 | 1251 | foldseq fadd (% k. fmul (Rep_matrix A j k) (Rep_matrix B k i)) (max (ncols A) (nrows B))" | 
| 14662 | 1252 | apply (simp add: mult_matrix_def mult_matrix_n_def) | 
| 14593 | 1253 | apply (subst RepAbs_matrix) | 
| 1254 | apply (rule exI[of _ "nrows A"], simp! add: nrows foldseq_zero) | |
| 1255 | apply (rule exI[of _ "ncols B"], simp! add: ncols foldseq_zero) | |
| 1256 | by simp | |
| 1257 | ||
| 1258 | lemma transpose_mult_matrix: | |
| 14662 | 1259 | assumes | 
| 1260 | "! a. fmul 0 a = 0" | |
| 14593 | 1261 | "! a. fmul a 0 = 0" | 
| 1262 | "fadd 0 0 = 0" | |
| 14662 | 1263 | "! x y. fmul y x = fmul x y" | 
| 14593 | 1264 | shows | 
| 1265 | "transpose_matrix (mult_matrix fmul fadd A B) = mult_matrix fmul fadd (transpose_matrix B) (transpose_matrix A)" | |
| 1266 | apply (simp add: Rep_matrix_inject[THEN sym]) | |
| 1267 | apply (rule ext)+ | |
| 1268 | by (simp! add: Rep_mult_matrix max_ac) | |
| 1269 | ||
| 14940 | 1270 | lemma column_transpose_matrix: "column_of_matrix (transpose_matrix A) n = transpose_matrix (row_of_matrix A n)" | 
| 1271 | apply (simp add: Rep_matrix_inject[THEN sym]) | |
| 1272 | apply (rule ext)+ | |
| 1273 | by simp | |
| 1274 | ||
| 1275 | lemma take_columns_transpose_matrix: "take_columns (transpose_matrix A) n = transpose_matrix (take_rows A n)" | |
| 1276 | apply (simp add: Rep_matrix_inject[THEN sym]) | |
| 1277 | apply (rule ext)+ | |
| 1278 | by simp | |
| 1279 | ||
| 14691 | 1280 | instance matrix :: ("{ord, zero}") ord ..
 | 
| 14593 | 1281 | |
| 1282 | defs (overloaded) | |
| 1283 |   le_matrix_def: "(A::('a::{ord,zero}) matrix) <= B == ! j i. (Rep_matrix A j i) <= (Rep_matrix B j i)"
 | |
| 1284 |   less_def: "(A::('a::{ord,zero}) matrix) < B == (A <= B) & (A \<noteq> B)"
 | |
| 1285 | ||
| 14662 | 1286 | instance matrix :: ("{order, zero}") order
 | 
| 14593 | 1287 | apply intro_classes | 
| 1288 | apply (simp_all add: le_matrix_def less_def) | |
| 1289 | apply (auto) | |
| 1290 | apply (drule_tac x=j in spec, drule_tac x=j in spec) | |
| 1291 | apply (drule_tac x=i in spec, drule_tac x=i in spec) | |
| 1292 | apply (simp) | |
| 1293 | apply (simp add: Rep_matrix_inject[THEN sym]) | |
| 1294 | apply (rule ext)+ | |
| 1295 | apply (drule_tac x=xa in spec, drule_tac x=xa in spec) | |
| 1296 | apply (drule_tac x=xb in spec, drule_tac x=xb in spec) | |
| 1297 | by simp | |
| 1298 | ||
| 14662 | 1299 | lemma le_apply_matrix: | 
| 14593 | 1300 | assumes | 
| 1301 | "f 0 = 0" | |
| 1302 | "! x y. x <= y \<longrightarrow> f x <= f y" | |
| 1303 |   "(a::('a::{ord, zero}) matrix) <= b"
 | |
| 1304 | shows | |
| 1305 | "apply_matrix f a <= apply_matrix f b" | |
| 1306 | by (simp! add: le_matrix_def) | |
| 1307 | ||
| 1308 | lemma le_combine_matrix: | |
| 1309 | assumes | |
| 1310 | "f 0 0 = 0" | |
| 1311 | "! a b c d. a <= b & c <= d \<longrightarrow> f a c <= f b d" | |
| 1312 | "A <= B" | |
| 1313 | "C <= D" | |
| 1314 | shows | |
| 1315 | "combine_matrix f A C <= combine_matrix f B D" | |
| 1316 | by (simp! add: le_matrix_def) | |
| 1317 | ||
| 1318 | lemma le_left_combine_matrix: | |
| 1319 | assumes | |
| 1320 | "f 0 0 = 0" | |
| 14940 | 1321 | "! a b c. a <= b \<longrightarrow> f c a <= f c b" | 
| 14593 | 1322 | "A <= B" | 
| 14662 | 1323 | shows | 
| 14593 | 1324 | "combine_matrix f C A <= combine_matrix f C B" | 
| 1325 | by (simp! add: le_matrix_def) | |
| 1326 | ||
| 1327 | lemma le_right_combine_matrix: | |
| 1328 | assumes | |
| 1329 | "f 0 0 = 0" | |
| 14940 | 1330 | "! a b c. a <= b \<longrightarrow> f a c <= f b c" | 
| 14593 | 1331 | "A <= B" | 
| 14662 | 1332 | shows | 
| 14593 | 1333 | "combine_matrix f A C <= combine_matrix f B C" | 
| 1334 | by (simp! add: le_matrix_def) | |
| 1335 | ||
| 1336 | lemma le_transpose_matrix: "(A <= B) = (transpose_matrix A <= transpose_matrix B)" | |
| 1337 | by (simp add: le_matrix_def, auto) | |
| 1338 | ||
| 1339 | lemma le_foldseq: | |
| 1340 | assumes | |
| 14662 | 1341 | "! a b c d . a <= b & c <= d \<longrightarrow> f a c <= f b d" | 
| 14593 | 1342 | "! i. i <= n \<longrightarrow> s i <= t i" | 
| 1343 | shows | |
| 1344 | "foldseq f s n <= foldseq f t n" | |
| 1345 | proof - | |
| 1346 | 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!) | |
| 1347 | then show "foldseq f s n <= foldseq f t n" by (simp!) | |
| 14662 | 1348 | qed | 
| 14593 | 1349 | |
| 1350 | lemma le_left_mult: | |
| 1351 | assumes | |
| 1352 | "! a b c d. a <= b & c <= d \<longrightarrow> fadd a c <= fadd b d" | |
| 14940 | 1353 | "! c a b. 0 <= c & a <= b \<longrightarrow> fmul c a <= fmul c b" | 
| 14662 | 1354 | "! a. fmul 0 a = 0" | 
| 14593 | 1355 | "! a. fmul a 0 = 0" | 
| 14662 | 1356 | "fadd 0 0 = 0" | 
| 14593 | 1357 | "0 <= C" | 
| 1358 | "A <= B" | |
| 1359 | shows | |
| 1360 | "mult_matrix fmul fadd C A <= mult_matrix fmul fadd C B" | |
| 1361 | apply (simp! add: le_matrix_def Rep_mult_matrix) | |
| 1362 | apply (auto) | |
| 15481 | 1363 | apply (simplesubst foldseq_zerotail[of _ _ _ "max (ncols C) (max (nrows A) (nrows B))"], simp_all add: nrows ncols max1 max2)+ | 
| 14593 | 1364 | apply (rule le_foldseq) | 
| 1365 | by (auto) | |
| 1366 | ||
| 1367 | lemma le_right_mult: | |
| 1368 | assumes | |
| 1369 | "! a b c d. a <= b & c <= d \<longrightarrow> fadd a c <= fadd b d" | |
| 1370 | "! c a b. 0 <= c & a <= b \<longrightarrow> fmul a c <= fmul b c" | |
| 14662 | 1371 | "! a. fmul 0 a = 0" | 
| 14593 | 1372 | "! a. fmul a 0 = 0" | 
| 14662 | 1373 | "fadd 0 0 = 0" | 
| 14593 | 1374 | "0 <= C" | 
| 1375 | "A <= B" | |
| 1376 | shows | |
| 1377 | "mult_matrix fmul fadd A C <= mult_matrix fmul fadd B C" | |
| 1378 | apply (simp! add: le_matrix_def Rep_mult_matrix) | |
| 1379 | apply (auto) | |
| 15481 | 1380 | apply (simplesubst foldseq_zerotail[of _ _ _ "max (nrows C) (max (ncols A) (ncols B))"], simp_all add: nrows ncols max1 max2)+ | 
| 14593 | 1381 | apply (rule le_foldseq) | 
| 1382 | by (auto) | |
| 1383 | ||
| 15178 | 1384 | lemma spec2: "! j i. P j i \<Longrightarrow> P j i" by blast | 
| 1385 | lemma neg_imp: "(\<not> Q \<longrightarrow> \<not> P) \<Longrightarrow> P \<longrightarrow> Q" by blast | |
| 1386 | ||
| 14940 | 1387 | lemma singleton_matrix_le[simp]: "(singleton_matrix j i a <= singleton_matrix j i b) = (a <= (b::_::order))" | 
| 1388 | by (auto simp add: le_matrix_def) | |
| 1389 | ||
| 15178 | 1390 | lemma singleton_le_zero[simp]: "(singleton_matrix j i x <= 0) = (x <= (0::'a::{order,zero}))"
 | 
| 1391 | apply (auto) | |
| 1392 | apply (simp add: le_matrix_def) | |
| 1393 | apply (drule_tac j=j and i=i in spec2) | |
| 1394 | apply (simp) | |
| 1395 | apply (simp add: le_matrix_def) | |
| 1396 | done | |
| 1397 | ||
| 1398 | lemma singleton_ge_zero[simp]: "(0 <= singleton_matrix j i x) = ((0::'a::{order,zero}) <= x)"
 | |
| 1399 | apply (auto) | |
| 1400 | apply (simp add: le_matrix_def) | |
| 1401 | apply (drule_tac j=j and i=i in spec2) | |
| 1402 | apply (simp) | |
| 1403 | apply (simp add: le_matrix_def) | |
| 1404 | done | |
| 1405 | ||
| 1406 | lemma move_matrix_le_zero[simp]: "0 <= j \<Longrightarrow> 0 <= i \<Longrightarrow> (move_matrix A j i <= 0) = (A <= (0::('a::{order,zero}) matrix))"
 | |
| 1407 | apply (auto simp add: le_matrix_def neg_def) | |
| 1408 | apply (drule_tac j="ja+(nat j)" and i="ia+(nat i)" in spec2) | |
| 1409 | apply (auto) | |
| 1410 | done | |
| 1411 | ||
| 1412 | lemma move_matrix_zero_le[simp]: "0 <= j \<Longrightarrow> 0 <= i \<Longrightarrow> (0 <= move_matrix A j i) = ((0::('a::{order,zero}) matrix) <= A)"
 | |
| 1413 | apply (auto simp add: le_matrix_def neg_def) | |
| 1414 | apply (drule_tac j="ja+(nat j)" and i="ia+(nat i)" in spec2) | |
| 1415 | apply (auto) | |
| 1416 | done | |
| 1417 | ||
| 1418 | 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))"
 | |
| 1419 | apply (auto simp add: le_matrix_def neg_def) | |
| 1420 | apply (drule_tac j="ja+(nat j)" and i="ia+(nat i)" in spec2) | |
| 1421 | apply (auto) | |
| 1422 | done | |
| 1423 | ||
| 14593 | 1424 | end |