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