author | nipkow |
Tue, 17 Jun 2025 14:11:40 +0200 | |
changeset 82733 | 8b537e1af2ec |
parent 80756 | 4d592706086e |
permissions | -rw-r--r-- |
47455 | 1 |
(* Title: HOL/Matrix_LP/Matrix.thy |
80756 | 2 |
Author: Steven Obua; updated to Isar by LCP |
14593 | 3 |
*) |
4 |
||
17915 | 5 |
theory Matrix |
66453
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
63566
diff
changeset
|
6 |
imports Main "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:
42463
diff
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:
44174
diff
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:
44174
diff
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:
44174
diff
changeset
|
17 |
unfolding matrix_def |
4a8743618257
prefer typedef without extra definition and alternative name;
wenzelm
parents:
44174
diff
changeset
|
18 |
proof |
4a8743618257
prefer typedef without extra definition and alternative name;
wenzelm
parents:
44174
diff
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 |
||
80736 | 25 |
lemma matrix_eqI: |
26 |
fixes A B :: "'a::zero matrix" |
|
27 |
assumes "\<And>m n. Rep_matrix A m n = Rep_matrix B m n" |
|
28 |
shows "A=B" |
|
29 |
using Rep_matrix_inject assms by blast |
|
30 |
||
27484 | 31 |
lemma finite_nonzero_positions : "finite (nonzero_positions (Rep_matrix A))" |
45694
4a8743618257
prefer typedef without extra definition and alternative name;
wenzelm
parents:
44174
diff
changeset
|
32 |
by (induct A) (simp add: Abs_matrix_inverse matrix_def) |
27484 | 33 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35032
diff
changeset
|
34 |
definition nrows :: "('a::zero) matrix \<Rightarrow> nat" where |
27484 | 35 |
"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:
35032
diff
changeset
|
36 |
|
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35032
diff
changeset
|
37 |
definition ncols :: "('a::zero) matrix \<Rightarrow> nat" where |
27484 | 38 |
"ncols A == if nonzero_positions(Rep_matrix A) = {} then 0 else Suc(Max ((image snd) (nonzero_positions (Rep_matrix A))))" |
39 |
||
40 |
lemma nrows: |
|
41 |
assumes hyp: "nrows A \<le> m" |
|
35612 | 42 |
shows "(Rep_matrix A m n) = 0" |
27484 | 43 |
proof cases |
44 |
assume "nonzero_positions(Rep_matrix A) = {}" |
|
45 |
then show "(Rep_matrix A m n) = 0" by (simp add: nonzero_positions_def) |
|
46 |
next |
|
47 |
assume a: "nonzero_positions(Rep_matrix A) \<noteq> {}" |
|
48 |
let ?S = "fst`(nonzero_positions(Rep_matrix A))" |
|
49 |
have c: "finite (?S)" by (simp add: finite_nonzero_positions) |
|
50 |
from hyp have d: "Max (?S) < m" by (simp add: a nrows_def) |
|
51 |
have "m \<notin> ?S" |
|
52 |
proof - |
|
80736 | 53 |
have "m \<in> ?S \<Longrightarrow> m \<le> Max(?S)" by (simp add: Max_ge [OF c]) |
54 |
moreover from d have "~(m \<le> Max ?S)" by (simp) |
|
27484 | 55 |
ultimately show "m \<notin> ?S" by (auto) |
56 |
qed |
|
57 |
thus "Rep_matrix A m n = 0" by (simp add: nonzero_positions_def image_Collect) |
|
58 |
qed |
|
59 |
||
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35032
diff
changeset
|
60 |
definition transpose_infmatrix :: "'a infmatrix \<Rightarrow> 'a infmatrix" where |
27484 | 61 |
"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:
35032
diff
changeset
|
62 |
|
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35032
diff
changeset
|
63 |
definition transpose_matrix :: "('a::zero) matrix \<Rightarrow> 'a matrix" where |
27484 | 64 |
"transpose_matrix == Abs_matrix o transpose_infmatrix o Rep_matrix" |
65 |
||
66 |
declare transpose_infmatrix_def[simp] |
|
67 |
||
68 |
lemma transpose_infmatrix_twice[simp]: "transpose_infmatrix (transpose_infmatrix A) = A" |
|
69 |
by ((rule ext)+, simp) |
|
70 |
||
80736 | 71 |
lemma transpose_infmatrix: "transpose_infmatrix (\<lambda>j i. P j i) = (\<lambda>j i. P i j)" |
80756 | 72 |
by force |
27484 | 73 |
|
74 |
lemma transpose_infmatrix_closed[simp]: "Rep_matrix (Abs_matrix (transpose_infmatrix (Rep_matrix x))) = transpose_infmatrix (Rep_matrix x)" |
|
75 |
proof - |
|
76 |
let ?A = "{pos. Rep_matrix x (snd pos) (fst pos) \<noteq> 0}" |
|
80756 | 77 |
let ?B = "{pos. Rep_matrix x (fst pos) (snd pos) \<noteq> 0}" |
80736 | 78 |
let ?swap = "\<lambda>pos. (snd pos, fst pos)" |
80756 | 79 |
have "finite ?A" |
80 |
proof - |
|
81 |
have swap_image: "?swap`?A = ?B" |
|
82 |
by (force simp add: image_def) |
|
83 |
then have "finite (?swap`?A)" |
|
84 |
by (metis (full_types) finite_nonzero_positions nonzero_positions_def) |
|
85 |
moreover |
|
86 |
have "inj_on ?swap ?A" by (simp add: inj_on_def) |
|
87 |
ultimately show "finite ?A" |
|
88 |
using finite_imageD by blast |
|
89 |
qed |
|
90 |
then show ?thesis |
|
91 |
by (simp add: Abs_matrix_inverse matrix_def nonzero_positions_def) |
|
27484 | 92 |
qed |
93 |
||
80736 | 94 |
lemma infmatrixforward: "(x::'a infmatrix) = y \<Longrightarrow> \<forall> a b. x a b = y a b" |
95 |
by auto |
|
27484 | 96 |
|
97 |
lemma transpose_infmatrix_inject: "(transpose_infmatrix A = transpose_infmatrix B) = (A = B)" |
|
80736 | 98 |
by (metis transpose_infmatrix_twice) |
27484 | 99 |
|
100 |
lemma transpose_matrix_inject: "(transpose_matrix A = transpose_matrix B) = (A = B)" |
|
80736 | 101 |
unfolding transpose_matrix_def o_def |
102 |
by (metis Rep_matrix_inject transpose_infmatrix_closed transpose_infmatrix_inject) |
|
27484 | 103 |
|
104 |
lemma transpose_matrix[simp]: "Rep_matrix(transpose_matrix A) j i = Rep_matrix A i j" |
|
80736 | 105 |
by (simp add: transpose_matrix_def) |
27484 | 106 |
|
107 |
lemma transpose_transpose_id[simp]: "transpose_matrix (transpose_matrix A) = A" |
|
80736 | 108 |
by (simp add: transpose_matrix_def) |
27484 | 109 |
|
110 |
lemma nrows_transpose[simp]: "nrows (transpose_matrix A) = ncols A" |
|
80736 | 111 |
by (simp add: nrows_def ncols_def nonzero_positions_def transpose_matrix_def image_def) |
27484 | 112 |
|
113 |
lemma ncols_transpose[simp]: "ncols (transpose_matrix A) = nrows A" |
|
80736 | 114 |
by (metis nrows_transpose transpose_transpose_id) |
27484 | 115 |
|
80736 | 116 |
lemma ncols: "ncols A \<le> n \<Longrightarrow> Rep_matrix A m n = 0" |
117 |
by (metis nrows nrows_transpose transpose_matrix) |
|
27484 | 118 |
|
80736 | 119 |
lemma ncols_le: "(ncols A \<le> n) \<longleftrightarrow> (\<forall>j i. n \<le> i \<longrightarrow> (Rep_matrix A j i) = 0)" (is "_ = ?st") |
80756 | 120 |
proof - |
121 |
have "Rep_matrix A j i = 0" |
|
122 |
if "ncols A \<le> n" "n \<le> i" for j i |
|
123 |
by (meson that le_trans ncols) |
|
124 |
moreover have "ncols A \<le> n" |
|
125 |
if "\<forall>j i. n \<le> i \<longrightarrow> Rep_matrix A j i = 0" |
|
126 |
unfolding ncols_def |
|
127 |
proof (clarsimp split: if_split_asm) |
|
128 |
assume \<section>: "nonzero_positions (Rep_matrix A) \<noteq> {}" |
|
129 |
let ?P = "nonzero_positions (Rep_matrix A)" |
|
130 |
let ?p = "snd`?P" |
|
131 |
have a:"finite ?p" by (simp add: finite_nonzero_positions) |
|
132 |
let ?m = "Max ?p" |
|
133 |
show "Suc (Max (snd ` nonzero_positions (Rep_matrix A))) \<le> n" |
|
134 |
using \<section> that obtains_MAX [OF finite_nonzero_positions] |
|
135 |
by (metis (mono_tags, lifting) mem_Collect_eq nonzero_positions_def not_less_eq_eq) |
|
136 |
qed |
|
137 |
ultimately show ?thesis |
|
138 |
by auto |
|
27484 | 139 |
qed |
140 |
||
80756 | 141 |
lemma less_ncols: "(n < ncols A) = (\<exists>j i. n \<le> i \<and> (Rep_matrix A j i) \<noteq> 0)" |
142 |
by (meson linorder_not_le ncols_le) |
|
27484 | 143 |
|
80736 | 144 |
lemma le_ncols: "(n \<le> ncols A) = (\<forall> m. (\<forall> j i. m \<le> i \<longrightarrow> (Rep_matrix A j i) = 0) \<longrightarrow> n \<le> m)" |
80756 | 145 |
by (meson le_trans ncols ncols_le) |
27484 | 146 |
|
80736 | 147 |
lemma nrows_le: "(nrows A \<le> n) = (\<forall>j i. n \<le> j \<longrightarrow> (Rep_matrix A j i) = 0)" (is ?s) |
80756 | 148 |
by (metis ncols_le ncols_transpose transpose_matrix) |
27484 | 149 |
|
80756 | 150 |
lemma less_nrows: "(m < nrows A) = (\<exists>j i. m \<le> j \<and> (Rep_matrix A j i) \<noteq> 0)" |
151 |
by (meson linorder_not_le nrows_le) |
|
27484 | 152 |
|
80736 | 153 |
lemma le_nrows: "(n \<le> nrows A) = (\<forall> m. (\<forall> j i. m \<le> j \<longrightarrow> (Rep_matrix A j i) = 0) \<longrightarrow> n \<le> m)" |
154 |
by (meson order.trans nrows nrows_le) |
|
27484 | 155 |
|
156 |
lemma nrows_notzero: "Rep_matrix A m n \<noteq> 0 \<Longrightarrow> m < nrows A" |
|
80736 | 157 |
by (meson leI nrows) |
27484 | 158 |
|
159 |
lemma ncols_notzero: "Rep_matrix A m n \<noteq> 0 \<Longrightarrow> n < ncols A" |
|
80736 | 160 |
by (meson leI ncols) |
27484 | 161 |
|
162 |
lemma finite_natarray1: "finite {x. x < (n::nat)}" |
|
80756 | 163 |
by simp |
27484 | 164 |
|
80756 | 165 |
lemma finite_natarray2: "finite {(x, y). x < (m::nat) \<and> y < (n::nat)}" |
80736 | 166 |
by simp |
27484 | 167 |
|
168 |
lemma RepAbs_matrix: |
|
80756 | 169 |
assumes "\<exists>m. \<forall>j i. m \<le> j \<longrightarrow> x j i = 0" |
170 |
and "\<exists>n. \<forall>j i. (n \<le> i \<longrightarrow> x j i = 0)" |
|
27484 | 171 |
shows "(Rep_matrix (Abs_matrix x)) = x" |
172 |
proof - |
|
80756 | 173 |
have "finite {pos. x (fst pos) (snd pos) \<noteq> 0}" |
174 |
proof - |
|
175 |
from assms obtain m n where a: "\<forall>j i. m \<le> j \<longrightarrow> x j i = 0" |
|
176 |
and b: "\<forall>j i. n \<le> i \<longrightarrow> x j i = 0" by (blast) |
|
177 |
let ?u = "{(i, j). x i j \<noteq> 0}" |
|
178 |
let ?v = "{(i, j). i < m \<and> j < n}" |
|
179 |
have c: "\<And>(m::nat) a. ~(m \<le> a) \<Longrightarrow> a < m" by (arith) |
|
180 |
with a b have d: "?u \<subseteq> ?v" by blast |
|
181 |
moreover have "finite ?v" by (simp add: finite_natarray2) |
|
182 |
moreover have "{pos. x (fst pos) (snd pos) \<noteq> 0} = ?u" by auto |
|
183 |
ultimately show "finite {pos. x (fst pos) (snd pos) \<noteq> 0}" |
|
184 |
by (metis (lifting) finite_subset) |
|
185 |
qed |
|
186 |
then show ?thesis |
|
187 |
by (simp add: Abs_matrix_inverse matrix_def nonzero_positions_def) |
|
27484 | 188 |
qed |
189 |
||
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35032
diff
changeset
|
190 |
definition apply_infmatrix :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a infmatrix \<Rightarrow> 'b infmatrix" where |
80736 | 191 |
"apply_infmatrix f == \<lambda>A. (\<lambda>j i. f (A j i))" |
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35032
diff
changeset
|
192 |
|
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35032
diff
changeset
|
193 |
definition apply_matrix :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a::zero) matrix \<Rightarrow> ('b::zero) matrix" where |
80736 | 194 |
"apply_matrix f == \<lambda>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:
35032
diff
changeset
|
195 |
|
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35032
diff
changeset
|
196 |
definition combine_infmatrix :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a infmatrix \<Rightarrow> 'b infmatrix \<Rightarrow> 'c infmatrix" where |
80736 | 197 |
"combine_infmatrix f == \<lambda>A B. (\<lambda>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:
35032
diff
changeset
|
198 |
|
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35032
diff
changeset
|
199 |
definition combine_matrix :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> ('a::zero) matrix \<Rightarrow> ('b::zero) matrix \<Rightarrow> ('c::zero) matrix" where |
80736 | 200 |
"combine_matrix f == \<lambda>A B. Abs_matrix (combine_infmatrix f (Rep_matrix A) (Rep_matrix B))" |
27484 | 201 |
|
202 |
lemma expand_apply_infmatrix[simp]: "apply_infmatrix f A j i = f (A j i)" |
|
80736 | 203 |
by (simp add: apply_infmatrix_def) |
27484 | 204 |
|
205 |
lemma expand_combine_infmatrix[simp]: "combine_infmatrix f A B j i = f (A j i) (B j i)" |
|
80736 | 206 |
by (simp add: combine_infmatrix_def) |
27484 | 207 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35032
diff
changeset
|
208 |
definition commutative :: "('a \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> bool" where |
80736 | 209 |
"commutative f == \<forall>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:
35032
diff
changeset
|
210 |
|
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35032
diff
changeset
|
211 |
definition associative :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> bool" where |
80736 | 212 |
"associative f == \<forall>x y z. f (f x y) z = f x (f y z)" |
27484 | 213 |
|
63167 | 214 |
text\<open> |
27484 | 215 |
To reason about associativity and commutativity of operations on matrices, |
216 |
let's take a step back and look at the general situtation: Assume that we have |
|
217 |
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. |
|
218 |
Each function $f: A \times A \rightarrow A$ now induces a function $f': B \times B \rightarrow B$ by $f' = u \circ f$. |
|
219 |
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.$ |
|
63167 | 220 |
\<close> |
27484 | 221 |
|
222 |
lemma combine_infmatrix_commute: |
|
223 |
"commutative f \<Longrightarrow> commutative (combine_infmatrix f)" |
|
224 |
by (simp add: commutative_def combine_infmatrix_def) |
|
225 |
||
226 |
lemma combine_matrix_commute: |
|
227 |
"commutative f \<Longrightarrow> commutative (combine_matrix f)" |
|
228 |
by (simp add: combine_matrix_def commutative_def combine_infmatrix_def) |
|
229 |
||
63167 | 230 |
text\<open> |
73463 | 231 |
On the contrary, given an associative function $f$ we cannot expect $f'$ to be associative. A counterexample is given by $A=\bbbZ$, $B=\{-1, 0, 1\}$, |
232 |
as $f$ we take addition on $\bbbZ$, which is clearly associative. The abstraction is given by $u(a) = 0$ for $a \notin B$. Then we have |
|
27484 | 233 |
\[ f' (f' 1 1) -1 = u(f (u (f 1 1)) -1) = u(f (u 2) -1) = u (f 0 -1) = -1, \] |
234 |
but on the other hand we have |
|
235 |
\[ f' 1 (f' 1 -1) = u (f 1 (u (f 1 -1))) = u (f 1 0) = 1.\] |
|
236 |
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: |
|
63167 | 237 |
\<close> |
27484 | 238 |
|
239 |
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)" |
|
80736 | 240 |
by (smt (verit) UnCI expand_combine_infmatrix mem_Collect_eq nonzero_positions_def subsetI) |
27484 | 241 |
|
242 |
lemma finite_nonzero_positions_Rep[simp]: "finite (nonzero_positions (Rep_matrix A))" |
|
80736 | 243 |
by (simp add: finite_nonzero_positions) |
27484 | 244 |
|
245 |
lemma combine_infmatrix_closed [simp]: |
|
246 |
"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)" |
|
80756 | 247 |
apply (rule Abs_matrix_inverse) |
248 |
apply (simp add: matrix_def) |
|
249 |
by (meson finite_Un finite_nonzero_positions_Rep finite_subset nonzero_positions_combine_infmatrix) |
|
27484 | 250 |
|
63167 | 251 |
text \<open>We need the next two lemmas only later, but it is analog to the above one, so we prove them now:\<close> |
27484 | 252 |
lemma nonzero_positions_apply_infmatrix[simp]: "f 0 = 0 \<Longrightarrow> nonzero_positions (apply_infmatrix f A) \<subseteq> nonzero_positions A" |
253 |
by (rule subsetI, simp add: nonzero_positions_def apply_infmatrix_def, auto) |
|
254 |
||
255 |
lemma apply_infmatrix_closed [simp]: |
|
256 |
"f 0 = 0 \<Longrightarrow> Rep_matrix (Abs_matrix (apply_infmatrix f (Rep_matrix A))) = apply_infmatrix f (Rep_matrix A)" |
|
257 |
apply (rule Abs_matrix_inverse) |
|
258 |
apply (simp add: matrix_def) |
|
80756 | 259 |
by (meson finite_nonzero_positions_Rep finite_subset nonzero_positions_apply_infmatrix) |
27484 | 260 |
|
261 |
lemma combine_infmatrix_assoc[simp]: "f 0 0 = 0 \<Longrightarrow> associative f \<Longrightarrow> associative (combine_infmatrix f)" |
|
80736 | 262 |
by (simp add: associative_def combine_infmatrix_def) |
27484 | 263 |
|
264 |
lemma combine_matrix_assoc: "f 0 0 = 0 \<Longrightarrow> associative f \<Longrightarrow> associative (combine_matrix f)" |
|
80756 | 265 |
by (smt (verit) associative_def combine_infmatrix_assoc combine_infmatrix_closed combine_matrix_def) |
27484 | 266 |
|
267 |
lemma Rep_apply_matrix[simp]: "f 0 = 0 \<Longrightarrow> Rep_matrix (apply_matrix f A) j i = f (Rep_matrix A j i)" |
|
80756 | 268 |
by (simp add: apply_matrix_def) |
27484 | 269 |
|
270 |
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)" |
|
271 |
by(simp add: combine_matrix_def) |
|
272 |
||
80736 | 273 |
lemma combine_nrows_max: "f 0 0 = 0 \<Longrightarrow> nrows (combine_matrix f A B) \<le> max (nrows A) (nrows B)" |
80756 | 274 |
by (simp add: nrows_le) |
27484 | 275 |
|
80736 | 276 |
lemma combine_ncols_max: "f 0 0 = 0 \<Longrightarrow> ncols (combine_matrix f A B) \<le> max (ncols A) (ncols B)" |
80756 | 277 |
by (simp add: ncols_le) |
27484 | 278 |
|
80736 | 279 |
lemma combine_nrows: "f 0 0 = 0 \<Longrightarrow> nrows A \<le> q \<Longrightarrow> nrows B \<le> q \<Longrightarrow> nrows(combine_matrix f A B) \<le> q" |
27484 | 280 |
by (simp add: nrows_le) |
281 |
||
80736 | 282 |
lemma combine_ncols: "f 0 0 = 0 \<Longrightarrow> ncols A \<le> q \<Longrightarrow> ncols B \<le> q \<Longrightarrow> ncols(combine_matrix f A B) \<le> q" |
27484 | 283 |
by (simp add: ncols_le) |
284 |
||
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35032
diff
changeset
|
285 |
definition zero_r_neutral :: "('a \<Rightarrow> 'b::zero \<Rightarrow> 'a) \<Rightarrow> bool" where |
67613 | 286 |
"zero_r_neutral f == \<forall>a. f a 0 = a" |
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35032
diff
changeset
|
287 |
|
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35032
diff
changeset
|
288 |
definition zero_l_neutral :: "('a::zero \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> bool" where |
67613 | 289 |
"zero_l_neutral f == \<forall>a. f 0 a = a" |
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35032
diff
changeset
|
290 |
|
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35032
diff
changeset
|
291 |
definition zero_closed :: "(('a::zero) \<Rightarrow> ('b::zero) \<Rightarrow> ('c::zero)) \<Rightarrow> bool" where |
80756 | 292 |
"zero_closed f == (\<forall>x. f x 0 = 0) \<and> (\<forall>y. f 0 y = 0)" |
27484 | 293 |
|
38273 | 294 |
primrec foldseq :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> (nat \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a" |
295 |
where |
|
27484 | 296 |
"foldseq f s 0 = s 0" |
80736 | 297 |
| "foldseq f s (Suc n) = f (s 0) (foldseq f (\<lambda>k. s(Suc k)) n)" |
27484 | 298 |
|
38273 | 299 |
primrec foldseq_transposed :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> (nat \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a" |
300 |
where |
|
27484 | 301 |
"foldseq_transposed f s 0 = s 0" |
38273 | 302 |
| "foldseq_transposed f s (Suc n) = f (foldseq_transposed f s n) (s (Suc n))" |
27484 | 303 |
|
80756 | 304 |
lemma foldseq_assoc: |
305 |
assumes a:"associative f" |
|
306 |
shows "associative f \<Longrightarrow> foldseq f = foldseq_transposed f" |
|
27484 | 307 |
proof - |
80756 | 308 |
have "N \<le> n \<Longrightarrow> foldseq f s N = foldseq_transposed f s N" for N s n |
309 |
proof (induct n arbitrary: N s) |
|
310 |
case 0 |
|
311 |
then show ?case |
|
312 |
by auto |
|
313 |
next |
|
314 |
case (Suc n) |
|
315 |
show ?case |
|
316 |
proof cases |
|
317 |
assume "N \<le> n" |
|
318 |
then show ?thesis |
|
319 |
by (simp add: Suc.hyps) |
|
27484 | 320 |
next |
80756 | 321 |
assume "~(N \<le> n)" |
322 |
then have Nsuceq: "N = Suc n" |
|
323 |
using Suc.prems by linarith |
|
324 |
have neqz: "n \<noteq> 0 \<Longrightarrow> \<exists>m. n = Suc m \<and> Suc m \<le> n" |
|
325 |
by arith |
|
326 |
have assocf: "!! x y z. f x (f y z) = f (f x y) z" |
|
327 |
by (metis a associative_def) |
|
328 |
have "f (f (s 0) (foldseq_transposed f (\<lambda>k. s (Suc k)) m)) (s (Suc (Suc m))) = |
|
329 |
f (f (foldseq_transposed f s m) (s (Suc m))) (s (Suc (Suc m)))" |
|
330 |
if "n = Suc m" for m |
|
331 |
proof - |
|
332 |
have \<section>: "foldseq_transposed f (\<lambda>k. s (Suc k)) m = foldseq f (\<lambda>k. s (Suc k)) m" (is "?T1 = ?T2") |
|
333 |
by (simp add: Suc.hyps that) |
|
334 |
have "f (s 0) ?T2 = foldseq f s (Suc m)" by simp |
|
335 |
also have "\<dots> = foldseq_transposed f s (Suc m)" |
|
336 |
using Suc.hyps that by blast |
|
337 |
also have "\<dots> = f (foldseq_transposed f s m) (s (Suc m))" |
|
338 |
by simp |
|
339 |
finally show ?thesis |
|
340 |
by (simp add: \<section>) |
|
27484 | 341 |
qed |
80756 | 342 |
then show "foldseq f s N = foldseq_transposed f s N" |
343 |
unfolding Nsuceq using assocf Suc.hyps neqz by force |
|
27484 | 344 |
qed |
345 |
qed |
|
80756 | 346 |
then show ?thesis |
347 |
by blast |
|
348 |
qed |
|
27484 | 349 |
|
80756 | 350 |
lemma foldseq_distr: |
351 |
assumes assoc: "associative f" and comm: "commutative f" |
|
352 |
shows "foldseq f (\<lambda>k. f (u k) (v k)) n = f (foldseq f u n) (foldseq f v n)" |
|
27484 | 353 |
proof - |
354 |
from assoc have a:"!! x y z. f (f x y) z = f x (f y z)" by (simp add: associative_def) |
|
355 |
from comm have b: "!! x y. f x y = f y x" by (simp add: commutative_def) |
|
356 |
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) |
|
80756 | 357 |
have "(\<forall>u v. foldseq f (\<lambda>k. f (u k) (v k)) n = f (foldseq f u n) (foldseq f v n))" for n |
358 |
by (induct n) (simp_all add: assoc b c foldseq_assoc) |
|
80736 | 359 |
then show "foldseq f (\<lambda>k. f (u k) (v k)) n = f (foldseq f u n) (foldseq f v n)" by simp |
27484 | 360 |
qed |
361 |
||
67613 | 362 |
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); \<exists>x y. (f x) \<noteq> (f y); \<exists>x y. (g x) \<noteq> (g y); f x x = x; g x x = x\<rbrakk> \<Longrightarrow> f=g | (\<forall>y. f y x = y) | (\<forall>y. g y x = y)" |
27484 | 363 |
oops |
364 |
(* Model found |
|
365 |
||
366 |
Trying to find a model that refutes: \<lbrakk>associative f; associative g; |
|
367 |
\<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; |
|
368 |
\<exists>x y. g x \<noteq> g y; f x x = x; g x x = x\<rbrakk> |
|
369 |
\<Longrightarrow> f = g \<or> (\<forall>y. f y x = y) \<or> (\<forall>y. g y x = y) |
|
370 |
Searching for a model of size 1, translating term... invoking SAT solver... no model found. |
|
371 |
Searching for a model of size 2, translating term... invoking SAT solver... no model found. |
|
372 |
Searching for a model of size 3, translating term... invoking SAT solver... |
|
373 |
Model found: |
|
374 |
Size of types: 'a: 3 |
|
375 |
x: a1 |
|
376 |
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)) |
|
377 |
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)) |
|
378 |
*) |
|
379 |
||
380 |
lemma foldseq_zero: |
|
80756 | 381 |
assumes fz: "f 0 0 = 0" and sz: "\<forall>i. i \<le> n \<longrightarrow> s i = 0" |
382 |
shows "foldseq f s n = 0" |
|
27484 | 383 |
proof - |
80756 | 384 |
have "\<forall>s. (\<forall>i. i \<le> n \<longrightarrow> s i = 0) \<longrightarrow> foldseq f s n = 0" for n |
385 |
by (induct n) (simp_all add: fz) |
|
386 |
then show ?thesis |
|
387 |
by (simp add: sz) |
|
27484 | 388 |
qed |
389 |
||
390 |
lemma foldseq_significant_positions: |
|
80736 | 391 |
assumes p: "\<forall>i. i \<le> N \<longrightarrow> S i = T i" |
35612 | 392 |
shows "foldseq f S N = foldseq f T N" |
80756 | 393 |
using assms |
394 |
proof (induction N arbitrary: S T) |
|
395 |
case 0 |
|
396 |
then show ?case by simp |
|
397 |
next |
|
398 |
case (Suc N) |
|
399 |
then show ?case |
|
400 |
unfolding foldseq.simps by (metis not_less_eq_eq le0) |
|
27484 | 401 |
qed |
402 |
||
35612 | 403 |
lemma foldseq_tail: |
80736 | 404 |
assumes "M \<le> N" |
405 |
shows "foldseq f S N = foldseq f (\<lambda>k. (if k < M then (S k) else (foldseq f (\<lambda>k. S(k+M)) (N-M)))) M" |
|
80756 | 406 |
using assms |
407 |
proof (induction N arbitrary: M S) |
|
408 |
case 0 |
|
409 |
then show ?case by auto |
|
410 |
next |
|
411 |
case (Suc N) |
|
412 |
show ?case |
|
413 |
proof (cases "M = Suc N") |
|
414 |
case True |
|
415 |
then show ?thesis |
|
416 |
by (auto intro!: arg_cong [of concl: "f (S 0)"] foldseq_significant_positions) |
|
417 |
next |
|
418 |
case False |
|
419 |
then have "M\<le>N" |
|
420 |
using Suc.prems by force |
|
421 |
show ?thesis |
|
422 |
proof (cases "M = 0") |
|
423 |
case True |
|
424 |
then show ?thesis |
|
425 |
by auto |
|
426 |
next |
|
427 |
case False |
|
428 |
then obtain M' where M': "M = Suc M'" "M' \<le> N" |
|
429 |
by (metis Suc_leD \<open>M \<le> N\<close> nat.nchotomy) |
|
430 |
then show ?thesis |
|
431 |
apply (simp add: Suc.IH [OF \<open>M'\<le>N\<close>]) |
|
432 |
using add_Suc_right diff_Suc_Suc by presburger |
|
27484 | 433 |
qed |
80756 | 434 |
qed |
27484 | 435 |
qed |
436 |
||
437 |
lemma foldseq_zerotail: |
|
80736 | 438 |
assumes fz: "f 0 0 = 0" and sz: "\<forall>i. n \<le> i \<longrightarrow> s i = 0" and nm: "n \<le> m" |
439 |
shows "foldseq f s n = foldseq f s m" |
|
440 |
unfolding foldseq_tail[OF nm] |
|
441 |
by (metis (no_types, lifting) foldseq_zero fz le_add2 linorder_not_le sz) |
|
27484 | 442 |
|
443 |
lemma foldseq_zerotail2: |
|
67613 | 444 |
assumes "\<forall>x. f x 0 = x" |
445 |
and "\<forall>i. n < i \<longrightarrow> s i = 0" |
|
80736 | 446 |
and nm: "n \<le> m" |
80756 | 447 |
shows "foldseq f s n = foldseq f s m" |
27484 | 448 |
proof - |
80756 | 449 |
have "s i = (if i < n then s i else foldseq f (\<lambda>k. s (k + n)) (m - n))" |
450 |
if "i\<le>n" for i |
|
451 |
proof (cases "m=n") |
|
452 |
case True |
|
453 |
then show ?thesis |
|
454 |
using that by auto |
|
455 |
next |
|
456 |
case False |
|
457 |
then obtain k where "m-n = Suc k" |
|
458 |
by (metis Suc_diff_Suc le_neq_implies_less nm) |
|
459 |
then show ?thesis |
|
460 |
apply simp |
|
461 |
by (simp add: assms(1,2) foldseq_zero nat_less_le that) |
|
462 |
qed |
|
463 |
then show ?thesis |
|
464 |
unfolding foldseq_tail[OF nm] |
|
465 |
by (auto intro: foldseq_significant_positions) |
|
27484 | 466 |
qed |
467 |
||
468 |
lemma foldseq_zerostart: |
|
80756 | 469 |
assumes f00x: "\<forall>x. f 0 (f 0 x) = f 0 x" and 0: "\<forall>i. i \<le> n \<longrightarrow> s i = 0" |
470 |
shows "foldseq f s (Suc n) = f 0 (s (Suc n))" |
|
471 |
using 0 |
|
472 |
proof (induction n arbitrary: s) |
|
473 |
case 0 |
|
474 |
then show ?case by auto |
|
475 |
next |
|
476 |
case (Suc n s) |
|
477 |
then show ?case |
|
478 |
apply (simp add: le_Suc_eq) |
|
479 |
by (smt (verit, ccfv_threshold) Suc.prems Suc_le_mono f00x foldseq_significant_positions le0) |
|
27484 | 480 |
qed |
481 |
||
482 |
lemma foldseq_zerostart2: |
|
80756 | 483 |
assumes x: "\<forall>x. f 0 x = x" and 0: "\<forall>i. i < n \<longrightarrow> s i = 0" |
484 |
shows "foldseq f s n = s n" |
|
27484 | 485 |
proof - |
486 |
show "foldseq f s n = s n" |
|
80756 | 487 |
proof (cases n) |
488 |
case 0 |
|
489 |
then show ?thesis |
|
490 |
by auto |
|
491 |
next |
|
492 |
case (Suc n') |
|
493 |
then show ?thesis |
|
494 |
by (metis "0" foldseq_zerostart le_imp_less_Suc x) |
|
495 |
qed |
|
27484 | 496 |
qed |
497 |
||
498 |
lemma foldseq_almostzero: |
|
67613 | 499 |
assumes f0x: "\<forall>x. f 0 x = x" and fx0: "\<forall>x. f x 0 = x" and s0: "\<forall>i. i \<noteq> j \<longrightarrow> s i = 0" |
80736 | 500 |
shows "foldseq f s n = (if (j \<le> n) then (s j) else 0)" |
80756 | 501 |
by (smt (verit, ccfv_SIG) f0x foldseq_zerostart2 foldseq_zerotail2 fx0 le_refl nat_less_le s0) |
27484 | 502 |
|
503 |
lemma foldseq_distr_unary: |
|
80756 | 504 |
assumes "\<And>a b. g (f a b) = f (g a) (g b)" |
80736 | 505 |
shows "g(foldseq f s n) = foldseq f (\<lambda>x. g(s x)) n" |
80756 | 506 |
proof (induction n arbitrary: s) |
507 |
case 0 |
|
508 |
then show ?case |
|
509 |
by auto |
|
510 |
next |
|
511 |
case (Suc n) |
|
512 |
then show ?case |
|
513 |
using assms by fastforce |
|
27484 | 514 |
qed |
515 |
||
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35032
diff
changeset
|
516 |
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 |
80736 | 517 |
"mult_matrix_n n fmul fadd A B == Abs_matrix(\<lambda>j i. foldseq fadd (\<lambda>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:
35032
diff
changeset
|
518 |
|
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35032
diff
changeset
|
519 |
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 | 520 |
"mult_matrix fmul fadd A B == mult_matrix_n (max (ncols A) (nrows B)) fmul fadd A B" |
521 |
||
522 |
lemma mult_matrix_n: |
|
80756 | 523 |
assumes "ncols A \<le> n" "nrows B \<le> n" "fadd 0 0 = 0" "fmul 0 0 = 0" |
524 |
shows "mult_matrix fmul fadd A B = mult_matrix_n n fmul fadd A B" |
|
27484 | 525 |
proof - |
80756 | 526 |
have "foldseq fadd (\<lambda>k. fmul (Rep_matrix A j k) (Rep_matrix B k i)) |
527 |
(max (ncols A) (nrows B)) = |
|
528 |
foldseq fadd (\<lambda>k. fmul (Rep_matrix A j k) (Rep_matrix B k i)) n" for i j |
|
529 |
using assms by (simp add: foldseq_zerotail nrows_le ncols_le) |
|
530 |
then show ?thesis |
|
531 |
by (simp add: mult_matrix_def mult_matrix_n_def) |
|
27484 | 532 |
qed |
533 |
||
534 |
lemma mult_matrix_nm: |
|
80736 | 535 |
assumes "ncols A \<le> n" "nrows B \<le> n" "ncols A \<le> m" "nrows B \<le> m" "fadd 0 0 = 0" "fmul 0 0 = 0" |
27484 | 536 |
shows "mult_matrix_n n fmul fadd A B = mult_matrix_n m fmul fadd A B" |
537 |
proof - |
|
35612 | 538 |
from assms have "mult_matrix_n n fmul fadd A B = mult_matrix fmul fadd A B" |
539 |
by (simp add: mult_matrix_n) |
|
540 |
also from assms have "\<dots> = mult_matrix_n m fmul fadd A B" |
|
541 |
by (simp add: mult_matrix_n[THEN sym]) |
|
27484 | 542 |
finally show "mult_matrix_n n fmul fadd A B = mult_matrix_n m fmul fadd A B" by simp |
543 |
qed |
|
544 |
||
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35032
diff
changeset
|
545 |
definition r_distributive :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> bool" where |
67613 | 546 |
"r_distributive fmul fadd == \<forall>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:
35032
diff
changeset
|
547 |
|
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35032
diff
changeset
|
548 |
definition l_distributive :: "('a \<Rightarrow> 'b \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> bool" where |
67613 | 549 |
"l_distributive fmul fadd == \<forall>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:
35032
diff
changeset
|
550 |
|
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35032
diff
changeset
|
551 |
definition distributive :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> bool" where |
80756 | 552 |
"distributive fmul fadd == l_distributive fmul fadd \<and> r_distributive fmul fadd" |
27484 | 553 |
|
80736 | 554 |
lemma max1: "!! a x y. (a::nat) \<le> x \<Longrightarrow> a \<le> max x y" by (arith) |
555 |
lemma max2: "!! b x y. (b::nat) \<le> y \<Longrightarrow> b \<le> max x y" by (arith) |
|
27484 | 556 |
|
557 |
lemma r_distributive_matrix: |
|
35612 | 558 |
assumes |
27484 | 559 |
"r_distributive fmul fadd" |
560 |
"associative fadd" |
|
561 |
"commutative fadd" |
|
562 |
"fadd 0 0 = 0" |
|
67613 | 563 |
"\<forall>a. fmul a 0 = 0" |
564 |
"\<forall>a. fmul 0 a = 0" |
|
35612 | 565 |
shows "r_distributive (mult_matrix fmul fadd) (combine_matrix fadd)" |
27484 | 566 |
proof - |
35612 | 567 |
from assms show ?thesis |
27484 | 568 |
apply (simp add: r_distributive_def mult_matrix_def, auto) |
569 |
proof - |
|
570 |
fix a::"'a matrix" |
|
571 |
fix u::"'b matrix" |
|
572 |
fix v::"'b matrix" |
|
573 |
let ?mx = "max (ncols a) (max (nrows u) (nrows v))" |
|
35612 | 574 |
from assms show "mult_matrix_n (max (ncols a) (nrows (combine_matrix fadd u v))) fmul fadd a (combine_matrix fadd u v) = |
27484 | 575 |
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)" |
576 |
apply (subst mult_matrix_nm[of _ _ _ ?mx fadd fmul]) |
|
577 |
apply (simp add: max1 max2 combine_nrows combine_ncols)+ |
|
578 |
apply (subst mult_matrix_nm[of _ _ v ?mx fadd fmul]) |
|
579 |
apply (simp add: max1 max2 combine_nrows combine_ncols)+ |
|
580 |
apply (subst mult_matrix_nm[of _ _ u ?mx fadd fmul]) |
|
581 |
apply (simp add: max1 max2 combine_nrows combine_ncols)+ |
|
582 |
apply (simp add: mult_matrix_n_def r_distributive_def foldseq_distr[of fadd]) |
|
583 |
apply (simp add: combine_matrix_def combine_infmatrix_def) |
|
80756 | 584 |
apply (intro ext arg_cong[of concl: "Abs_matrix"]) |
27484 | 585 |
apply (simplesubst RepAbs_matrix) |
586 |
apply (simp, auto) |
|
587 |
apply (rule exI[of _ "nrows a"], simp add: nrows_le foldseq_zero) |
|
588 |
apply (rule exI[of _ "ncols v"], simp add: ncols_le foldseq_zero) |
|
589 |
apply (subst RepAbs_matrix) |
|
590 |
apply (simp, auto) |
|
591 |
apply (rule exI[of _ "nrows a"], simp add: nrows_le foldseq_zero) |
|
592 |
apply (rule exI[of _ "ncols u"], simp add: ncols_le foldseq_zero) |
|
593 |
done |
|
594 |
qed |
|
595 |
qed |
|
596 |
||
597 |
lemma l_distributive_matrix: |
|
35612 | 598 |
assumes |
27484 | 599 |
"l_distributive fmul fadd" |
600 |
"associative fadd" |
|
601 |
"commutative fadd" |
|
602 |
"fadd 0 0 = 0" |
|
67613 | 603 |
"\<forall>a. fmul a 0 = 0" |
604 |
"\<forall>a. fmul 0 a = 0" |
|
35612 | 605 |
shows "l_distributive (mult_matrix fmul fadd) (combine_matrix fadd)" |
27484 | 606 |
proof - |
35612 | 607 |
from assms show ?thesis |
27484 | 608 |
apply (simp add: l_distributive_def mult_matrix_def, auto) |
609 |
proof - |
|
610 |
fix a::"'b matrix" |
|
611 |
fix u::"'a matrix" |
|
612 |
fix v::"'a matrix" |
|
613 |
let ?mx = "max (nrows a) (max (ncols u) (ncols v))" |
|
35612 | 614 |
from assms show "mult_matrix_n (max (ncols (combine_matrix fadd u v)) (nrows a)) fmul fadd (combine_matrix fadd u v) a = |
27484 | 615 |
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)" |
616 |
apply (subst mult_matrix_nm[of v _ _ ?mx fadd fmul]) |
|
617 |
apply (simp add: max1 max2 combine_nrows combine_ncols)+ |
|
618 |
apply (subst mult_matrix_nm[of u _ _ ?mx fadd fmul]) |
|
619 |
apply (simp add: max1 max2 combine_nrows combine_ncols)+ |
|
620 |
apply (subst mult_matrix_nm[of _ _ _ ?mx fadd fmul]) |
|
621 |
apply (simp add: max1 max2 combine_nrows combine_ncols)+ |
|
622 |
apply (simp add: mult_matrix_n_def l_distributive_def foldseq_distr[of fadd]) |
|
623 |
apply (simp add: combine_matrix_def combine_infmatrix_def) |
|
80756 | 624 |
apply (intro ext arg_cong[of concl: "Abs_matrix"]) |
27484 | 625 |
apply (simplesubst RepAbs_matrix) |
626 |
apply (simp, auto) |
|
627 |
apply (rule exI[of _ "nrows v"], simp add: nrows_le foldseq_zero) |
|
628 |
apply (rule exI[of _ "ncols a"], simp add: ncols_le foldseq_zero) |
|
629 |
apply (subst RepAbs_matrix) |
|
630 |
apply (simp, auto) |
|
631 |
apply (rule exI[of _ "nrows u"], simp add: nrows_le foldseq_zero) |
|
632 |
apply (rule exI[of _ "ncols a"], simp add: ncols_le foldseq_zero) |
|
633 |
done |
|
634 |
qed |
|
635 |
qed |
|
636 |
||
637 |
instantiation matrix :: (zero) zero |
|
638 |
begin |
|
639 |
||
37765 | 640 |
definition zero_matrix_def: "0 = Abs_matrix (\<lambda>j i. 0)" |
27484 | 641 |
|
642 |
instance .. |
|
643 |
||
644 |
end |
|
645 |
||
646 |
lemma Rep_zero_matrix_def[simp]: "Rep_matrix 0 j i = 0" |
|
80756 | 647 |
by (simp add: RepAbs_matrix zero_matrix_def) |
27484 | 648 |
|
649 |
lemma zero_matrix_def_nrows[simp]: "nrows 0 = 0" |
|
80756 | 650 |
using nrows_le by force |
27484 | 651 |
|
652 |
lemma zero_matrix_def_ncols[simp]: "ncols 0 = 0" |
|
80756 | 653 |
using ncols_le by fastforce |
27484 | 654 |
|
655 |
lemma combine_matrix_zero_l_neutral: "zero_l_neutral f \<Longrightarrow> zero_l_neutral (combine_matrix f)" |
|
656 |
by (simp add: zero_l_neutral_def combine_matrix_def combine_infmatrix_def) |
|
657 |
||
658 |
lemma combine_matrix_zero_r_neutral: "zero_r_neutral f \<Longrightarrow> zero_r_neutral (combine_matrix f)" |
|
659 |
by (simp add: zero_r_neutral_def combine_matrix_def combine_infmatrix_def) |
|
660 |
||
661 |
lemma mult_matrix_zero_closed: "\<lbrakk>fadd 0 0 = 0; zero_closed fmul\<rbrakk> \<Longrightarrow> zero_closed (mult_matrix fmul fadd)" |
|
662 |
apply (simp add: zero_closed_def mult_matrix_def mult_matrix_n_def) |
|
80756 | 663 |
by (simp add: foldseq_zero zero_matrix_def) |
27484 | 664 |
|
67613 | 665 |
lemma mult_matrix_n_zero_right[simp]: "\<lbrakk>fadd 0 0 = 0; \<forall>a. fmul a 0 = 0\<rbrakk> \<Longrightarrow> mult_matrix_n n fmul fadd A 0 = 0" |
80756 | 666 |
by (simp add: RepAbs_matrix foldseq_zero matrix_eqI mult_matrix_n_def) |
27484 | 667 |
|
67613 | 668 |
lemma mult_matrix_n_zero_left[simp]: "\<lbrakk>fadd 0 0 = 0; \<forall>a. fmul 0 a = 0\<rbrakk> \<Longrightarrow> mult_matrix_n n fmul fadd 0 A = 0" |
80756 | 669 |
by (simp add: RepAbs_matrix foldseq_zero matrix_eqI mult_matrix_n_def) |
27484 | 670 |
|
67613 | 671 |
lemma mult_matrix_zero_left[simp]: "\<lbrakk>fadd 0 0 = 0; \<forall>a. fmul 0 a = 0\<rbrakk> \<Longrightarrow> mult_matrix fmul fadd 0 A = 0" |
80756 | 672 |
by (simp add: mult_matrix_def) |
27484 | 673 |
|
67613 | 674 |
lemma mult_matrix_zero_right[simp]: "\<lbrakk>fadd 0 0 = 0; \<forall>a. fmul a 0 = 0\<rbrakk> \<Longrightarrow> mult_matrix fmul fadd A 0 = 0" |
80756 | 675 |
by (simp add: mult_matrix_def) |
27484 | 676 |
|
677 |
lemma apply_matrix_zero[simp]: "f 0 = 0 \<Longrightarrow> apply_matrix f 0 = 0" |
|
80756 | 678 |
by (simp add: matrix_eqI) |
27484 | 679 |
|
680 |
lemma combine_matrix_zero: "f 0 0 = 0 \<Longrightarrow> combine_matrix f 0 0 = 0" |
|
80756 | 681 |
by (simp add: matrix_eqI) |
27484 | 682 |
|
683 |
lemma transpose_matrix_zero[simp]: "transpose_matrix 0 = 0" |
|
80756 | 684 |
by (simp add: matrix_eqI) |
27484 | 685 |
|
80736 | 686 |
lemma apply_zero_matrix_def[simp]: "apply_matrix (\<lambda>x. 0) A = 0" |
80756 | 687 |
by (simp add: matrix_eqI) |
27484 | 688 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35032
diff
changeset
|
689 |
definition singleton_matrix :: "nat \<Rightarrow> nat \<Rightarrow> ('a::zero) \<Rightarrow> 'a matrix" where |
80756 | 690 |
"singleton_matrix j i a == Abs_matrix(\<lambda>m n. if j = m \<and> i = n then a else 0)" |
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35032
diff
changeset
|
691 |
|
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35032
diff
changeset
|
692 |
definition move_matrix :: "('a::zero) matrix \<Rightarrow> int \<Rightarrow> int \<Rightarrow> 'a matrix" where |
80736 | 693 |
"move_matrix A y x == Abs_matrix(\<lambda>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:
35032
diff
changeset
|
694 |
|
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35032
diff
changeset
|
695 |
definition take_rows :: "('a::zero) matrix \<Rightarrow> nat \<Rightarrow> 'a matrix" where |
80736 | 696 |
"take_rows A r == Abs_matrix(\<lambda>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:
35032
diff
changeset
|
697 |
|
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35032
diff
changeset
|
698 |
definition take_columns :: "('a::zero) matrix \<Rightarrow> nat \<Rightarrow> 'a matrix" where |
80736 | 699 |
"take_columns A c == Abs_matrix(\<lambda>j i. if (i < c) then (Rep_matrix A j i) else 0)" |
27484 | 700 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35032
diff
changeset
|
701 |
definition column_of_matrix :: "('a::zero) matrix \<Rightarrow> nat \<Rightarrow> 'a matrix" where |
27484 | 702 |
"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:
35032
diff
changeset
|
703 |
|
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35032
diff
changeset
|
704 |
definition row_of_matrix :: "('a::zero) matrix \<Rightarrow> nat \<Rightarrow> 'a matrix" where |
27484 | 705 |
"row_of_matrix A m == take_rows (move_matrix A (- int m) 0) 1" |
706 |
||
80756 | 707 |
lemma Rep_singleton_matrix[simp]: "Rep_matrix (singleton_matrix j i e) m n = (if j = m \<and> i = n then e else 0)" |
708 |
unfolding singleton_matrix_def |
|
709 |
by (smt (verit, del_insts) RepAbs_matrix Suc_n_not_le_n) |
|
27484 | 710 |
|
711 |
lemma apply_singleton_matrix[simp]: "f 0 = 0 \<Longrightarrow> apply_matrix f (singleton_matrix j i x) = (singleton_matrix j i (f x))" |
|
80736 | 712 |
by (simp add: matrix_eqI) |
27484 | 713 |
|
714 |
lemma singleton_matrix_zero[simp]: "singleton_matrix j i 0 = 0" |
|
715 |
by (simp add: singleton_matrix_def zero_matrix_def) |
|
716 |
||
717 |
lemma nrows_singleton[simp]: "nrows(singleton_matrix j i e) = (if e = 0 then 0 else Suc j)" |
|
80756 | 718 |
proof - |
719 |
have "e \<noteq> 0 \<Longrightarrow> Suc j \<le> nrows (singleton_matrix j i e)" |
|
720 |
by (metis Rep_singleton_matrix not_less_eq_eq nrows) |
|
721 |
then show ?thesis |
|
722 |
by (simp add: le_antisym nrows_le) |
|
27484 | 723 |
qed |
724 |
||
725 |
lemma ncols_singleton[simp]: "ncols(singleton_matrix j i e) = (if e = 0 then 0 else Suc i)" |
|
80756 | 726 |
by (simp add: Suc_leI le_antisym ncols_le ncols_notzero) |
27484 | 727 |
|
728 |
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)" |
|
80756 | 729 |
apply (simp add: singleton_matrix_def combine_matrix_def combine_infmatrix_def) |
730 |
apply (intro ext arg_cong[of concl: "Abs_matrix"]) |
|
731 |
by (metis Rep_singleton_matrix singleton_matrix_def) |
|
27484 | 732 |
|
733 |
lemma transpose_singleton[simp]: "transpose_matrix (singleton_matrix j i a) = singleton_matrix i j a" |
|
80736 | 734 |
by (simp add: matrix_eqI) |
27484 | 735 |
|
736 |
lemma Rep_move_matrix[simp]: |
|
737 |
"Rep_matrix (move_matrix A y x) j i = |
|
46702 | 738 |
(if (((int j)-y) < 0) | (((int i)-x) < 0) then 0 else Rep_matrix A (nat((int j)-y)) (nat((int i)-x)))" |
80756 | 739 |
apply (simp add: move_matrix_def) |
27484 | 740 |
by (subst RepAbs_matrix, |
61945 | 741 |
rule exI[of _ "(nrows A)+(nat \<bar>y\<bar>)"], auto, rule nrows, arith, |
742 |
rule exI[of _ "(ncols A)+(nat \<bar>x\<bar>)"], auto, rule ncols, arith)+ |
|
27484 | 743 |
|
744 |
lemma move_matrix_0_0[simp]: "move_matrix A 0 0 = A" |
|
80756 | 745 |
by (simp add: move_matrix_def) |
27484 | 746 |
|
747 |
lemma move_matrix_ortho: "move_matrix A j i = move_matrix (move_matrix A j 0) 0 i" |
|
80736 | 748 |
by (simp add: matrix_eqI) |
27484 | 749 |
|
750 |
lemma transpose_move_matrix[simp]: |
|
751 |
"transpose_matrix (move_matrix A x y) = move_matrix (transpose_matrix A) y x" |
|
80736 | 752 |
by (simp add: matrix_eqI) |
27484 | 753 |
|
80736 | 754 |
lemma move_matrix_singleton[simp]: "move_matrix (singleton_matrix u v x) j i = |
27484 | 755 |
(if (j + int u < 0) | (i + int v < 0) then 0 else (singleton_matrix (nat (j + int u)) (nat (i + int v)) x))" |
80756 | 756 |
by (auto intro!: matrix_eqI split: if_split_asm) |
27484 | 757 |
|
758 |
lemma Rep_take_columns[simp]: |
|
80756 | 759 |
"Rep_matrix (take_columns A c) j i = (if i < c then (Rep_matrix A j i) else 0)" |
760 |
unfolding take_columns_def |
|
761 |
by (smt (verit, best) RepAbs_matrix leD nrows) |
|
27484 | 762 |
|
763 |
lemma Rep_take_rows[simp]: |
|
80756 | 764 |
"Rep_matrix (take_rows A r) j i = (if j < r then (Rep_matrix A j i) else 0)" |
765 |
unfolding take_rows_def |
|
766 |
by (smt (verit, best) RepAbs_matrix leD ncols) |
|
27484 | 767 |
|
768 |
lemma Rep_column_of_matrix[simp]: |
|
769 |
"Rep_matrix (column_of_matrix A c) j i = (if i = 0 then (Rep_matrix A j c) else 0)" |
|
770 |
by (simp add: column_of_matrix_def) |
|
771 |
||
772 |
lemma Rep_row_of_matrix[simp]: |
|
773 |
"Rep_matrix (row_of_matrix A r) j i = (if j = 0 then (Rep_matrix A r i) else 0)" |
|
774 |
by (simp add: row_of_matrix_def) |
|
775 |
||
80736 | 776 |
lemma column_of_matrix: "ncols A \<le> n \<Longrightarrow> column_of_matrix A n = 0" |
777 |
by (simp add: matrix_eqI ncols) |
|
27484 | 778 |
|
80736 | 779 |
lemma row_of_matrix: "nrows A \<le> n \<Longrightarrow> row_of_matrix A n = 0" |
780 |
by (simp add: matrix_eqI nrows) |
|
27484 | 781 |
|
782 |
lemma mult_matrix_singleton_right[simp]: |
|
80756 | 783 |
assumes "\<forall>x. fmul x 0 = 0" "\<forall>x. fmul 0 x = 0" "\<forall>x. fadd 0 x = x" "\<forall>x. fadd x 0 = x" |
80736 | 784 |
shows "(mult_matrix fmul fadd A (singleton_matrix j i e)) = apply_matrix (\<lambda>x. fmul x e) (move_matrix (column_of_matrix A j) 0 (int i))" |
80756 | 785 |
using assms |
786 |
unfolding mult_matrix_def |
|
787 |
apply (subst mult_matrix_nm[of _ _ _ "max (ncols A) (Suc j)"]; |
|
788 |
simp add: mult_matrix_n_def apply_matrix_def apply_infmatrix_def) |
|
789 |
apply (intro ext arg_cong[of concl: "Abs_matrix"]) |
|
790 |
by (simp add: max_def assms foldseq_almostzero[of _ j]) |
|
27484 | 791 |
|
792 |
lemma mult_matrix_ext: |
|
793 |
assumes |
|
794 |
eprem: |
|
67613 | 795 |
"\<exists>e. (\<forall>a b. a \<noteq> b \<longrightarrow> fmul a e \<noteq> fmul b e)" |
27484 | 796 |
and fprems: |
67613 | 797 |
"\<forall>a. fmul 0 a = 0" |
798 |
"\<forall>a. fmul a 0 = 0" |
|
799 |
"\<forall>a. fadd a 0 = a" |
|
800 |
"\<forall>a. fadd 0 a = a" |
|
80736 | 801 |
and contraprems: "mult_matrix fmul fadd A = mult_matrix fmul fadd B" |
802 |
shows "A = B" |
|
80756 | 803 |
proof(rule ccontr) |
804 |
assume "A \<noteq> B" |
|
805 |
then obtain J I where ne: "(Rep_matrix A J I) \<noteq> (Rep_matrix B J I)" |
|
806 |
by (meson matrix_eqI) |
|
67613 | 807 |
from eprem obtain e where eprops:"(\<forall>a b. a \<noteq> b \<longrightarrow> fmul a e \<noteq> fmul b e)" by blast |
27484 | 808 |
let ?S = "singleton_matrix I 0 e" |
809 |
let ?comp = "mult_matrix fmul fadd" |
|
810 |
have d: "!!x f g. f = g \<Longrightarrow> f x = g x" by blast |
|
80736 | 811 |
have e: "(\<lambda>x. fmul x e) 0 = 0" by (simp add: assms) |
812 |
have "Rep_matrix (apply_matrix (\<lambda>x. fmul x e) (column_of_matrix A I)) \<noteq> |
|
813 |
Rep_matrix (apply_matrix (\<lambda>x. fmul x e) (column_of_matrix B I))" |
|
814 |
using fprems |
|
80756 | 815 |
by (metis Rep_apply_matrix Rep_column_of_matrix eprops ne) |
816 |
then have "?comp A ?S \<noteq> ?comp B ?S" |
|
80736 | 817 |
by (simp add: fprems eprops Rep_matrix_inject) |
27484 | 818 |
with contraprems show "False" by simp |
819 |
qed |
|
820 |
||
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35032
diff
changeset
|
821 |
definition foldmatrix :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a infmatrix) \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'a" where |
80736 | 822 |
"foldmatrix f g A m n == foldseq_transposed g (\<lambda>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:
35032
diff
changeset
|
823 |
|
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35032
diff
changeset
|
824 |
definition foldmatrix_transposed :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a infmatrix) \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'a" where |
80736 | 825 |
"foldmatrix_transposed f g A m n == foldseq g (\<lambda>j. foldseq_transposed f (A j) n) m" |
27484 | 826 |
|
827 |
lemma foldmatrix_transpose: |
|
80756 | 828 |
assumes "\<forall>a b c d. g(f a b) (f c d) = f (g a c) (g b d)" |
829 |
shows "foldmatrix f g A m n = foldmatrix_transposed g f (transpose_infmatrix A) n m" |
|
27484 | 830 |
proof - |
67613 | 831 |
have forall:"\<And>P x. (\<forall>x. P x) \<Longrightarrow> P x" by auto |
832 |
have tworows:"\<forall>A. foldmatrix f g A 1 n = foldmatrix_transposed g f (transpose_infmatrix A) n 1" |
|
80756 | 833 |
proof (induct n) |
834 |
case 0 |
|
835 |
then show ?case |
|
836 |
by (simp add: foldmatrix_def foldmatrix_transposed_def) |
|
837 |
next |
|
838 |
case (Suc n) |
|
839 |
then show ?case |
|
840 |
apply (clarsimp simp: foldmatrix_def foldmatrix_transposed_def assms) |
|
841 |
apply (rule arg_cong [of concl: "f _"]) |
|
842 |
by meson |
|
843 |
qed |
|
844 |
have "foldseq_transposed g (\<lambda>j. foldseq f (A j) n) m = |
|
845 |
foldseq f (\<lambda>j. foldseq_transposed g (transpose_infmatrix A j) m) n" |
|
846 |
proof (induct m) |
|
847 |
case 0 |
|
848 |
then show ?case by auto |
|
849 |
next |
|
850 |
case (Suc m) |
|
851 |
then show ?case |
|
852 |
using tworows |
|
853 |
apply (drule_tac x="\<lambda>j i. (if j = 0 then (foldseq_transposed g (\<lambda>u. A u i) m) else (A (Suc m) i))" in spec) |
|
854 |
by (simp add: Suc foldmatrix_def foldmatrix_transposed_def) |
|
855 |
qed |
|
856 |
then show "foldmatrix f g A m n = foldmatrix_transposed g f (transpose_infmatrix A) n m" |
|
27484 | 857 |
by (simp add: foldmatrix_def foldmatrix_transposed_def) |
858 |
qed |
|
859 |
||
860 |
lemma foldseq_foldseq: |
|
80756 | 861 |
assumes "associative f" "associative g" "\<forall>a b c d. g(f a b) (f c d) = f (g a c) (g b d)" |
27484 | 862 |
shows |
80736 | 863 |
"foldseq g (\<lambda>j. foldseq f (A j) n) m = foldseq f (\<lambda>j. foldseq g ((transpose_infmatrix A) j) m) n" |
80756 | 864 |
using foldmatrix_transpose[of g f A m n] |
35612 | 865 |
by (simp add: foldmatrix_def foldmatrix_transposed_def foldseq_assoc[THEN sym] assms) |
27484 | 866 |
|
867 |
lemma mult_n_nrows: |
|
80756 | 868 |
assumes "\<forall>a. fmul 0 a = 0" "\<forall>a. fmul a 0 = 0" "fadd 0 0 = 0" |
869 |
shows "nrows (mult_matrix_n n fmul fadd A B) \<le> nrows A" |
|
870 |
unfolding nrows_le mult_matrix_n_def |
|
871 |
apply (subst RepAbs_matrix) |
|
872 |
apply (rule_tac x="nrows A" in exI) |
|
873 |
apply (simp add: nrows assms foldseq_zero) |
|
874 |
apply (rule_tac x="ncols B" in exI) |
|
875 |
apply (simp add: ncols assms foldseq_zero) |
|
876 |
apply (simp add: nrows assms foldseq_zero) |
|
877 |
done |
|
27484 | 878 |
|
879 |
lemma mult_n_ncols: |
|
80756 | 880 |
assumes "\<forall>a. fmul 0 a = 0" "\<forall>a. fmul a 0 = 0" "fadd 0 0 = 0" |
881 |
shows "ncols (mult_matrix_n n fmul fadd A B) \<le> ncols B" |
|
882 |
unfolding ncols_le mult_matrix_n_def |
|
883 |
apply (subst RepAbs_matrix) |
|
884 |
apply (rule_tac x="nrows A" in exI) |
|
885 |
apply (simp add: nrows assms foldseq_zero) |
|
886 |
apply (rule_tac x="ncols B" in exI) |
|
887 |
apply (simp add: ncols assms foldseq_zero) |
|
888 |
apply (simp add: ncols assms foldseq_zero) |
|
889 |
done |
|
27484 | 890 |
|
891 |
lemma mult_nrows: |
|
80756 | 892 |
assumes |
893 |
"\<forall>a. fmul 0 a = 0" |
|
894 |
"\<forall>a. fmul a 0 = 0" |
|
895 |
"fadd 0 0 = 0" |
|
896 |
shows "nrows (mult_matrix fmul fadd A B) \<le> nrows A" |
|
897 |
by (simp add: mult_matrix_def mult_n_nrows assms) |
|
27484 | 898 |
|
899 |
lemma mult_ncols: |
|
80756 | 900 |
assumes |
901 |
"\<forall>a. fmul 0 a = 0" |
|
902 |
"\<forall>a. fmul a 0 = 0" |
|
903 |
"fadd 0 0 = 0" |
|
904 |
shows "ncols (mult_matrix fmul fadd A B) \<le> ncols B" |
|
905 |
by (simp add: mult_matrix_def mult_n_ncols assms) |
|
27484 | 906 |
|
80736 | 907 |
lemma nrows_move_matrix_le: "nrows (move_matrix A j i) \<le> nat((int (nrows A)) + j)" |
80756 | 908 |
by (smt (verit) Rep_move_matrix int_nat_eq nrows nrows_le of_nat_le_iff) |
27484 | 909 |
|
80736 | 910 |
lemma ncols_move_matrix_le: "ncols (move_matrix A j i) \<le> nat((int (ncols A)) + i)" |
80756 | 911 |
by (metis nrows_move_matrix_le nrows_transpose transpose_move_matrix) |
27484 | 912 |
|
913 |
lemma mult_matrix_assoc: |
|
35612 | 914 |
assumes |
67613 | 915 |
"\<forall>a. fmul1 0 a = 0" |
916 |
"\<forall>a. fmul1 a 0 = 0" |
|
917 |
"\<forall>a. fmul2 0 a = 0" |
|
918 |
"\<forall>a. fmul2 a 0 = 0" |
|
27484 | 919 |
"fadd1 0 0 = 0" |
920 |
"fadd2 0 0 = 0" |
|
67613 | 921 |
"\<forall>a b c d. fadd2 (fadd1 a b) (fadd1 c d) = fadd1 (fadd2 a c) (fadd2 b d)" |
27484 | 922 |
"associative fadd1" |
923 |
"associative fadd2" |
|
67613 | 924 |
"\<forall>a b c. fmul2 (fmul1 a b) c = fmul1 a (fmul2 b c)" |
925 |
"\<forall>a b c. fmul2 (fadd1 a b) c = fadd1 (fmul2 a c) (fmul2 b c)" |
|
926 |
"\<forall>a b c. fmul1 c (fadd2 a b) = fadd2 (fmul1 c a) (fmul1 c b)" |
|
35612 | 927 |
shows "mult_matrix fmul2 fadd2 (mult_matrix fmul1 fadd1 A B) C = mult_matrix fmul1 fadd1 A (mult_matrix fmul2 fadd2 B C)" |
27484 | 928 |
proof - |
929 |
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 |
|
80736 | 930 |
have fmul2fadd1fold: "!! x s n. fmul2 (foldseq fadd1 s n) x = foldseq fadd1 (\<lambda>k. fmul2 (s k) x) n" |
931 |
by (rule_tac g1 = "\<lambda>y. fmul2 y x" in ssubst [OF foldseq_distr_unary], insert assms, simp_all) |
|
932 |
have fmul1fadd2fold: "!! x s n. fmul1 x (foldseq fadd2 s n) = foldseq fadd2 (\<lambda>k. fmul1 x (s k)) n" |
|
933 |
using assms by (rule_tac g1 = "\<lambda>y. fmul1 x y" in ssubst [OF foldseq_distr_unary], simp_all) |
|
27484 | 934 |
let ?N = "max (ncols A) (max (ncols B) (max (nrows B) (nrows C)))" |
35612 | 935 |
show ?thesis |
80736 | 936 |
apply (intro matrix_eqI) |
27484 | 937 |
apply (simp add: mult_matrix_def) |
938 |
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)"]) |
|
80756 | 939 |
apply (simp add: max1 max2 mult_n_ncols mult_n_nrows assms)+ |
35612 | 940 |
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)"]) |
80756 | 941 |
apply (simp add: max1 max2 mult_n_ncols mult_n_nrows assms)+ |
27484 | 942 |
apply (simplesubst mult_matrix_nm[of _ _ _ "?N"]) |
80756 | 943 |
apply (simp add: max1 max2 mult_n_ncols mult_n_nrows assms)+ |
27484 | 944 |
apply (simplesubst mult_matrix_nm[of _ _ _ "?N"]) |
80756 | 945 |
apply (simp add: max1 max2 mult_n_ncols mult_n_nrows assms)+ |
27484 | 946 |
apply (simplesubst mult_matrix_nm[of _ _ _ "?N"]) |
80756 | 947 |
apply (simp add: max1 max2 mult_n_ncols mult_n_nrows assms)+ |
27484 | 948 |
apply (simplesubst mult_matrix_nm[of _ _ _ "?N"]) |
80756 | 949 |
apply (simp add: max1 max2 mult_n_ncols mult_n_nrows assms)+ |
27484 | 950 |
apply (simp add: mult_matrix_n_def) |
951 |
apply (rule comb_left) |
|
952 |
apply ((rule ext)+, simp) |
|
953 |
apply (simplesubst RepAbs_matrix) |
|
80756 | 954 |
apply (rule exI[of _ "nrows B"]) |
955 |
apply (simp add: nrows assms foldseq_zero) |
|
956 |
apply (rule exI[of _ "ncols C"]) |
|
957 |
apply (simp add: assms ncols foldseq_zero) |
|
27484 | 958 |
apply (subst RepAbs_matrix) |
80756 | 959 |
apply (rule exI[of _ "nrows A"]) |
960 |
apply (simp add: nrows assms foldseq_zero) |
|
961 |
apply (rule exI[of _ "ncols B"]) |
|
962 |
apply (simp add: assms ncols foldseq_zero) |
|
35612 | 963 |
apply (simp add: fmul2fadd1fold fmul1fadd2fold assms) |
27484 | 964 |
apply (subst foldseq_foldseq) |
80756 | 965 |
apply (simp add: assms)+ |
35612 | 966 |
apply (simp add: transpose_infmatrix) |
967 |
done |
|
27484 | 968 |
qed |
969 |
||
970 |
lemma mult_matrix_assoc_simple: |
|
35612 | 971 |
assumes |
67613 | 972 |
"\<forall>a. fmul 0 a = 0" |
973 |
"\<forall>a. fmul a 0 = 0" |
|
27484 | 974 |
"associative fadd" |
975 |
"commutative fadd" |
|
976 |
"associative fmul" |
|
977 |
"distributive fmul fadd" |
|
35612 | 978 |
shows "mult_matrix fmul fadd (mult_matrix fmul fadd A B) C = mult_matrix fmul fadd A (mult_matrix fmul fadd B C)" |
80756 | 979 |
by (smt (verit) assms associative_def commutative_def distributive_def l_distributive_def mult_matrix_assoc r_distributive_def) |
27484 | 980 |
|
981 |
lemma transpose_apply_matrix: "f 0 = 0 \<Longrightarrow> transpose_matrix (apply_matrix f A) = apply_matrix f (transpose_matrix A)" |
|
80736 | 982 |
by (simp add: matrix_eqI) |
27484 | 983 |
|
984 |
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)" |
|
80736 | 985 |
by (simp add: matrix_eqI) |
27484 | 986 |
|
987 |
lemma Rep_mult_matrix: |
|
80756 | 988 |
assumes "\<forall>a. fmul 0 a = 0" "\<forall>a. fmul a 0 = 0" "fadd 0 0 = 0" |
27484 | 989 |
shows |
80756 | 990 |
"Rep_matrix(mult_matrix fmul fadd A B) j i = |
80736 | 991 |
foldseq fadd (\<lambda>k. fmul (Rep_matrix A j k) (Rep_matrix B k i)) (max (ncols A) (nrows B))" |
80756 | 992 |
using assms |
993 |
apply (simp add: mult_matrix_def mult_matrix_n_def) |
|
994 |
apply (subst RepAbs_matrix) |
|
995 |
apply (rule exI[of _ "nrows A"], simp add: nrows foldseq_zero) |
|
996 |
apply (rule exI[of _ "ncols B"], simp add: ncols foldseq_zero) |
|
997 |
apply simp |
|
998 |
done |
|
27484 | 999 |
|
1000 |
lemma transpose_mult_matrix: |
|
1001 |
assumes |
|
67613 | 1002 |
"\<forall>a. fmul 0 a = 0" |
1003 |
"\<forall>a. fmul a 0 = 0" |
|
27484 | 1004 |
"fadd 0 0 = 0" |
67613 | 1005 |
"\<forall>x y. fmul y x = fmul x y" |
27484 | 1006 |
shows |
1007 |
"transpose_matrix (mult_matrix fmul fadd A B) = mult_matrix fmul fadd (transpose_matrix B) (transpose_matrix A)" |
|
35612 | 1008 |
using assms |
80736 | 1009 |
by (simp add: matrix_eqI Rep_mult_matrix ac_simps) |
27484 | 1010 |
|
1011 |
lemma column_transpose_matrix: "column_of_matrix (transpose_matrix A) n = transpose_matrix (row_of_matrix A n)" |
|
80736 | 1012 |
by (simp add: matrix_eqI) |
27484 | 1013 |
|
1014 |
lemma take_columns_transpose_matrix: "take_columns (transpose_matrix A) n = transpose_matrix (take_rows A n)" |
|
80736 | 1015 |
by (simp add: matrix_eqI) |
27484 | 1016 |
|
27580 | 1017 |
instantiation matrix :: ("{zero, ord}") ord |
27484 | 1018 |
begin |
1019 |
||
1020 |
definition |
|
1021 |
le_matrix_def: "A \<le> B \<longleftrightarrow> (\<forall>j i. Rep_matrix A j i \<le> Rep_matrix B j i)" |
|
1022 |
||
1023 |
definition |
|
61076 | 1024 |
less_def: "A < (B::'a matrix) \<longleftrightarrow> A \<le> B \<and> \<not> B \<le> A" |
27484 | 1025 |
|
1026 |
instance .. |
|
1027 |
||
1028 |
end |
|
1029 |
||
27580 | 1030 |
instance matrix :: ("{zero, order}") order |
80736 | 1031 |
proof |
1032 |
fix x y z :: "'a matrix" |
|
1033 |
assume "x \<le> y" "y \<le> z" |
|
1034 |
show "x \<le> z" |
|
1035 |
by (meson \<open>x \<le> y\<close> \<open>y \<le> z\<close> le_matrix_def order_trans) |
|
1036 |
next |
|
1037 |
fix x y :: "'a matrix" |
|
1038 |
assume "x \<le> y" "y \<le> x" |
|
1039 |
show "x = y" |
|
1040 |
by (meson \<open>x \<le> y\<close> \<open>y \<le> x\<close> le_matrix_def matrix_eqI order_antisym) |
|
1041 |
qed (auto simp: less_def le_matrix_def) |
|
27484 | 1042 |
|
1043 |
lemma le_apply_matrix: |
|
1044 |
assumes |
|
1045 |
"f 0 = 0" |
|
80736 | 1046 |
"\<forall>x y. x \<le> y \<longrightarrow> f x \<le> f y" |
1047 |
"(a::('a::{ord, zero}) matrix) \<le> b" |
|
1048 |
shows "apply_matrix f a \<le> apply_matrix f b" |
|
35612 | 1049 |
using assms by (simp add: le_matrix_def) |
27484 | 1050 |
|
1051 |
lemma le_combine_matrix: |
|
1052 |
assumes |
|
1053 |
"f 0 0 = 0" |
|
80756 | 1054 |
"\<forall>a b c d. a \<le> b \<and> c \<le> d \<longrightarrow> f a c \<le> f b d" |
80736 | 1055 |
"A \<le> B" |
1056 |
"C \<le> D" |
|
1057 |
shows "combine_matrix f A C \<le> combine_matrix f B D" |
|
35612 | 1058 |
using assms by (simp add: le_matrix_def) |
27484 | 1059 |
|
1060 |
lemma le_left_combine_matrix: |
|
1061 |
assumes |
|
1062 |
"f 0 0 = 0" |
|
80736 | 1063 |
"\<forall>a b c. a \<le> b \<longrightarrow> f c a \<le> f c b" |
1064 |
"A \<le> B" |
|
80756 | 1065 |
shows "combine_matrix f C A \<le> combine_matrix f C B" |
35612 | 1066 |
using assms by (simp add: le_matrix_def) |
27484 | 1067 |
|
1068 |
lemma le_right_combine_matrix: |
|
1069 |
assumes |
|
1070 |
"f 0 0 = 0" |
|
80736 | 1071 |
"\<forall>a b c. a \<le> b \<longrightarrow> f a c \<le> f b c" |
1072 |
"A \<le> B" |
|
80756 | 1073 |
shows "combine_matrix f A C \<le> combine_matrix f B C" |
35612 | 1074 |
using assms by (simp add: le_matrix_def) |
27484 | 1075 |
|
80736 | 1076 |
lemma le_transpose_matrix: "(A \<le> B) = (transpose_matrix A \<le> transpose_matrix B)" |
27484 | 1077 |
by (simp add: le_matrix_def, auto) |
1078 |
||
1079 |
lemma le_foldseq: |
|
1080 |
assumes |
|
80756 | 1081 |
"\<forall>a b c d . a \<le> b \<and> c \<le> d \<longrightarrow> f a c \<le> f b d" |
80736 | 1082 |
"\<forall>i. i \<le> n \<longrightarrow> s i \<le> t i" |
80756 | 1083 |
shows "foldseq f s n \<le> foldseq f t n" |
27484 | 1084 |
proof - |
80736 | 1085 |
have "\<forall>s t. (\<forall>i. i<=n \<longrightarrow> s i \<le> t i) \<longrightarrow> foldseq f s n \<le> foldseq f t n" |
35612 | 1086 |
by (induct n) (simp_all add: assms) |
80736 | 1087 |
then show "foldseq f s n \<le> foldseq f t n" using assms by simp |
27484 | 1088 |
qed |
1089 |
||
1090 |
lemma le_left_mult: |
|
1091 |
assumes |
|
80756 | 1092 |
"\<forall>a b c d. a \<le> b \<and> c \<le> d \<longrightarrow> fadd a c \<le> fadd b d" |
1093 |
"\<forall>c a b. 0 \<le> c \<and> a \<le> b \<longrightarrow> fmul c a \<le> fmul c b" |
|
67613 | 1094 |
"\<forall>a. fmul 0 a = 0" |
1095 |
"\<forall>a. fmul a 0 = 0" |
|
27484 | 1096 |
"fadd 0 0 = 0" |
80736 | 1097 |
"0 \<le> C" |
1098 |
"A \<le> B" |
|
80756 | 1099 |
shows "mult_matrix fmul fadd C A \<le> mult_matrix fmul fadd C B" |
35612 | 1100 |
using assms |
80756 | 1101 |
apply (auto simp: le_matrix_def Rep_mult_matrix) |
27484 | 1102 |
apply (simplesubst foldseq_zerotail[of _ _ _ "max (ncols C) (max (nrows A) (nrows B))"], simp_all add: nrows ncols max1 max2)+ |
1103 |
apply (rule le_foldseq) |
|
35612 | 1104 |
apply (auto) |
1105 |
done |
|
27484 | 1106 |
|
1107 |
lemma le_right_mult: |
|
1108 |
assumes |
|
80756 | 1109 |
"\<forall>a b c d. a \<le> b \<and> c \<le> d \<longrightarrow> fadd a c \<le> fadd b d" |
1110 |
"\<forall>c a b. 0 \<le> c \<and> a \<le> b \<longrightarrow> fmul a c \<le> fmul b c" |
|
67613 | 1111 |
"\<forall>a. fmul 0 a = 0" |
1112 |
"\<forall>a. fmul a 0 = 0" |
|
27484 | 1113 |
"fadd 0 0 = 0" |
80736 | 1114 |
"0 \<le> C" |
1115 |
"A \<le> B" |
|
80756 | 1116 |
shows "mult_matrix fmul fadd A C \<le> mult_matrix fmul fadd B C" |
35612 | 1117 |
using assms |
80756 | 1118 |
apply (auto simp: le_matrix_def Rep_mult_matrix) |
27484 | 1119 |
apply (simplesubst foldseq_zerotail[of _ _ _ "max (nrows C) (max (ncols A) (ncols B))"], simp_all add: nrows ncols max1 max2)+ |
1120 |
apply (rule le_foldseq) |
|
35612 | 1121 |
apply (auto) |
1122 |
done |
|
27484 | 1123 |
|
67613 | 1124 |
lemma spec2: "\<forall>j i. P j i \<Longrightarrow> P j i" by blast |
27484 | 1125 |
|
80736 | 1126 |
lemma singleton_matrix_le[simp]: "(singleton_matrix j i a \<le> singleton_matrix j i b) = (a \<le> (b::_::order))" |
1127 |
by (auto simp: le_matrix_def) |
|
27484 | 1128 |
|
80736 | 1129 |
lemma singleton_le_zero[simp]: "(singleton_matrix j i x \<le> 0) = (x \<le> (0::'a::{order,zero}))" |
80756 | 1130 |
by (metis singleton_matrix_le singleton_matrix_zero) |
27484 | 1131 |
|
80736 | 1132 |
lemma singleton_ge_zero[simp]: "(0 \<le> singleton_matrix j i x) = ((0::'a::{order,zero}) \<le> x)" |
80756 | 1133 |
by (metis singleton_matrix_le singleton_matrix_zero) |
1134 |
||
1135 |
lemma move_matrix_le_zero[simp]: |
|
1136 |
fixes A:: "'a::{order,zero} matrix" |
|
1137 |
assumes "0 \<le> j" "0 \<le> i" |
|
1138 |
shows "(move_matrix A j i \<le> 0) = (A \<le> 0)" |
|
1139 |
proof - |
|
1140 |
have "Rep_matrix A j' i' \<le> 0" |
|
1141 |
if "\<forall>n m. \<not> int n < j \<and> \<not> int m < i \<longrightarrow> Rep_matrix A (nat (int n - j)) (nat (int m - i)) \<le> 0" |
|
1142 |
for j' i' |
|
1143 |
using that[rule_format, of "j' + nat j" "i' + nat i"] by (simp add: assms) |
|
1144 |
then show ?thesis |
|
1145 |
by (auto simp: le_matrix_def) |
|
1146 |
qed |
|
27484 | 1147 |
|
80756 | 1148 |
lemma move_matrix_zero_le[simp]: |
1149 |
fixes A:: "'a::{order,zero} matrix" |
|
1150 |
assumes "0 \<le> j" "0 \<le> i" |
|
1151 |
shows "(0 \<le> move_matrix A j i) = (0 \<le> A)" |
|
1152 |
proof - |
|
1153 |
have "0 \<le> Rep_matrix A j' i'" |
|
1154 |
if "\<forall>n m. \<not> int n < j \<and> \<not> int m < i \<longrightarrow> 0 \<le> Rep_matrix A (nat (int n - j)) (nat (int m - i))" |
|
1155 |
for j' i' |
|
1156 |
using that[rule_format, of "j' + nat j" "i' + nat i"] by (simp add: assms) |
|
1157 |
then show ?thesis |
|
1158 |
by (auto simp: le_matrix_def) |
|
1159 |
qed |
|
27484 | 1160 |
|
80756 | 1161 |
lemma move_matrix_le_move_matrix_iff[simp]: |
1162 |
fixes A:: "'a::{order,zero} matrix" |
|
1163 |
assumes "0 \<le> j" "0 \<le> i" |
|
1164 |
shows "(move_matrix A j i \<le> move_matrix B j i) = (A \<le> B)" |
|
1165 |
proof - |
|
1166 |
have "Rep_matrix A j' i' \<le> Rep_matrix B j' i'" |
|
1167 |
if "\<forall>n m. \<not> int n < j \<and> \<not> int m < i \<longrightarrow> Rep_matrix A (nat (int n - j)) (nat (int m - i)) \<le> Rep_matrix B (nat (int n - j)) (nat (int m - i))" |
|
1168 |
for j' i' |
|
1169 |
using that[rule_format, of "j' + nat j" "i' + nat i"] by (simp add: assms) |
|
1170 |
then show ?thesis |
|
1171 |
by (auto simp: le_matrix_def) |
|
1172 |
qed |
|
27484 | 1173 |
|
27580 | 1174 |
instantiation matrix :: ("{lattice, zero}") lattice |
25764 | 1175 |
begin |
1176 |
||
37765 | 1177 |
definition "inf = combine_matrix inf" |
25764 | 1178 |
|
37765 | 1179 |
definition "sup = combine_matrix sup" |
25764 | 1180 |
|
1181 |
instance |
|
80736 | 1182 |
by standard (auto simp: 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:
22422
diff
changeset
|
1183 |
|
25764 | 1184 |
end |
1185 |
||
1186 |
instantiation matrix :: ("{plus, zero}") plus |
|
1187 |
begin |
|
1188 |
||
1189 |
definition |
|
67399 | 1190 |
plus_matrix_def: "A + B = combine_matrix (+) A B" |
25764 | 1191 |
|
1192 |
instance .. |
|
1193 |
||
1194 |
end |
|
1195 |
||
1196 |
instantiation matrix :: ("{uminus, zero}") uminus |
|
1197 |
begin |
|
1198 |
||
1199 |
definition |
|
37765 | 1200 |
minus_matrix_def: "- A = apply_matrix uminus A" |
25764 | 1201 |
|
1202 |
instance .. |
|
1203 |
||
1204 |
end |
|
1205 |
||
1206 |
instantiation matrix :: ("{minus, zero}") minus |
|
1207 |
begin |
|
14593 | 1208 |
|
25764 | 1209 |
definition |
67399 | 1210 |
diff_matrix_def: "A - B = combine_matrix (-) A B" |
25764 | 1211 |
|
1212 |
instance .. |
|
1213 |
||
1214 |
end |
|
1215 |
||
1216 |
instantiation matrix :: ("{plus, times, zero}") times |
|
1217 |
begin |
|
1218 |
||
1219 |
definition |
|
69064
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents:
67613
diff
changeset
|
1220 |
times_matrix_def: "A * B = mult_matrix ((*)) (+) A B" |
14940 | 1221 |
|
25764 | 1222 |
instance .. |
1223 |
||
1224 |
end |
|
1225 |
||
27653 | 1226 |
instantiation matrix :: ("{lattice, uminus, zero}") abs |
25764 | 1227 |
begin |
14940 | 1228 |
|
25764 | 1229 |
definition |
61945 | 1230 |
abs_matrix_def: "\<bar>A :: 'a matrix\<bar> = sup A (- A)" |
25764 | 1231 |
|
1232 |
instance .. |
|
1233 |
||
1234 |
end |
|
23879 | 1235 |
|
27653 | 1236 |
instance matrix :: (monoid_add) monoid_add |
1237 |
proof |
|
1238 |
fix A B C :: "'a matrix" |
|
80756 | 1239 |
show "A + B + C = A + (B + C)" |
1240 |
by (simp add: add.assoc matrix_eqI plus_matrix_def) |
|
27653 | 1241 |
show "0 + A = A" |
80756 | 1242 |
by (simp add: matrix_eqI plus_matrix_def) |
27653 | 1243 |
show "A + 0 = A" |
80756 | 1244 |
by (simp add: matrix_eqI plus_matrix_def) |
27653 | 1245 |
qed |
1246 |
||
1247 |
instance matrix :: (comm_monoid_add) comm_monoid_add |
|
1248 |
proof |
|
1249 |
fix A B :: "'a matrix" |
|
14940 | 1250 |
show "A + B = B + A" |
80756 | 1251 |
by (simp add: add.commute matrix_eqI plus_matrix_def) |
14940 | 1252 |
show "0 + A = A" |
80756 | 1253 |
by (simp add: plus_matrix_def matrix_eqI) |
27653 | 1254 |
qed |
1255 |
||
1256 |
instance matrix :: (group_add) group_add |
|
1257 |
proof |
|
1258 |
fix A B :: "'a matrix" |
|
1259 |
show "- A + A = 0" |
|
80736 | 1260 |
by (simp add: plus_matrix_def minus_matrix_def matrix_eqI) |
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
50027
diff
changeset
|
1261 |
show "A + - B = A - B" |
80736 | 1262 |
by (simp add: plus_matrix_def diff_matrix_def minus_matrix_def matrix_eqI) |
27653 | 1263 |
qed |
1264 |
||
1265 |
instance matrix :: (ab_group_add) ab_group_add |
|
1266 |
proof |
|
1267 |
fix A B :: "'a matrix" |
|
14940 | 1268 |
show "- A + A = 0" |
80736 | 1269 |
by (simp add: plus_matrix_def minus_matrix_def matrix_eqI) |
14940 | 1270 |
show "A - B = A + - B" |
80736 | 1271 |
by (simp add: plus_matrix_def diff_matrix_def minus_matrix_def matrix_eqI) |
27653 | 1272 |
qed |
1273 |
||
35028
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents:
34872
diff
changeset
|
1274 |
instance matrix :: (ordered_ab_group_add) ordered_ab_group_add |
27653 | 1275 |
proof |
1276 |
fix A B C :: "'a matrix" |
|
80736 | 1277 |
assume "A \<le> B" |
1278 |
then show "C + A \<le> C + B" |
|
80756 | 1279 |
by (simp add: le_matrix_def plus_matrix_def) |
14940 | 1280 |
qed |
27653 | 1281 |
|
35028
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents:
34872
diff
changeset
|
1282 |
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:
34872
diff
changeset
|
1283 |
instance matrix :: (lattice_ab_group_add) semilattice_sup_ab_group_add .. |
14593 | 1284 |
|
34872 | 1285 |
instance matrix :: (semiring_0) semiring_0 |
14940 | 1286 |
proof |
27653 | 1287 |
fix A B C :: "'a matrix" |
14940 | 1288 |
show "A * B * C = A * (B * C)" |
80756 | 1289 |
unfolding times_matrix_def |
1290 |
by (smt (verit, best) add.assoc associative_def distrib_left distrib_right group_cancel.add2 mult.assoc mult_matrix_assoc mult_not_zero) |
|
14940 | 1291 |
show "(A + B) * C = A * C + B * C" |
80756 | 1292 |
unfolding times_matrix_def plus_matrix_def |
1293 |
using l_distributive_matrix |
|
1294 |
by (metis (full_types) add.assoc add.commute associative_def commutative_def distrib_right l_distributive_def mult_not_zero) |
|
14940 | 1295 |
show "A * (B + C) = A * B + A * C" |
80756 | 1296 |
unfolding times_matrix_def plus_matrix_def |
1297 |
using r_distributive_matrix |
|
1298 |
by (metis (no_types, lifting) add.assoc add.commute associative_def commutative_def distrib_left mult_zero_left mult_zero_right r_distributive_def) |
|
1299 |
qed (auto simp: times_matrix_def) |
|
34872 | 1300 |
|
1301 |
instance matrix :: (ring) ring .. |
|
27653 | 1302 |
|
35028
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents:
34872
diff
changeset
|
1303 |
instance matrix :: (ordered_ring) ordered_ring |
27653 | 1304 |
proof |
1305 |
fix A B C :: "'a matrix" |
|
80756 | 1306 |
assume \<section>: "A \<le> B" "0 \<le> C" |
1307 |
from \<section> show "C * A \<le> C * B" |
|
1308 |
by (simp add: times_matrix_def add_mono le_left_mult mult_left_mono) |
|
1309 |
from \<section> show "A * C \<le> B * C" |
|
1310 |
by (simp add: times_matrix_def add_mono le_right_mult mult_right_mono) |
|
27653 | 1311 |
qed |
1312 |
||
35028
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents:
34872
diff
changeset
|
1313 |
instance matrix :: (lattice_ring) lattice_ring |
27653 | 1314 |
proof |
35028
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents:
34872
diff
changeset
|
1315 |
fix A B C :: "('a :: lattice_ring) matrix" |
61945 | 1316 |
show "\<bar>A\<bar> = sup A (-A)" |
27653 | 1317 |
by (simp add: abs_matrix_def) |
1318 |
qed |
|
14593 | 1319 |
|
80736 | 1320 |
instance matrix :: (lattice_ab_group_add_abs) lattice_ab_group_add_abs |
1321 |
proof |
|
1322 |
show "\<And>a:: 'a matrix. \<bar>a\<bar> = sup a (- a)" |
|
1323 |
by (simp add: abs_matrix_def) |
|
1324 |
qed |
|
1325 |
||
25303
0699e20feabd
renamed lordered_*_* to lordered_*_add_*; further localization
haftmann
parents:
23879
diff
changeset
|
1326 |
lemma Rep_matrix_add[simp]: |
27653 | 1327 |
"Rep_matrix ((a::('a::monoid_add)matrix)+b) j i = (Rep_matrix a j i) + (Rep_matrix b j i)" |
1328 |
by (simp add: plus_matrix_def) |
|
14593 | 1329 |
|
34872 | 1330 |
lemma Rep_matrix_mult: "Rep_matrix ((a::('a::semiring_0) matrix) * b) j i = |
80736 | 1331 |
foldseq (+) (\<lambda>k. (Rep_matrix a j k) * (Rep_matrix b k i)) (max (ncols a) (nrows b))" |
80756 | 1332 |
by (simp add: times_matrix_def Rep_mult_matrix) |
14593 | 1333 |
|
67613 | 1334 |
lemma apply_matrix_add: "\<forall>x y. f (x+y) = (f x) + (f y) \<Longrightarrow> f 0 = (0::'a) |
27653 | 1335 |
\<Longrightarrow> apply_matrix f ((a::('a::monoid_add) matrix) + b) = (apply_matrix f a) + (apply_matrix f b)" |
80736 | 1336 |
by (simp add: matrix_eqI) |
14593 | 1337 |
|
27653 | 1338 |
lemma singleton_matrix_add: "singleton_matrix j i ((a::_::monoid_add)+b) = (singleton_matrix j i a) + (singleton_matrix j i b)" |
80736 | 1339 |
by (simp add: matrix_eqI) |
14593 | 1340 |
|
80736 | 1341 |
lemma nrows_mult: "nrows ((A::('a::semiring_0) matrix) * B) \<le> nrows A" |
80756 | 1342 |
by (simp add: times_matrix_def mult_nrows) |
14593 | 1343 |
|
80736 | 1344 |
lemma ncols_mult: "ncols ((A::('a::semiring_0) matrix) * B) \<le> ncols B" |
80756 | 1345 |
by (simp add: times_matrix_def mult_ncols) |
14593 | 1346 |
|
22422
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
21312
diff
changeset
|
1347 |
definition |
ee19cdb07528
stepping towards uniform lattice theory development in HOL
haftmann
parents:
21312
diff
changeset
|
1348 |
one_matrix :: "nat \<Rightarrow> ('a::{zero,one}) matrix" where |
80756 | 1349 |
"one_matrix n = Abs_matrix (\<lambda>j i. if j = i \<and> j < n then 1 else 0)" |
14593 | 1350 |
|
80756 | 1351 |
lemma Rep_one_matrix[simp]: "Rep_matrix (one_matrix n) j i = (if (j = i \<and> j < n) then 1 else 0)" |
1352 |
unfolding one_matrix_def |
|
1353 |
by (smt (verit, del_insts) RepAbs_matrix not_le) |
|
14593 | 1354 |
|
20633 | 1355 |
lemma nrows_one_matrix[simp]: "nrows ((one_matrix n) :: ('a::zero_neq_one)matrix) = n" (is "?r = _") |
14593 | 1356 |
proof - |
80736 | 1357 |
have "?r \<le> n" by (simp add: nrows_le) |
1358 |
moreover have "n \<le> ?r" by (simp add:le_nrows, arith) |
|
14593 | 1359 |
ultimately show "?r = n" by simp |
1360 |
qed |
|
1361 |
||
20633 | 1362 |
lemma ncols_one_matrix[simp]: "ncols ((one_matrix n) :: ('a::zero_neq_one)matrix) = n" (is "?r = _") |
14593 | 1363 |
proof - |
80736 | 1364 |
have "?r \<le> n" by (simp add: ncols_le) |
1365 |
moreover have "n \<le> ?r" by (simp add: le_ncols, arith) |
|
14593 | 1366 |
ultimately show "?r = n" by simp |
1367 |
qed |
|
1368 |
||
80756 | 1369 |
lemma one_matrix_mult_right[simp]: |
1370 |
fixes A :: "('a::semiring_1) matrix" |
|
1371 |
shows "ncols A \<le> n \<Longrightarrow> A * (one_matrix n) = A" |
|
80736 | 1372 |
apply (intro matrix_eqI) |
1373 |
apply (simp add: times_matrix_def Rep_mult_matrix) |
|
1374 |
apply (subst foldseq_almostzero, auto simp: ncols) |
|
1375 |
done |
|
14593 | 1376 |
|
80756 | 1377 |
lemma one_matrix_mult_left[simp]: |
1378 |
fixes A :: "('a::semiring_1) matrix" |
|
1379 |
shows "nrows A \<le> n \<Longrightarrow> (one_matrix n) * A = A" |
|
80736 | 1380 |
apply (intro matrix_eqI) |
1381 |
apply (simp add: times_matrix_def Rep_mult_matrix) |
|
1382 |
apply (subst foldseq_almostzero, auto simp: nrows) |
|
1383 |
done |
|
14593 | 1384 |
|
80756 | 1385 |
lemma transpose_matrix_mult: |
1386 |
fixes A :: "('a::comm_ring) matrix" |
|
1387 |
shows "transpose_matrix (A*B) = (transpose_matrix B) * (transpose_matrix A)" |
|
1388 |
by (simp add: times_matrix_def transpose_mult_matrix mult.commute) |
|
1389 |
||
1390 |
lemma transpose_matrix_add: |
|
1391 |
fixes A :: "('a::monoid_add) matrix" |
|
1392 |
shows "transpose_matrix (A+B) = transpose_matrix A + transpose_matrix B" |
|
1393 |
by (simp add: plus_matrix_def transpose_combine_matrix) |
|
14940 | 1394 |
|
80756 | 1395 |
lemma transpose_matrix_diff: |
1396 |
fixes A :: "('a::group_add) matrix" |
|
1397 |
shows "transpose_matrix (A-B) = transpose_matrix A - transpose_matrix B" |
|
1398 |
by (simp add: diff_matrix_def transpose_combine_matrix) |
|
14940 | 1399 |
|
80756 | 1400 |
lemma transpose_matrix_minus: |
1401 |
fixes A :: "('a::group_add) matrix" |
|
1402 |
shows "transpose_matrix (-A) = - transpose_matrix (A::'a matrix)" |
|
1403 |
by (simp add: minus_matrix_def transpose_apply_matrix) |
|
14940 | 1404 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35032
diff
changeset
|
1405 |
definition right_inverse_matrix :: "('a::{ring_1}) matrix \<Rightarrow> 'a matrix \<Rightarrow> bool" where |
14940 | 1406 |
"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:
35032
diff
changeset
|
1407 |
|
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35032
diff
changeset
|
1408 |
definition left_inverse_matrix :: "('a::{ring_1}) matrix \<Rightarrow> 'a matrix \<Rightarrow> bool" where |
14940 | 1409 |
"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:
35032
diff
changeset
|
1410 |
|
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35032
diff
changeset
|
1411 |
definition inverse_matrix :: "('a::{ring_1}) matrix \<Rightarrow> 'a matrix \<Rightarrow> bool" where |
14940 | 1412 |
"inverse_matrix A X == (right_inverse_matrix A X) \<and> (left_inverse_matrix A X)" |
14593 | 1413 |
|
1414 |
lemma right_inverse_matrix_dim: "right_inverse_matrix A X \<Longrightarrow> nrows A = ncols X" |
|
80756 | 1415 |
using ncols_mult[of A X] nrows_mult[of A X] |
1416 |
by (simp add: right_inverse_matrix_def) |
|
14593 | 1417 |
|
14940 | 1418 |
lemma left_inverse_matrix_dim: "left_inverse_matrix A Y \<Longrightarrow> ncols A = nrows Y" |
80756 | 1419 |
using ncols_mult[of Y A] nrows_mult[of Y A] |
1420 |
by (simp add: left_inverse_matrix_def) |
|
14940 | 1421 |
|
1422 |
lemma left_right_inverse_matrix_unique: |
|
1423 |
assumes "left_inverse_matrix A Y" "right_inverse_matrix A X" |
|
1424 |
shows "X = Y" |
|
1425 |
proof - |
|
80756 | 1426 |
have "Y = Y * one_matrix (nrows A)" |
1427 |
by (metis assms(1) left_inverse_matrix_def one_matrix_mult_right) |
|
1428 |
also have "\<dots> = Y * (A * X)" |
|
1429 |
by (metis assms(2) max.idem right_inverse_matrix_def right_inverse_matrix_dim) |
|
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
54864
diff
changeset
|
1430 |
also have "\<dots> = (Y * A) * X" by (simp add: mult.assoc) |
80756 | 1431 |
also have "\<dots> = X" |
1432 |
using assms left_inverse_matrix_def right_inverse_matrix_def |
|
1433 |
by (metis left_inverse_matrix_dim max.idem one_matrix_mult_left) |
|
14940 | 1434 |
ultimately show "X = Y" by (simp) |
1435 |
qed |
|
1436 |
||
1437 |
lemma inverse_matrix_inject: "\<lbrakk> inverse_matrix A X; inverse_matrix A Y \<rbrakk> \<Longrightarrow> X = Y" |
|
80736 | 1438 |
by (auto simp: inverse_matrix_def left_right_inverse_matrix_unique) |
14940 | 1439 |
|
1440 |
lemma one_matrix_inverse: "inverse_matrix (one_matrix n) (one_matrix n)" |
|
1441 |
by (simp add: inverse_matrix_def left_inverse_matrix_def right_inverse_matrix_def) |
|
1442 |
||
34872 | 1443 |
lemma zero_imp_mult_zero: "(a::'a::semiring_0) = 0 | b = 0 \<Longrightarrow> a * b = 0" |
80756 | 1444 |
by auto |
14940 | 1445 |
|
1446 |
lemma Rep_matrix_zero_imp_mult_zero: |
|
67613 | 1447 |
"\<forall>j i k. (Rep_matrix A j k = 0) | (Rep_matrix B k i) = 0 \<Longrightarrow> A * B = (0::('a::lattice_ring) matrix)" |
80736 | 1448 |
by (simp add: matrix_eqI Rep_matrix_mult foldseq_zero zero_imp_mult_zero) |
14940 | 1449 |
|
80736 | 1450 |
lemma add_nrows: "nrows (A::('a::monoid_add) matrix) \<le> u \<Longrightarrow> nrows B \<le> u \<Longrightarrow> nrows (A + B) \<le> u" |
80756 | 1451 |
by (simp add: nrows_le) |
14940 | 1452 |
|
80756 | 1453 |
lemma move_matrix_row_mult: |
1454 |
fixes A :: "('a::semiring_0) matrix" |
|
1455 |
shows "move_matrix (A * B) j 0 = (move_matrix A j 0) * B" |
|
80736 | 1456 |
proof - |
1457 |
have "\<And>m. \<not> int m < j \<Longrightarrow> ncols (move_matrix A j 0) \<le> max (ncols A) (nrows B)" |
|
1458 |
by (smt (verit, best) max1 nat_int ncols_move_matrix_le) |
|
1459 |
then show ?thesis |
|
1460 |
apply (intro matrix_eqI) |
|
1461 |
apply (auto simp: Rep_matrix_mult foldseq_zero) |
|
1462 |
apply (rule_tac foldseq_zerotail[symmetric]) |
|
1463 |
apply (auto simp: nrows zero_imp_mult_zero max2) |
|
1464 |
done |
|
1465 |
qed |
|
14940 | 1466 |
|
80756 | 1467 |
lemma move_matrix_col_mult: |
1468 |
fixes A :: "('a::semiring_0) matrix" |
|
1469 |
shows "move_matrix (A * B) 0 i = A * (move_matrix B 0 i)" |
|
80736 | 1470 |
proof - |
1471 |
have "\<And>n. \<not> int n < i \<Longrightarrow> nrows (move_matrix B 0 i) \<le> max (ncols A) (nrows B)" |
|
1472 |
by (smt (verit, del_insts) max2 nat_int nrows_move_matrix_le) |
|
1473 |
then show ?thesis |
|
1474 |
apply (intro matrix_eqI) |
|
1475 |
apply (auto simp: Rep_matrix_mult foldseq_zero) |
|
1476 |
apply (rule_tac foldseq_zerotail[symmetric]) |
|
1477 |
apply (auto simp: ncols zero_imp_mult_zero max1) |
|
1478 |
done |
|
1479 |
qed |
|
14940 | 1480 |
|
27653 | 1481 |
lemma move_matrix_add: "((move_matrix (A + B) j i)::(('a::monoid_add) matrix)) = (move_matrix A j i) + (move_matrix B j i)" |
80736 | 1482 |
by (simp add: matrix_eqI) |
14940 | 1483 |
|
34872 | 1484 |
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 | 1485 |
by (simp add: move_matrix_ortho[of "A*B"] move_matrix_col_mult move_matrix_row_mult) |
1486 |
||
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35032
diff
changeset
|
1487 |
definition scalar_mult :: "('a::ring) \<Rightarrow> 'a matrix \<Rightarrow> 'a matrix" where |
69064
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents:
67613
diff
changeset
|
1488 |
"scalar_mult a m == apply_matrix ((*) a) m" |
14940 | 1489 |
|
1490 |
lemma scalar_mult_zero[simp]: "scalar_mult y 0 = 0" |
|
80736 | 1491 |
by (simp add: scalar_mult_def) |
14940 | 1492 |
|
1493 |
lemma scalar_mult_add: "scalar_mult y (a+b) = (scalar_mult y a) + (scalar_mult y b)" |
|
80736 | 1494 |
by (simp add: scalar_mult_def apply_matrix_add algebra_simps) |
14940 | 1495 |
|
1496 |
lemma Rep_scalar_mult[simp]: "Rep_matrix (scalar_mult y a) j i = y * (Rep_matrix a j i)" |
|
80736 | 1497 |
by (simp add: scalar_mult_def) |
14940 | 1498 |
|
1499 |
lemma scalar_mult_singleton[simp]: "scalar_mult y (singleton_matrix j i x) = singleton_matrix j i (y * x)" |
|
80736 | 1500 |
by (simp add: scalar_mult_def) |
14940 | 1501 |
|
27653 | 1502 |
lemma Rep_minus[simp]: "Rep_matrix (-(A::_::group_add)) x y = - (Rep_matrix A x y)" |
80736 | 1503 |
by (simp add: minus_matrix_def) |
14940 | 1504 |
|
61945 | 1505 |
lemma Rep_abs[simp]: "Rep_matrix \<bar>A::_::lattice_ab_group_add\<bar> x y = \<bar>Rep_matrix A x y\<bar>" |
80736 | 1506 |
by (simp add: abs_lattice sup_matrix_def) |
14940 | 1507 |
|
14593 | 1508 |
end |