Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven...

80
Matrix Steven Obua June 9, 2019 theory Matrix imports Main HOL-Library .Lattice-Algebras begin type-synonym 0 a infmatrix = nat nat 0 a definition nonzero-positions :: (nat nat 0 a ::zero ) (nat × nat ) set where nonzero-positions A = {pos . A (fst pos )(snd pos ) = 0 } definition matrix = {(f ::(nat nat 0 a ::zero )). finite (nonzero-positions f )} typedef (overloaded) 0 a matrix = matrix :: (nat nat 0 a ::zero ) set unfolding matrix-def proof show (λji . 0 ) ∈{(f ::(nat nat 0 a ::zero )). finite (nonzero-positions f )} by (simp add : nonzero-positions-def ) qed declare Rep-matrix-inverse [simp ] lemma finite-nonzero-positions : finite (nonzero-positions (Rep-matrix A)) by (induct A)(simp add : Abs-matrix-inverse matrix-def ) definition nrows :: ( 0 a ::zero ) matrix nat where nrows A == if nonzero-positions (Rep-matrix A)= {} then 0 else Suc (Max ((image fst )(nonzero-positions (Rep-matrix A)))) definition ncols :: ( 0 a ::zero ) matrix nat where ncols A == if nonzero-positions (Rep-matrix A)= {} then 0 else Suc (Max ((image snd )(nonzero-positions (Rep-matrix A)))) lemma nrows : assumes hyp : nrows A m shows (Rep-matrix A m n )= 0 proof cases assume nonzero-positions (Rep-matrix A)= {} 1

Transcript of Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven...

Page 1: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

Matrix

Steven Obua

June 9, 2019

theory Matriximports Main HOL−Library .Lattice-Algebrasbegin

type-synonym ′a infmatrix = nat ⇒ nat ⇒ ′a

definition nonzero-positions :: (nat ⇒ nat ⇒ ′a::zero) ⇒ (nat × nat) set wherenonzero-positions A = {pos. A (fst pos) (snd pos) ∼= 0}

definition matrix = {(f ::(nat ⇒ nat ⇒ ′a::zero)). finite (nonzero-positions f )}

typedef (overloaded) ′a matrix = matrix :: (nat ⇒ nat ⇒ ′a::zero) setunfolding matrix-def

proofshow (λj i . 0 ) ∈ {(f ::(nat ⇒ nat ⇒ ′a::zero)). finite (nonzero-positions f )}

by (simp add : nonzero-positions-def )qed

declare Rep-matrix-inverse[simp]

lemma finite-nonzero-positions : finite (nonzero-positions (Rep-matrix A))by (induct A) (simp add : Abs-matrix-inverse matrix-def )

definition nrows :: ( ′a::zero) matrix ⇒ nat wherenrows A == if nonzero-positions(Rep-matrix A) = {} then 0 else Suc(Max

((image fst) (nonzero-positions (Rep-matrix A))))

definition ncols :: ( ′a::zero) matrix ⇒ nat wherencols A == if nonzero-positions(Rep-matrix A) = {} then 0 else Suc(Max ((image

snd) (nonzero-positions (Rep-matrix A))))

lemma nrows:assumes hyp: nrows A ≤ mshows (Rep-matrix A m n) = 0

proof casesassume nonzero-positions(Rep-matrix A) = {}

1

Page 2: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

then show (Rep-matrix A m n) = 0 by (simp add : nonzero-positions-def )next

assume a: nonzero-positions(Rep-matrix A) 6= {}let ?S = fst‘ (nonzero-positions(Rep-matrix A))have c: finite (?S ) by (simp add : finite-nonzero-positions)from hyp have d : Max (?S ) < m by (simp add : a nrows-def )have m /∈ ?S

proof −have m ∈ ?S =⇒ m <= Max (?S ) by (simp add : Max-ge [OF c])moreover from d have ∼(m <= Max ?S ) by (simp)ultimately show m /∈ ?S by (auto)

qedthus Rep-matrix A m n = 0 by (simp add : nonzero-positions-def image-Collect)

qed

definition transpose-infmatrix :: ′a infmatrix ⇒ ′a infmatrix wheretranspose-infmatrix A j i == A i j

definition transpose-matrix :: ( ′a::zero) matrix ⇒ ′a matrix wheretranspose-matrix == Abs-matrix o transpose-infmatrix o Rep-matrix

declare transpose-infmatrix-def [simp]

lemma transpose-infmatrix-twice[simp]: transpose-infmatrix (transpose-infmatrixA) = Aby ((rule ext)+, simp)

lemma transpose-infmatrix : transpose-infmatrix (% j i . P j i) = (% j i . P i j )apply (rule ext)+by simp

lemma transpose-infmatrix-closed [simp]: Rep-matrix (Abs-matrix (transpose-infmatrix(Rep-matrix x ))) = transpose-infmatrix (Rep-matrix x )apply (rule Abs-matrix-inverse)apply (simp add : matrix-def nonzero-positions-def image-def )proof −

let ?A = {pos. Rep-matrix x (snd pos) (fst pos) 6= 0}let ?swap = % pos. (snd pos, fst pos)let ?B = {pos. Rep-matrix x (fst pos) (snd pos) 6= 0}have swap-image: ?swap‘?A = ?B

apply (simp add : image-def )apply (rule set-eqI )apply (simp)proof

fix yassume hyp: ∃ a b. Rep-matrix x b a 6= 0 ∧ y = (b, a)thus Rep-matrix x (fst y) (snd y) 6= 0

proof −from hyp obtain a b where (Rep-matrix x b a 6= 0 & y = (b,a)) by blast

2

Page 3: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

then show Rep-matrix x (fst y) (snd y) 6= 0 by (simp)qed

nextfix yassume hyp: Rep-matrix x (fst y) (snd y) 6= 0show ∃ a b. (Rep-matrix x b a 6= 0 & y = (b,a))

by (rule exI [of - snd y ], rule exI [of - fst y ]) (simp add : hyp)qed

then have finite (?swap‘?A)proof −have finite (nonzero-positions (Rep-matrix x )) by (simp add : finite-nonzero-positions)then have finite ?B by (simp add : nonzero-positions-def )with swap-image show finite (?swap‘?A) by (simp)

qedmoreoverhave inj-on ?swap ?A by (simp add : inj-on-def )ultimately show finite ?Aby (rule finite-imageD [of ?swap ?A])

qed

lemma infmatrixforward : (x :: ′a infmatrix ) = y =⇒ ∀ a b. x a b = y a b by auto

lemma transpose-infmatrix-inject : (transpose-infmatrix A = transpose-infmatrixB) = (A = B)apply (auto)apply (rule ext)+apply (simp add : transpose-infmatrix )apply (drule infmatrixforward)apply (simp)done

lemma transpose-matrix-inject : (transpose-matrix A = transpose-matrix B) = (A= B)apply (simp add : transpose-matrix-def )apply (subst Rep-matrix-inject [THEN sym])+apply (simp only : transpose-infmatrix-closed transpose-infmatrix-inject)done

lemma transpose-matrix [simp]: Rep-matrix (transpose-matrix A) j i = Rep-matrixA i jby (simp add : transpose-matrix-def )

lemma transpose-transpose-id [simp]: transpose-matrix (transpose-matrix A) = Aby (simp add : transpose-matrix-def )

lemma nrows-transpose[simp]: nrows (transpose-matrix A) = ncols Aby (simp add : nrows-def ncols-def nonzero-positions-def transpose-matrix-def image-def )

lemma ncols-transpose[simp]: ncols (transpose-matrix A) = nrows Aby (simp add : nrows-def ncols-def nonzero-positions-def transpose-matrix-def image-def )

3

Page 4: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

lemma ncols: ncols A <= n =⇒ Rep-matrix A m n = 0proof −

assume ncols A <= nthen have nrows (transpose-matrix A) <= n by (simp)then have Rep-matrix (transpose-matrix A) n m = 0 by (rule nrows)thus Rep-matrix A m n = 0 by (simp add : transpose-matrix-def )

qed

lemma ncols-le: (ncols A <= n) = (∀ j i . n <= i −→ (Rep-matrix A j i) = 0 )(is - = ?st)apply (auto)apply (simp add : ncols)proof (simp add : ncols-def , auto)

let ?P = nonzero-positions (Rep-matrix A)let ?p = snd‘?Phave a:finite ?p by (simp add : finite-nonzero-positions)let ?m = Max ?passume ∼(Suc (?m) <= n)then have b:n <= ?m by (simp)fix a bassume (a,b) ∈ ?Pthen have ?p 6= {} by (auto)with a have ?m ∈ ?p by (simp)moreover have ∀ x . (x ∈ ?p −→ (∃ y . (Rep-matrix A y x ) 6= 0 )) by (simp add :

nonzero-positions-def image-def )ultimately have ∃ y . (Rep-matrix A y ?m) 6= 0 by (simp)moreover assume ?stultimately show False using b by (simp)

qed

lemma less-ncols: (n < ncols A) = (∃ j i . n <= i & (Rep-matrix A j i) 6= 0 )proof −

have a: !! (a::nat) b. (a < b) = (∼(b <= a)) by arithshow ?thesis by (simp add : a ncols-le)

qed

lemma le-ncols: (n <= ncols A) = (∀ m. (∀ j i . m <= i −→ (Rep-matrix A j i)= 0 ) −→ n <= m)apply (auto)apply (subgoal-tac ncols A <= m)apply (simp)apply (simp add : ncols-le)apply (drule-tac x=ncols A in spec)by (simp add : ncols)

lemma nrows-le: (nrows A <= n) = (∀ j i . n <= j −→ (Rep-matrix A j i) = 0 )(is ?s)proof −

4

Page 5: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

have (nrows A <= n) = (ncols (transpose-matrix A) <= n) by (simp)also have . . . = (∀ j i . n <= i −→ (Rep-matrix (transpose-matrix A) j i = 0 ))

by (rule ncols-le)also have . . . = (∀ j i . n <= i −→ (Rep-matrix A i j ) = 0 ) by (simp)finally show (nrows A <= n) = (∀ j i . n <= j −→ (Rep-matrix A j i) = 0 ) by

(auto)qed

lemma less-nrows: (m < nrows A) = (∃ j i . m <= j & (Rep-matrix A j i) 6= 0 )proof −

have a: !! (a::nat) b. (a < b) = (∼(b <= a)) by arithshow ?thesis by (simp add : a nrows-le)

qed

lemma le-nrows: (n <= nrows A) = (∀ m. (∀ j i . m <= j −→ (Rep-matrix A ji) = 0 ) −→ n <= m)apply (auto)apply (subgoal-tac nrows A <= m)apply (simp)apply (simp add : nrows-le)apply (drule-tac x=nrows A in spec)by (simp add : nrows)

lemma nrows-notzero: Rep-matrix A m n 6= 0 =⇒ m < nrows Aapply (case-tac nrows A <= m)apply (simp-all add : nrows)done

lemma ncols-notzero: Rep-matrix A m n 6= 0 =⇒ n < ncols Aapply (case-tac ncols A <= n)apply (simp-all add : ncols)done

lemma finite-natarray1 : finite {x . x < (n::nat)}apply (induct n)apply (simp)proof −

fix nhave {x . x < Suc n} = insert n {x . x < n} by (rule set-eqI , simp, arith)moreover assume finite {x . x < n}ultimately show finite {x . x < Suc n} by (simp)

qed

lemma finite-natarray2 : finite {(x , y). x < (m::nat) & y < (n::nat)}by simp

lemma RepAbs-matrix :assumes aem: ∃m. ∀ j i . m <= j −→ x j i = 0 (is ?em) and aen:∃n. ∀ j i . (n

<= i −→ x j i = 0 ) (is ?en)

5

Page 6: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

shows (Rep-matrix (Abs-matrix x )) = xapply (rule Abs-matrix-inverse)apply (simp add : matrix-def nonzero-positions-def )proof −

from aem obtain m where a: ∀ j i . m <= j −→ x j i = 0 by (blast)from aen obtain n where b: ∀ j i . n <= i −→ x j i = 0 by (blast)let ?u = {(i , j ). x i j 6= 0}let ?v = {(i , j ). i < m & j < n}have c: !! (m::nat) a. ∼(m <= a) =⇒ a < m by (arith)from a b have (?u ∩ (−?v)) = {}

apply (simp)apply (rule set-eqI )apply (simp)apply autoby (rule c, auto)+

then have d : ?u ⊆ ?v by blastmoreover have finite ?v by (simp add : finite-natarray2 )moreover have {pos. x (fst pos) (snd pos) 6= 0} = ?u by autoultimately show finite {pos. x (fst pos) (snd pos) 6= 0}

by (metis (lifting) finite-subset)qed

definition apply-infmatrix :: ( ′a ⇒ ′b) ⇒ ′a infmatrix ⇒ ′b infmatrix whereapply-infmatrix f == % A. (% j i . f (A j i))

definition apply-matrix :: ( ′a ⇒ ′b) ⇒ ( ′a::zero) matrix ⇒ ( ′b::zero) matrixwhere

apply-matrix f == % A. Abs-matrix (apply-infmatrix f (Rep-matrix A))

definition combine-infmatrix :: ( ′a ⇒ ′b ⇒ ′c) ⇒ ′a infmatrix ⇒ ′b infmatrix ⇒′c infmatrix where

combine-infmatrix f == % A B . (% j i . f (A j i) (B j i))

definition combine-matrix :: ( ′a ⇒ ′b ⇒ ′c) ⇒ ( ′a::zero) matrix ⇒ ( ′b::zero)matrix ⇒ ( ′c::zero) matrix where

combine-matrix f == % A B . Abs-matrix (combine-infmatrix f (Rep-matrix A)(Rep-matrix B))

lemma expand-apply-infmatrix [simp]: apply-infmatrix f A j i = f (A j i)by (simp add : apply-infmatrix-def )

lemma expand-combine-infmatrix [simp]: combine-infmatrix f A B j i = f (A j i)(B j i)by (simp add : combine-infmatrix-def )

definition commutative :: ( ′a ⇒ ′a ⇒ ′b) ⇒ bool wherecommutative f == ∀ x y . f x y = f y x

definition associative :: ( ′a ⇒ ′a ⇒ ′a) ⇒ bool where

6

Page 7: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

associative f == ∀ x y z . f (f x y) z = f x (f y z )

To reason about associativity and commutativity of operations on matrices,let’s take a step back and look at the general situtation: Assume that wehave sets A and B with B ⊂ A and an abstraction u : A → B. Thisabstraction has to fulfill u(b) = b for all b ∈ B, but is arbitrary otherwise.Each function f : A × A → A now induces a function f ′ : B × B → B byf ′ = u ◦ f . It is obvious that commutativity of f implies commutativity off ′: f ′xy = u(fxy) = u(fyx) = f ′yx.

lemma combine-infmatrix-commute:commutative f =⇒ commutative (combine-infmatrix f )

by (simp add : commutative-def combine-infmatrix-def )

lemma combine-matrix-commute:commutative f =⇒ commutative (combine-matrix f )by (simp add : combine-matrix-def commutative-def combine-infmatrix-def )

On the contrary, given an associative function f we cannot expect f ′ to beassociative. A counterexample is given by A = ZZ, B = {−1, 0, 1}, as f wetake addition on ZZ, which is clearly associative. The abstraction is given byu(a) = 0 for a /∈ B. Then we have

f ′(f ′11)− 1 = u(f(u(f11))− 1) = u(f(u2)− 1) = u(f0− 1) = −1,

but on the other hand we have

f ′1(f ′1− 1) = u(f1(u(f1− 1))) = u(f10) = 1.

A way out of this problem is to assume that f(A× A) ⊂ A holds, and thisis what we are going to do:

lemma nonzero-positions-combine-infmatrix [simp]: f 0 0 = 0 =⇒ nonzero-positions(combine-infmatrix f A B) ⊆ (nonzero-positions A) ∪ (nonzero-positions B)by (rule subsetI , simp add : nonzero-positions-def combine-infmatrix-def , auto)

lemma finite-nonzero-positions-Rep[simp]: finite (nonzero-positions (Rep-matrixA))by (insert Rep-matrix [of A], simp add : matrix-def )

lemma combine-infmatrix-closed [simp]:f 0 0 = 0 =⇒ Rep-matrix (Abs-matrix (combine-infmatrix f (Rep-matrix A)

(Rep-matrix B))) = combine-infmatrix f (Rep-matrix A) (Rep-matrix B)apply (rule Abs-matrix-inverse)apply (simp add : matrix-def )apply (rule finite-subset [of - (nonzero-positions (Rep-matrix A)) ∪ (nonzero-positions(Rep-matrix B))])by (simp-all)

We need the next two lemmas only later, but it is analog to the above one,so we prove them now:

7

Page 8: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

lemma nonzero-positions-apply-infmatrix [simp]: f 0 = 0 =⇒ nonzero-positions(apply-infmatrix f A) ⊆ nonzero-positions Aby (rule subsetI , simp add : nonzero-positions-def apply-infmatrix-def , auto)

lemma apply-infmatrix-closed [simp]:f 0 = 0 =⇒ Rep-matrix (Abs-matrix (apply-infmatrix f (Rep-matrix A))) =

apply-infmatrix f (Rep-matrix A)apply (rule Abs-matrix-inverse)apply (simp add : matrix-def )apply (rule finite-subset [of - nonzero-positions (Rep-matrix A)])by (simp-all)

lemma combine-infmatrix-assoc[simp]: f 0 0 = 0 =⇒ associative f =⇒ associative(combine-infmatrix f )by (simp add : associative-def combine-infmatrix-def )

lemma comb: f = g =⇒ x = y =⇒ f x = g yby (auto)

lemma combine-matrix-assoc: f 0 0 = 0 =⇒ associative f =⇒ associative (combine-matrixf )apply (simp(no-asm) add : associative-def combine-matrix-def , auto)apply (rule comb [of Abs-matrix Abs-matrix ])by (auto, insert combine-infmatrix-assoc[of f ], simp add : associative-def )

lemma Rep-apply-matrix [simp]: f 0 = 0 =⇒ Rep-matrix (apply-matrix f A) j i =f (Rep-matrix A j i)by (simp add : apply-matrix-def )

lemma Rep-combine-matrix [simp]: f 0 0 = 0 =⇒ Rep-matrix (combine-matrix fA B) j i = f (Rep-matrix A j i) (Rep-matrix B j i)

by(simp add : combine-matrix-def )

lemma combine-nrows-max : f 0 0 = 0 =⇒ nrows (combine-matrix f A B) <=max (nrows A) (nrows B)by (simp add : nrows-le)

lemma combine-ncols-max : f 0 0 = 0 =⇒ ncols (combine-matrix f A B) <= max(ncols A) (ncols B)by (simp add : ncols-le)

lemma combine-nrows: f 0 0 = 0 =⇒ nrows A <= q =⇒ nrows B <= q =⇒nrows(combine-matrix f A B) <= q

by (simp add : nrows-le)

lemma combine-ncols: f 0 0 = 0 =⇒ ncols A <= q =⇒ ncols B <= q =⇒ncols(combine-matrix f A B) <= q

by (simp add : ncols-le)

8

Page 9: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

definition zero-r-neutral :: ( ′a ⇒ ′b::zero ⇒ ′a) ⇒ bool wherezero-r-neutral f == ∀ a. f a 0 = a

definition zero-l-neutral :: ( ′a::zero ⇒ ′b ⇒ ′b) ⇒ bool wherezero-l-neutral f == ∀ a. f 0 a = a

definition zero-closed :: (( ′a::zero) ⇒ ( ′b::zero) ⇒ ( ′c::zero)) ⇒ bool wherezero-closed f == (∀ x . f x 0 = 0 ) & (∀ y . f 0 y = 0 )

primrec foldseq :: ( ′a ⇒ ′a ⇒ ′a) ⇒ (nat ⇒ ′a) ⇒ nat ⇒ ′awhere

foldseq f s 0 = s 0| foldseq f s (Suc n) = f (s 0 ) (foldseq f (% k . s(Suc k)) n)

primrec foldseq-transposed :: ( ′a ⇒ ′a ⇒ ′a) ⇒ (nat ⇒ ′a) ⇒ nat ⇒ ′awhere

foldseq-transposed f s 0 = s 0| foldseq-transposed f s (Suc n) = f (foldseq-transposed f s n) (s (Suc n))

lemma foldseq-assoc : associative f =⇒ foldseq f = foldseq-transposed fproof −

assume a:associative fthen have sublemma:

∧n. ∀N s. N <= n −→ foldseq f s N = foldseq-transposed

f s Nproof −

fix nshow ∀N s. N <= n −→ foldseq f s N = foldseq-transposed f s Nproof (induct n)

show ∀N s. N <= 0 −→ foldseq f s N = foldseq-transposed f s N by simpnext

fix nassume b: ∀N s. N <= n −→ foldseq f s N = foldseq-transposed f s Nhave c:

∧N s. N <= n =⇒ foldseq f s N = foldseq-transposed f s N by (simp

add : b)show ∀N t . N <= Suc n −→ foldseq f t N = foldseq-transposed f t Nproof (auto)

fix N tassume Nsuc: N <= Suc nshow foldseq f t N = foldseq-transposed f t Nproof cases

assume N <= nthen show foldseq f t N = foldseq-transposed f t N by (simp add : b)

nextassume ∼(N <= n)with Nsuc have Nsuceq : N = Suc n by simphave neqz : n 6= 0 =⇒ ∃m. n = Suc m & Suc m <= n by arithhave assocf : !! x y z . f x (f y z ) = f (f x y) z by (insert a, simp add :

associative-def )show foldseq f t N = foldseq-transposed f t N

9

Page 10: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

apply (simp add : Nsuceq)apply (subst c)apply (simp)apply (case-tac n = 0 )apply (simp)apply (drule neqz )apply (erule exE )apply (simp)apply (subst assocf )proof −

fix massume n = Suc m & Suc m <= nthen have mless: Suc m <= n by ariththen have step1 : foldseq-transposed f (% k . t (Suc k)) m = foldseq f

(% k . t (Suc k)) m (is ?T1 = ?T2 )apply (subst c)by simp+

have step2 : f (t 0 ) ?T2 = foldseq f t (Suc m) (is - = ?T3 ) by simphave step3 : ?T3 = foldseq-transposed f t (Suc m) (is - = ?T4 )

apply (subst c)by (simp add : mless)+

have step4 : ?T4 = f (foldseq-transposed f t m) (t (Suc m)) (is -=?T5 )by simp

from step1 step2 step3 step4 show sowhat : f (f (t 0 ) ?T1 ) (t (Suc(Suc m))) = f ?T5 (t (Suc (Suc m))) by simp

qedqed

qedqed

qedshow foldseq f = foldseq-transposed f by ((rule ext)+, insert sublemma, auto)

qed

lemma foldseq-distr : [[associative f ; commutative f ]] =⇒ foldseq f (% k . f (u k) (vk)) n = f (foldseq f u n) (foldseq f v n)proof −

assume assoc: associative fassume comm: commutative ffrom assoc have a:!! x y z . f (f x y) z = f x (f y z ) by (simp add : associative-def )from comm have b: !! x y . f x y = f y x by (simp add : commutative-def )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 )have

∧n. (∀ u v . foldseq f (%k . f (u k) (v k)) n = f (foldseq f u n) (foldseq f v

n))apply (induct-tac n)apply (simp+, auto)by (simp add : a b c)

then show foldseq f (% k . f (u k) (v k)) n = f (foldseq f u n) (foldseq f v n) bysimp

10

Page 11: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

qed

theorem [[associative f ; associative g ; ∀ a b c d . g (f a b) (f c d) = f (g a c) (g bd); ∃ x y . (f x ) 6= (f y); ∃ x y . (g x ) 6= (g y); f x x = x ; g x x = x ]] =⇒ f =g | (∀ y .f y x = y) | (∀ y . g y x = y)oops

lemma foldseq-zero:assumes fz : f 0 0 = 0 and sz : ∀ i . i <= n −→ s i = 0shows foldseq f s n = 0proof −

have∧

n. ∀ s. (∀ i . i <= n −→ s i = 0 ) −→ foldseq f s n = 0apply (induct-tac n)apply (simp)by (simp add : fz )

then show foldseq f s n = 0 by (simp add : sz )qed

lemma foldseq-significant-positions:assumes p: ∀ i . i <= N −→ S i = T ishows foldseq f S N = foldseq f T N

proof −have

∧m. ∀ s t . (∀ i . i<=m −→ s i = t i) −→ foldseq f s m = foldseq f t m

apply (induct-tac m)apply (simp)apply (simp)apply (auto)proof −

fix nfix s::nat⇒ ′afix t ::nat⇒ ′aassume a: ∀ s t . (∀ i≤n. s i = t i) −→ foldseq f s n = foldseq f t nassume b: ∀ i≤Suc n. s i = t ihave c:!! a b. a = b =⇒ f (t 0 ) a = f (t 0 ) b by blasthave d :!! s t . (∀ i≤n. s i = t i) =⇒ foldseq f s n = foldseq f t n by (simp

add : a)show f (t 0 ) (foldseq f (λk . s (Suc k)) n) = f (t 0 ) (foldseq f (λk . t (Suc

k)) n) by (rule c, simp add : d b)qed

with p show ?thesis by simpqed

lemma foldseq-tail :assumes M <= Nshows foldseq f S N = foldseq f (% k . (if k < M then (S k) else (foldseq f (%

k . S (k+M )) (N−M )))) Mproof −

have suc:∧

a b. [[a <= Suc b; a 6= Suc b]] =⇒ a <= b by arith

11

Page 12: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

have a:∧

a b c . a = b =⇒ f c a = f c b by blasthave

∧n. ∀m s. m <= n −→ foldseq f s n = foldseq f (% k . (if k < m then (s

k) else (foldseq f (% k . s(k+m)) (n−m)))) mapply (induct-tac n)apply (simp)apply (simp)apply (auto)apply (case-tac m = Suc na)apply (simp)apply (rule a)apply (rule foldseq-significant-positions)apply (auto)apply (drule suc, simp+)proof −

fix na m sassume suba:∀m≤na. ∀ s. foldseq f s na = foldseq f (λk . if k < m then s k

else foldseq f (λk . s (k + m)) (na − m))massume subb:m <= nafrom suba have subc:!! m s. m <= na =⇒foldseq f s na = foldseq f (λk . if

k < m then s k else foldseq f (λk . s (k + m)) (na − m))m by simphave subd : foldseq f (λk . if k < m then s (Suc k) else foldseq f (λk . s (Suc

(k + m))) (na − m)) m =foldseq f (% k . s(Suc k)) naby (rule subc[of m % k . s(Suc k), THEN sym], simp add : subb)

from subb have sube: m 6= 0 =⇒ ∃mm. m = Suc mm & mm <= na byarith

show f (s 0 ) (foldseq f (λk . if k < m then s (Suc k) else foldseq f (λk . s (Suc(k + m))) (na − m)) m) =

foldseq f (λk . if k < m then s k else foldseq f (λk . s (k + m)) (Suc na −m)) m

apply (simp add : subd)apply (cases m = 0 )apply simp

apply (drule sube)apply autoapply (rule a)apply (simp add : subc cong del : if-weak-cong)done

qedthen show ?thesis using assms by simp

qed

lemma foldseq-zerotail :assumesfz : f 0 0 = 0and sz : ∀ i . n <= i −→ s i = 0and nm: n <= mshowsfoldseq f s n = foldseq f s m

12

Page 13: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

proof −show foldseq f s n = foldseq f s m

apply (simp add : foldseq-tail [OF nm, of f s])apply (rule foldseq-significant-positions)apply (auto)apply (subst foldseq-zero)by (simp add : fz sz )+

qed

lemma foldseq-zerotail2 :assumes ∀ x . f x 0 = xand ∀ i . n < i −→ s i = 0and nm: n <= mshows foldseq f s n = foldseq f s m

proof −have f 0 0 = 0 by (simp add : assms)have b:

∧m n. n <= m =⇒ m 6= n =⇒ ∃ k . m−n = Suc k by arith

have c: 0 <= m by simphave d :

∧k . k 6= 0 =⇒ ∃ l . k = Suc l by arith

show ?thesisapply (subst foldseq-tail [OF nm])apply (rule foldseq-significant-positions)apply (auto)apply (case-tac m=n)apply (simp+)apply (drule b[OF nm])apply (auto)apply (case-tac k=0 )apply (simp add : assms)apply (drule d)apply (auto)apply (simp add : assms foldseq-zero)done

qed

lemma foldseq-zerostart :∀ x . f 0 (f 0 x ) = f 0 x =⇒ ∀ i . i <= n −→ s i = 0 =⇒ foldseq f s (Suc n) = f

0 (s (Suc n))proof −

assume f00x : ∀ x . f 0 (f 0 x ) = f 0 xhave ∀ s. (∀ i . i<=n −→ s i = 0 ) −→ foldseq f s (Suc n) = f 0 (s (Suc n))

apply (induct n)apply (simp)apply (rule allI , rule impI )proof −

fix nfix shave a:foldseq f s (Suc (Suc n)) = f (s 0 ) (foldseq f (% k . s(Suc k)) (Suc

n)) by simp

13

Page 14: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

assume b: ∀ s. ((∀ i≤n. s i = 0 ) −→ foldseq f s (Suc n) = f 0 (s (Suc n)))from b have c:!! s. (∀ i≤n. s i = 0 ) =⇒ foldseq f s (Suc n) = f 0 (s (Suc

n)) by simpassume d : ∀ i . i <= Suc n −→ s i = 0show foldseq f s (Suc (Suc n)) = f 0 (s (Suc (Suc n)))

apply (subst a)apply (subst c)by (simp add : d f00x )+

qedthen show ∀ i . i <= n −→ s i = 0 =⇒ foldseq f s (Suc n) = f 0 (s (Suc n))

by simpqed

lemma foldseq-zerostart2 :∀ x . f 0 x = x =⇒ ∀ i . i < n −→ s i = 0 =⇒ foldseq f s n = s n

proof −assume a: ∀ i . i<n −→ s i = 0assume x : ∀ x . f 0 x = xfrom x have f00x : ∀ x . f 0 (f 0 x ) = f 0 x by blasthave b:

∧i l . i < Suc l = (i <= l) by arith

have d :∧

k . k 6= 0 =⇒ ∃ l . k = Suc l by arithshow foldseq f s n = s napply (case-tac n=0 )apply (simp)apply (insert a)apply (drule d)apply (auto)apply (simp add : b)apply (insert f00x )apply (drule foldseq-zerostart)by (simp add : x )+

qed

lemma foldseq-almostzero:assumes f0x : ∀ x . f 0 x = x and fx0 : ∀ x . f x 0 = x and s0 : ∀ i . i 6= j −→ s i

= 0shows foldseq f s n = (if (j <= n) then (s j ) else 0 )

proof −from s0 have a: ∀ i . i < j −→ s i = 0 by simpfrom s0 have b: ∀ i . j < i −→ s i = 0 by simpshow ?thesis

apply autoapply (subst foldseq-zerotail2 [of f , OF fx0 , of j , OF b, of n, THEN sym])apply simpapply (subst foldseq-zerostart2 )apply (simp add : f0x a)+apply (subst foldseq-zero)by (simp add : s0 f0x )+

qed

14

Page 15: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

lemma foldseq-distr-unary :assumes !! a b. g (f a b) = f (g a) (g b)shows g(foldseq f s n) = foldseq f (% x . g(s x )) n

proof −have ∀ s. g(foldseq f s n) = foldseq f (% x . g(s x )) n

apply (induct-tac n)apply (simp)apply (simp)apply (auto)apply (drule-tac x=% k . s (Suc k) in spec)by (simp add : assms)

then show ?thesis by simpqed

definition mult-matrix-n :: nat ⇒ (( ′a::zero) ⇒ ( ′b::zero) ⇒ ( ′c::zero)) ⇒ ( ′c ⇒′c ⇒ ′c) ⇒ ′a matrix ⇒ ′b matrix ⇒ ′c matrix where

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)

definition mult-matrix :: (( ′a::zero) ⇒ ( ′b::zero) ⇒ ( ′c::zero)) ⇒ ( ′c ⇒ ′c ⇒ ′c)⇒ ′a matrix ⇒ ′b matrix ⇒ ′c matrix where

mult-matrix fmul fadd A B == mult-matrix-n (max (ncols A) (nrows B)) fmulfadd A B

lemma mult-matrix-n:assumes ncols A ≤ n (is ?An) nrows B ≤ n (is ?Bn) fadd 0 0 = 0 fmul 0 0

= 0shows c:mult-matrix fmul fadd A B = mult-matrix-n n fmul fadd A B

proof −show ?thesis using assms

apply (simp add : mult-matrix-def mult-matrix-n-def )apply (rule comb[of Abs-matrix Abs-matrix ], simp, (rule ext)+)apply (rule foldseq-zerotail , simp-all add : nrows-le ncols-le assms)done

qed

lemma mult-matrix-nm:assumes ncols A <= n nrows B <= n ncols A <= m nrows B <= m fadd 0 0

= 0 fmul 0 0 = 0shows mult-matrix-n n fmul fadd A B = mult-matrix-n m fmul fadd A B

proof −from assms have mult-matrix-n n fmul fadd A B = mult-matrix fmul fadd A B

by (simp add : mult-matrix-n)also from assms have . . . = mult-matrix-n m fmul fadd A B

by (simp add : mult-matrix-n[THEN sym])finally show mult-matrix-n n fmul fadd A B = mult-matrix-n m fmul fadd A B

by simpqed

15

Page 16: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

definition r-distributive :: ( ′a ⇒ ′b ⇒ ′b) ⇒ ( ′b ⇒ ′b ⇒ ′b) ⇒ bool wherer-distributive fmul fadd == ∀ a u v . fmul a (fadd u v) = fadd (fmul a u) (fmul

a v)

definition l-distributive :: ( ′a ⇒ ′b ⇒ ′a) ⇒ ( ′a ⇒ ′a ⇒ ′a) ⇒ bool wherel-distributive fmul fadd == ∀ a u v . fmul (fadd u v) a = fadd (fmul u a) (fmul v

a)

definition distributive :: ( ′a ⇒ ′a ⇒ ′a) ⇒ ( ′a ⇒ ′a ⇒ ′a) ⇒ bool wheredistributive fmul fadd == l-distributive fmul fadd & r-distributive fmul fadd

lemma max1 : !! a x y . (a::nat) <= x =⇒ a <= max x y by (arith)lemma max2 : !! b x y . (b::nat) <= y =⇒ b <= max x y by (arith)

lemma r-distributive-matrix :assumesr-distributive fmul faddassociative faddcommutative faddfadd 0 0 = 0∀ a. fmul a 0 = 0∀ a. fmul 0 a = 0

shows r-distributive (mult-matrix fmul fadd) (combine-matrix fadd)proof −

from assms show ?thesisapply (simp add : r-distributive-def mult-matrix-def , auto)proof −

fix a:: ′a matrixfix u:: ′b matrixfix v :: ′b matrixlet ?mx = max (ncols a) (max (nrows u) (nrows v))from assms show mult-matrix-n (max (ncols a) (nrows (combine-matrix fadd

u v))) fmul fadd a (combine-matrix fadd u v) =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)apply (subst mult-matrix-nm[of - - - ?mx fadd fmul ])apply (simp add : max1 max2 combine-nrows combine-ncols)+apply (subst mult-matrix-nm[of - - v ?mx fadd fmul ])apply (simp add : max1 max2 combine-nrows combine-ncols)+apply (subst mult-matrix-nm[of - - u ?mx fadd fmul ])apply (simp add : max1 max2 combine-nrows combine-ncols)+apply (simp add : mult-matrix-n-def r-distributive-def foldseq-distr [of fadd ])apply (simp add : combine-matrix-def combine-infmatrix-def )apply (rule comb[of Abs-matrix Abs-matrix ], simp, (rule ext)+)apply (simplesubst RepAbs-matrix )apply (simp, auto)apply (rule exI [of - nrows a], simp add : nrows-le foldseq-zero)apply (rule exI [of - ncols v ], simp add : ncols-le foldseq-zero)

16

Page 17: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

apply (subst RepAbs-matrix )apply (simp, auto)apply (rule exI [of - nrows a], simp add : nrows-le foldseq-zero)apply (rule exI [of - ncols u], simp add : ncols-le foldseq-zero)done

qedqed

lemma l-distributive-matrix :assumesl-distributive fmul faddassociative faddcommutative faddfadd 0 0 = 0∀ a. fmul a 0 = 0∀ a. fmul 0 a = 0

shows l-distributive (mult-matrix fmul fadd) (combine-matrix fadd)proof −

from assms show ?thesisapply (simp add : l-distributive-def mult-matrix-def , auto)proof −

fix a:: ′b matrixfix u:: ′a matrixfix v :: ′a matrixlet ?mx = max (nrows a) (max (ncols u) (ncols v))from assms show mult-matrix-n (max (ncols (combine-matrix fadd u v))

(nrows a)) fmul fadd (combine-matrix fadd u v) a =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)apply (subst mult-matrix-nm[of v - - ?mx fadd fmul ])apply (simp add : max1 max2 combine-nrows combine-ncols)+apply (subst mult-matrix-nm[of u - - ?mx fadd fmul ])apply (simp add : max1 max2 combine-nrows combine-ncols)+apply (subst mult-matrix-nm[of - - - ?mx fadd fmul ])apply (simp add : max1 max2 combine-nrows combine-ncols)+apply (simp add : mult-matrix-n-def l-distributive-def foldseq-distr [of fadd ])apply (simp add : combine-matrix-def combine-infmatrix-def )apply (rule comb[of Abs-matrix Abs-matrix ], simp, (rule ext)+)apply (simplesubst RepAbs-matrix )apply (simp, auto)apply (rule exI [of - nrows v ], simp add : nrows-le foldseq-zero)apply (rule exI [of - ncols a], simp add : ncols-le foldseq-zero)apply (subst RepAbs-matrix )apply (simp, auto)apply (rule exI [of - nrows u], simp add : nrows-le foldseq-zero)apply (rule exI [of - ncols a], simp add : ncols-le foldseq-zero)done

qedqed

17

Page 18: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

instantiation matrix :: (zero) zerobegin

definition zero-matrix-def : 0 = Abs-matrix (λj i . 0 )

instance ..

end

lemma Rep-zero-matrix-def [simp]: Rep-matrix 0 j i = 0apply (simp add : zero-matrix-def )apply (subst RepAbs-matrix )by (auto)

lemma zero-matrix-def-nrows[simp]: nrows 0 = 0proof −

have a:!! (x ::nat). x <= 0 =⇒ x = 0 by (arith)show nrows 0 = 0 by (rule a, subst nrows-le, simp)

qed

lemma zero-matrix-def-ncols[simp]: ncols 0 = 0proof −

have a:!! (x ::nat). x <= 0 =⇒ x = 0 by (arith)show ncols 0 = 0 by (rule a, subst ncols-le, simp)

qed

lemma combine-matrix-zero-l-neutral : zero-l-neutral f =⇒ zero-l-neutral (combine-matrixf )

by (simp add : zero-l-neutral-def combine-matrix-def combine-infmatrix-def )

lemma combine-matrix-zero-r-neutral : zero-r-neutral f =⇒ zero-r-neutral (combine-matrixf )

by (simp add : zero-r-neutral-def combine-matrix-def combine-infmatrix-def )

lemma mult-matrix-zero-closed : [[fadd 0 0 = 0 ; zero-closed fmul ]] =⇒ zero-closed(mult-matrix fmul fadd)

apply (simp add : zero-closed-def mult-matrix-def mult-matrix-n-def )apply (auto)by (subst foldseq-zero, (simp add : zero-matrix-def )+)+

lemma mult-matrix-n-zero-right [simp]: [[fadd 0 0 = 0 ; ∀ a. fmul a 0 = 0 ]] =⇒mult-matrix-n n fmul fadd A 0 = 0

apply (simp add : mult-matrix-n-def )apply (subst foldseq-zero)by (simp-all add : zero-matrix-def )

lemma mult-matrix-n-zero-left [simp]: [[fadd 0 0 = 0 ; ∀ a. fmul 0 a = 0 ]] =⇒mult-matrix-n n fmul fadd 0 A = 0

18

Page 19: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

apply (simp add : mult-matrix-n-def )apply (subst foldseq-zero)by (simp-all add : zero-matrix-def )

lemma mult-matrix-zero-left [simp]: [[fadd 0 0 = 0 ; ∀ a. fmul 0 a = 0 ]] =⇒ mult-matrixfmul fadd 0 A = 0by (simp add : mult-matrix-def )

lemma mult-matrix-zero-right [simp]: [[fadd 0 0 = 0 ; ∀ a. fmul a 0 = 0 ]] =⇒mult-matrix fmul fadd A 0 = 0by (simp add : mult-matrix-def )

lemma apply-matrix-zero[simp]: f 0 = 0 =⇒ apply-matrix f 0 = 0apply (simp add : apply-matrix-def apply-infmatrix-def )by (simp add : zero-matrix-def )

lemma combine-matrix-zero: f 0 0 = 0 =⇒ combine-matrix f 0 0 = 0apply (simp add : combine-matrix-def combine-infmatrix-def )by (simp add : zero-matrix-def )

lemma transpose-matrix-zero[simp]: transpose-matrix 0 = 0apply (simp add : transpose-matrix-def zero-matrix-def RepAbs-matrix )apply (subst Rep-matrix-inject [symmetric], (rule ext)+)apply (simp add : RepAbs-matrix )done

lemma apply-zero-matrix-def [simp]: apply-matrix (% x . 0 ) A = 0apply (simp add : apply-matrix-def apply-infmatrix-def )by (simp add : zero-matrix-def )

definition singleton-matrix :: nat ⇒ nat ⇒ ( ′a::zero) ⇒ ′a matrix wheresingleton-matrix j i a == Abs-matrix (% m n. if j = m & i = n then a else 0 )

definition move-matrix :: ( ′a::zero) matrix ⇒ int ⇒ int ⇒ ′a matrix wheremove-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 )))

definition take-rows :: ( ′a::zero) matrix ⇒ nat ⇒ ′a matrix wheretake-rows A r == Abs-matrix (% j i . if (j < r) then (Rep-matrix A j i) else 0 )

definition take-columns :: ( ′a::zero) matrix ⇒ nat ⇒ ′a matrix wheretake-columns A c == Abs-matrix (% j i . if (i < c) then (Rep-matrix A j i) else

0 )

definition column-of-matrix :: ( ′a::zero) matrix ⇒ nat ⇒ ′a matrix wherecolumn-of-matrix A n == take-columns (move-matrix A 0 (− int n)) 1

definition row-of-matrix :: ( ′a::zero) matrix ⇒ nat ⇒ ′a matrix whererow-of-matrix A m == take-rows (move-matrix A (− int m) 0 ) 1

19

Page 20: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

lemma Rep-singleton-matrix [simp]: Rep-matrix (singleton-matrix j i e) m n = (ifj = m & i = n then e else 0 )apply (simp add : singleton-matrix-def )apply (auto)apply (subst RepAbs-matrix )apply (rule exI [of - Suc m], simp)apply (rule exI [of - Suc n], simp+)by (subst RepAbs-matrix , rule exI [of - Suc j ], simp, rule exI [of - Suc i ], simp+)+

lemma apply-singleton-matrix [simp]: f 0 = 0 =⇒ apply-matrix f (singleton-matrixj i x ) = (singleton-matrix j i (f x ))apply (subst Rep-matrix-inject [symmetric])apply (rule ext)+apply (simp)done

lemma singleton-matrix-zero[simp]: singleton-matrix j i 0 = 0by (simp add : singleton-matrix-def zero-matrix-def )

lemma nrows-singleton[simp]: nrows(singleton-matrix j i e) = (if e = 0 then 0else Suc j )proof−have th: ¬ (∀m. m ≤ j ) ∃n. ¬ n ≤ i by arith+from th show ?thesisapply (auto)apply (rule le-antisym)apply (subst nrows-le)apply (simp add : singleton-matrix-def , auto)apply (subst RepAbs-matrix )apply autoapply (simp add : Suc-le-eq)apply (rule not-le-imp-less)apply (subst nrows-le)by simpqed

lemma ncols-singleton[simp]: ncols(singleton-matrix j i e) = (if e = 0 then 0 elseSuc i)proof−have th: ¬ (∀m. m ≤ j ) ∃n. ¬ n ≤ i by arith+from th show ?thesisapply (auto)apply (rule le-antisym)apply (subst ncols-le)apply (simp add : singleton-matrix-def , auto)apply (subst RepAbs-matrix )apply autoapply (simp add : Suc-le-eq)

20

Page 21: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

apply (rule not-le-imp-less)apply (subst ncols-le)by simpqed

lemma combine-singleton: f 0 0 = 0 =⇒ combine-matrix f (singleton-matrix j ia) (singleton-matrix j i b) = singleton-matrix j i (f a b)apply (simp add : singleton-matrix-def combine-matrix-def combine-infmatrix-def )apply (subst RepAbs-matrix )apply (rule exI [of - Suc j ], simp)apply (rule exI [of - Suc i ], simp)apply (rule comb[of Abs-matrix Abs-matrix ], simp, (rule ext)+)apply (subst RepAbs-matrix )apply (rule exI [of - Suc j ], simp)apply (rule exI [of - Suc i ], simp)by simp

lemma transpose-singleton[simp]: transpose-matrix (singleton-matrix j i a) = singleton-matrixi j aapply (subst Rep-matrix-inject [symmetric], (rule ext)+)apply (simp)done

lemma Rep-move-matrix [simp]:Rep-matrix (move-matrix A y x ) 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 )))apply (simp add : move-matrix-def )apply (auto)by (subst RepAbs-matrix ,

rule exI [of - (nrows A)+(nat |y |)], auto, rule nrows, arith,rule exI [of - (ncols A)+(nat |x |)], auto, rule ncols, arith)+

lemma move-matrix-0-0 [simp]: move-matrix A 0 0 = Aby (simp add : move-matrix-def )

lemma move-matrix-ortho: move-matrix A j i = move-matrix (move-matrix A j0 ) 0 iapply (subst Rep-matrix-inject [symmetric])apply (rule ext)+apply (simp)done

lemma transpose-move-matrix [simp]:transpose-matrix (move-matrix A x y) = move-matrix (transpose-matrix A) y x

apply (subst Rep-matrix-inject [symmetric], (rule ext)+)apply (simp)done

21

Page 22: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

lemma move-matrix-singleton[simp]: move-matrix (singleton-matrix u v x ) j i =(if (j + int u < 0 ) | (i + int v < 0 ) then 0 else (singleton-matrix (nat (j + int

u)) (nat (i + int v)) x ))apply (subst Rep-matrix-inject [symmetric])apply (rule ext)+apply (case-tac j + int u < 0 )apply (simp, arith)apply (case-tac i + int v < 0 )apply (simp, arith)apply simpapply arithdone

lemma Rep-take-columns[simp]:Rep-matrix (take-columns A c) j i =(if i < c then (Rep-matrix A j i) else 0 )

apply (simp add : take-columns-def )apply (simplesubst RepAbs-matrix )apply (rule exI [of - nrows A], auto, simp add : nrows-le)apply (rule exI [of - ncols A], auto, simp add : ncols-le)done

lemma Rep-take-rows[simp]:Rep-matrix (take-rows A r) j i =(if j < r then (Rep-matrix A j i) else 0 )

apply (simp add : take-rows-def )apply (simplesubst RepAbs-matrix )apply (rule exI [of - nrows A], auto, simp add : nrows-le)apply (rule exI [of - ncols A], auto, simp add : ncols-le)done

lemma Rep-column-of-matrix [simp]:Rep-matrix (column-of-matrix A c) j i = (if i = 0 then (Rep-matrix A j c) else

0 )by (simp add : column-of-matrix-def )

lemma Rep-row-of-matrix [simp]:Rep-matrix (row-of-matrix A r) j i = (if j = 0 then (Rep-matrix A r i) else 0 )by (simp add : row-of-matrix-def )

lemma column-of-matrix : ncols A <= n =⇒ column-of-matrix A n = 0apply (subst Rep-matrix-inject [THEN sym])apply (rule ext)+by (simp add : ncols)

lemma row-of-matrix : nrows A <= n =⇒ row-of-matrix A n = 0apply (subst Rep-matrix-inject [THEN sym])apply (rule ext)+by (simp add : nrows)

22

Page 23: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

lemma mult-matrix-singleton-right [simp]:assumes∀ x . fmul x 0 = 0∀ x . fmul 0 x = 0∀ x . fadd 0 x = x∀ x . fadd x 0 = xshows (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))apply (simp add : mult-matrix-def )apply (subst mult-matrix-nm[of - - - max (ncols A) (Suc j )])apply (auto)apply (simp add : assms)+apply (simp add : mult-matrix-n-def apply-matrix-def apply-infmatrix-def )apply (rule comb[of Abs-matrix Abs-matrix ], auto, (rule ext)+)apply (subst foldseq-almostzero[of - j ])apply (simp add : assms)+apply (auto)done

lemma mult-matrix-ext :assumeseprem:∃ e. (∀ a b. a 6= b −→ fmul a e 6= fmul b e)and fprems:∀ a. fmul 0 a = 0∀ a. fmul a 0 = 0∀ a. fadd a 0 = a∀ a. fadd 0 a = aand contraprems:mult-matrix fmul fadd A = mult-matrix fmul fadd BshowsA = B

proof(rule contrapos-np[of False], simp)assume a: A 6= Bhave b:

∧f g . (∀ x y . f x y = g x y) =⇒ f = g by ((rule ext)+, auto)

have ∃ j i . (Rep-matrix A j i) 6= (Rep-matrix B j i)apply (rule contrapos-np[of False], simp+)apply (insert b[of Rep-matrix A Rep-matrix B ], simp)by (simp add : Rep-matrix-inject a)

then obtain J I where c:(Rep-matrix A J I ) 6= (Rep-matrix B J I ) by blastfrom eprem obtain e where eprops:(∀ a b. a 6= b −→ fmul a e 6= fmul b e) by

blastlet ?S = singleton-matrix I 0 elet ?comp = mult-matrix fmul faddhave d : !!x f g . f = g =⇒ f x = g x by blasthave e: (% x . fmul x e) 0 = 0 by (simp add : assms)have ∼(?comp A ?S = ?comp B ?S )

apply (rule notI )

23

Page 24: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

apply (simp add : fprems eprops)apply (simp add : Rep-matrix-inject [THEN sym])apply (drule d [of - - J ], drule d [of - - 0 ])by (simp add : e c eprops)

with contraprems show False by simpqed

definition foldmatrix :: ( ′a ⇒ ′a ⇒ ′a) ⇒ ( ′a ⇒ ′a ⇒ ′a) ⇒ ( ′a infmatrix ) ⇒nat ⇒ nat ⇒ ′a where

foldmatrix f g A m n == foldseq-transposed g (% j . foldseq f (A j ) n) m

definition foldmatrix-transposed :: ( ′a ⇒ ′a ⇒ ′a) ⇒ ( ′a ⇒ ′a ⇒ ′a) ⇒ ( ′ainfmatrix ) ⇒ nat ⇒ nat ⇒ ′a where

foldmatrix-transposed f g A m n == foldseq g (% j . foldseq-transposed f (A j ) n)m

lemma foldmatrix-transpose:assumes∀ a b c d . g(f a b) (f c d) = f (g a c) (g b d)showsfoldmatrix f g A m n = foldmatrix-transposed g f (transpose-infmatrix A) n m

proof −have forall :

∧P x . (∀ x . P x ) =⇒ P x by auto

have tworows:∀A. foldmatrix f g A 1 n = foldmatrix-transposed g f (transpose-infmatrixA) n 1

apply (induct n)apply (simp add : foldmatrix-def foldmatrix-transposed-def assms)+apply (auto)by (drule-tac x=(% j i . A j (Suc i)) in forall , simp)

show foldmatrix f g A m n = foldmatrix-transposed g f (transpose-infmatrix A)n m

apply (simp add : foldmatrix-def foldmatrix-transposed-def )apply (induct m, simp)apply (simp)apply (insert tworows)apply (drule-tac x=% j i . (if j = 0 then (foldseq-transposed g (λu. A u i) m)

else (A (Suc m) i)) in spec)by (simp add : foldmatrix-def foldmatrix-transposed-def )

qed

lemma foldseq-foldseq :assumes

associative fassociative g∀ a b c d . g(f a b) (f c d) = f (g a c) (g b d)

showsfoldseq g (% j . foldseq f (A j ) n) m = foldseq f (% j . foldseq g ((transpose-infmatrix

A) j ) m) napply (insert foldmatrix-transpose[of g f A m n])

24

Page 25: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

by (simp add : foldmatrix-def foldmatrix-transposed-def foldseq-assoc[THEN sym]assms)

lemma mult-n-nrows:assumes∀ a. fmul 0 a = 0∀ a. fmul a 0 = 0fadd 0 0 = 0shows nrows (mult-matrix-n n fmul fadd A B) ≤ nrows Aapply (subst nrows-le)apply (simp add : mult-matrix-n-def )apply (subst RepAbs-matrix )apply (rule-tac x=nrows A in exI )apply (simp add : nrows assms foldseq-zero)apply (rule-tac x=ncols B in exI )apply (simp add : ncols assms foldseq-zero)apply (simp add : nrows assms foldseq-zero)done

lemma mult-n-ncols:assumes∀ a. fmul 0 a = 0∀ a. fmul a 0 = 0fadd 0 0 = 0shows ncols (mult-matrix-n n fmul fadd A B) ≤ ncols Bapply (subst ncols-le)apply (simp add : mult-matrix-n-def )apply (subst RepAbs-matrix )apply (rule-tac x=nrows A in exI )apply (simp add : nrows assms foldseq-zero)apply (rule-tac x=ncols B in exI )apply (simp add : ncols assms foldseq-zero)apply (simp add : ncols assms foldseq-zero)done

lemma mult-nrows:assumes∀ a. fmul 0 a = 0∀ a. fmul a 0 = 0fadd 0 0 = 0shows nrows (mult-matrix fmul fadd A B) ≤ nrows Aby (simp add : mult-matrix-def mult-n-nrows assms)

lemma mult-ncols:assumes∀ a. fmul 0 a = 0∀ a. fmul a 0 = 0fadd 0 0 = 0shows ncols (mult-matrix fmul fadd A B) ≤ ncols B

25

Page 26: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

by (simp add : mult-matrix-def mult-n-ncols assms)

lemma nrows-move-matrix-le: nrows (move-matrix A j i) <= nat((int (nrows A))+ j )

apply (auto simp add : nrows-le)apply (rule nrows)apply (arith)done

lemma ncols-move-matrix-le: ncols (move-matrix A j i) <= nat((int (ncols A))+ i)

apply (auto simp add : ncols-le)apply (rule ncols)apply (arith)done

lemma mult-matrix-assoc:assumes∀ a. fmul1 0 a = 0∀ a. fmul1 a 0 = 0∀ a. fmul2 0 a = 0∀ a. fmul2 a 0 = 0fadd1 0 0 = 0fadd2 0 0 = 0∀ a b c d . fadd2 (fadd1 a b) (fadd1 c d) = fadd1 (fadd2 a c) (fadd2 b d)associative fadd1associative fadd2∀ a b c. fmul2 (fmul1 a b) c = fmul1 a (fmul2 b c)∀ a b c. fmul2 (fadd1 a b) c = fadd1 (fmul2 a c) (fmul2 b c)∀ a b c. fmul1 c (fadd2 a b) = fadd2 (fmul1 c a) (fmul1 c b)shows mult-matrix fmul2 fadd2 (mult-matrix fmul1 fadd1 A B) C = mult-matrix

fmul1 fadd1 A (mult-matrix fmul2 fadd2 B C )proof −

have comb-left : !! A B x y . A = B =⇒ (Rep-matrix (Abs-matrix A)) x y =(Rep-matrix (Abs-matrix B)) x y by blast

have fmul2fadd1fold : !! x s n. fmul2 (foldseq fadd1 s n) x = foldseq fadd1 (%k . fmul2 (s k) x ) n

by (rule-tac g1 = % y . fmul2 y x in ssubst [OF foldseq-distr-unary ], insertassms, simp-all)

have fmul1fadd2fold : !! x s n. fmul1 x (foldseq fadd2 s n) = foldseq fadd2 (% k .fmul1 x (s k)) n

using assms by (rule-tac g1 = % y . fmul1 x y in ssubst [OF foldseq-distr-unary ],simp-all)

let ?N = max (ncols A) (max (ncols B) (max (nrows B) (nrows C )))show ?thesis

apply (simp add : Rep-matrix-inject [THEN sym])apply (rule ext)+apply (simp add : mult-matrix-def )apply (simplesubst mult-matrix-nm[of - max (ncols (mult-matrix-n (max (ncols

26

Page 27: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

A) (nrows B)) fmul1 fadd1 A B)) (nrows C ) - max (ncols B) (nrows C )])apply (simp add : max1 max2 mult-n-ncols mult-n-nrows assms)+apply (simplesubst mult-matrix-nm[of - max (ncols A) (nrows (mult-matrix-n

(max (ncols B) (nrows C )) fmul2 fadd2 B C )) - max (ncols A) (nrows B)])apply (simp add : max1 max2 mult-n-ncols mult-n-nrows assms)+apply (simplesubst mult-matrix-nm[of - - - ?N ])apply (simp add : max1 max2 mult-n-ncols mult-n-nrows assms)+apply (simplesubst mult-matrix-nm[of - - - ?N ])apply (simp add : max1 max2 mult-n-ncols mult-n-nrows assms)+apply (simplesubst mult-matrix-nm[of - - - ?N ])apply (simp add : max1 max2 mult-n-ncols mult-n-nrows assms)+apply (simplesubst mult-matrix-nm[of - - - ?N ])apply (simp add : max1 max2 mult-n-ncols mult-n-nrows assms)+apply (simp add : mult-matrix-n-def )apply (rule comb-left)apply ((rule ext)+, simp)apply (simplesubst RepAbs-matrix )apply (rule exI [of - nrows B ])apply (simp add : nrows assms foldseq-zero)apply (rule exI [of - ncols C ])apply (simp add : assms ncols foldseq-zero)apply (subst RepAbs-matrix )apply (rule exI [of - nrows A])apply (simp add : nrows assms foldseq-zero)apply (rule exI [of - ncols B ])apply (simp add : assms ncols foldseq-zero)apply (simp add : fmul2fadd1fold fmul1fadd2fold assms)apply (subst foldseq-foldseq)apply (simp add : assms)+apply (simp add : transpose-infmatrix )done

qed

lemmaassumes∀ a. fmul1 0 a = 0∀ a. fmul1 a 0 = 0∀ a. fmul2 0 a = 0∀ a. fmul2 a 0 = 0fadd1 0 0 = 0fadd2 0 0 = 0∀ a b c d . fadd2 (fadd1 a b) (fadd1 c d) = fadd1 (fadd2 a c) (fadd2 b d)associative fadd1associative fadd2∀ a b c. fmul2 (fmul1 a b) c = fmul1 a (fmul2 b c)∀ a b c. fmul2 (fadd1 a b) c = fadd1 (fmul2 a c) (fmul2 b c)∀ a b c. fmul1 c (fadd2 a b) = fadd2 (fmul1 c a) (fmul1 c b)shows(mult-matrix fmul1 fadd1 A) o (mult-matrix fmul2 fadd2 B) = mult-matrix fmul2

27

Page 28: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

fadd2 (mult-matrix fmul1 fadd1 A B)apply (rule ext)+apply (simp add : comp-def )apply (simp add : mult-matrix-assoc assms)done

lemma mult-matrix-assoc-simple:assumes∀ a. fmul 0 a = 0∀ a. fmul a 0 = 0fadd 0 0 = 0associative faddcommutative faddassociative fmuldistributive fmul faddshows mult-matrix fmul fadd (mult-matrix fmul fadd A B) C = mult-matrix fmul

fadd A (mult-matrix fmul fadd B C )proof −

have !! a b c d . fadd (fadd a b) (fadd c d) = fadd (fadd a c) (fadd b d)using assms by (simp add : associative-def commutative-def )

then show ?thesisapply (subst mult-matrix-assoc)using assmsapply simp-all

apply (simp-all add : associative-def distributive-def l-distributive-def r-distributive-def )done

qed

lemma transpose-apply-matrix : f 0 = 0 =⇒ transpose-matrix (apply-matrix f A)= apply-matrix f (transpose-matrix A)apply (simp add : Rep-matrix-inject [THEN sym])apply (rule ext)+by simp

lemma transpose-combine-matrix : f 0 0 = 0 =⇒ transpose-matrix (combine-matrixf A B) = combine-matrix f (transpose-matrix A) (transpose-matrix B)apply (simp add : Rep-matrix-inject [THEN sym])apply (rule ext)+by simp

lemma Rep-mult-matrix :assumes∀ a. fmul 0 a = 0∀ a. fmul a 0 = 0fadd 0 0 = 0showsRep-matrix (mult-matrix fmul fadd A B) j i =foldseq fadd (% k . fmul (Rep-matrix A j k) (Rep-matrix B k i)) (max (ncols A)

(nrows B))

28

Page 29: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

apply (simp add : mult-matrix-def mult-matrix-n-def )apply (subst RepAbs-matrix )apply (rule exI [of - nrows A], insert assms, simp add : nrows foldseq-zero)apply (rule exI [of - ncols B ], insert assms, simp add : ncols foldseq-zero)apply simpdone

lemma transpose-mult-matrix :assumes∀ a. fmul 0 a = 0∀ a. fmul a 0 = 0fadd 0 0 = 0∀ x y . fmul y x = fmul x yshowstranspose-matrix (mult-matrix fmul fadd A B) = mult-matrix fmul fadd (transpose-matrix

B) (transpose-matrix A)apply (simp add : Rep-matrix-inject [THEN sym])apply (rule ext)+using assmsapply (simp add : Rep-mult-matrix ac-simps)done

lemma column-transpose-matrix : column-of-matrix (transpose-matrix A) n = transpose-matrix(row-of-matrix A n)apply (simp add : Rep-matrix-inject [THEN sym])apply (rule ext)+by simp

lemma take-columns-transpose-matrix : take-columns (transpose-matrix A) n =transpose-matrix (take-rows A n)apply (simp add : Rep-matrix-inject [THEN sym])apply (rule ext)+by simp

instantiation matrix :: ({zero, ord}) ordbegin

definitionle-matrix-def : A ≤ B ←→ (∀ j i . Rep-matrix A j i ≤ Rep-matrix B j i)

definitionless-def : A < (B :: ′a matrix ) ←→ A ≤ B ∧ ¬ B ≤ A

instance ..

end

instance matrix :: ({zero, order}) orderapply intro-classes

29

Page 30: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

apply (simp-all add : le-matrix-def less-def )apply (auto)apply (drule-tac x=j in spec, drule-tac x=j in spec)apply (drule-tac x=i in spec, drule-tac x=i in spec)apply (simp)apply (simp add : Rep-matrix-inject [THEN sym])apply (rule ext)+apply (drule-tac x=xa in spec, drule-tac x=xa in spec)apply (drule-tac x=xb in spec, drule-tac x=xb in spec)apply simpdone

lemma le-apply-matrix :assumesf 0 = 0∀ x y . x <= y −→ f x <= f y(a::( ′a::{ord , zero}) matrix ) <= bshowsapply-matrix f a <= apply-matrix f busing assms by (simp add : le-matrix-def )

lemma le-combine-matrix :assumesf 0 0 = 0∀ a b c d . a <= b & c <= d −→ f a c <= f b dA <= BC <= Dshowscombine-matrix f A C <= combine-matrix f B Dusing assms by (simp add : le-matrix-def )

lemma le-left-combine-matrix :assumesf 0 0 = 0∀ a b c. a <= b −→ f c a <= f c bA <= Bshowscombine-matrix f C A <= combine-matrix f C Busing assms by (simp add : le-matrix-def )

lemma le-right-combine-matrix :assumesf 0 0 = 0∀ a b c. a <= b −→ f a c <= f b cA <= Bshowscombine-matrix f A C <= combine-matrix f B Cusing assms by (simp add : le-matrix-def )

30

Page 31: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

lemma le-transpose-matrix : (A <= B) = (transpose-matrix A <= transpose-matrixB)

by (simp add : le-matrix-def , auto)

lemma le-foldseq :assumes∀ a b c d . a <= b & c <= d −→ f a c <= f b d∀ i . i <= n −→ s i <= t ishowsfoldseq f s n <= foldseq f t n

proof −have ∀ s t . (∀ i . i<=n −→ s i <= t i) −→ foldseq f s n <= foldseq f t n

by (induct n) (simp-all add : assms)then show foldseq f s n <= foldseq f t n using assms by simp

qed

lemma le-left-mult :assumes∀ a b c d . a <= b & c <= d −→ fadd a c <= fadd b d∀ c a b. 0 <= c & a <= b −→ fmul c a <= fmul c b∀ a. fmul 0 a = 0∀ a. fmul a 0 = 0fadd 0 0 = 00 <= CA <= Bshowsmult-matrix fmul fadd C A <= mult-matrix fmul fadd C Busing assmsapply (simp add : le-matrix-def Rep-mult-matrix )apply (auto)apply (simplesubst foldseq-zerotail [of - - - max (ncols C ) (max (nrows A) (nrows

B))], simp-all add : nrows ncols max1 max2 )+apply (rule le-foldseq)apply (auto)done

lemma le-right-mult :assumes∀ a b c d . a <= b & c <= d −→ fadd a c <= fadd b d∀ c a b. 0 <= c & a <= b −→ fmul a c <= fmul b c∀ a. fmul 0 a = 0∀ a. fmul a 0 = 0fadd 0 0 = 00 <= CA <= Bshowsmult-matrix fmul fadd A C <= mult-matrix fmul fadd B Cusing assmsapply (simp add : le-matrix-def Rep-mult-matrix )

31

Page 32: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

apply (auto)apply (simplesubst foldseq-zerotail [of - - - max (nrows C ) (max (ncols A) (ncols

B))], simp-all add : nrows ncols max1 max2 )+apply (rule le-foldseq)apply (auto)done

lemma spec2 : ∀ j i . P j i =⇒ P j i by blastlemma neg-imp: (¬ Q −→ ¬ P) =⇒ P −→ Q by blast

lemma singleton-matrix-le[simp]: (singleton-matrix j i a <= singleton-matrix j ib) = (a <= (b::-::order))

by (auto simp add : le-matrix-def )

lemma singleton-le-zero[simp]: (singleton-matrix j i x <= 0 ) = (x <= (0 :: ′a::{order ,zero}))apply (auto)apply (simp add : le-matrix-def )apply (drule-tac j =j and i=i in spec2 )apply (simp)apply (simp add : le-matrix-def )done

lemma singleton-ge-zero[simp]: (0 <= singleton-matrix j i x ) = ((0 :: ′a::{order ,zero})<= x )

apply (auto)apply (simp add : le-matrix-def )apply (drule-tac j =j and i=i in spec2 )apply (simp)apply (simp add : le-matrix-def )done

lemma move-matrix-le-zero[simp]: 0 <= j =⇒ 0 <= i =⇒ (move-matrix A j i<= 0 ) = (A <= (0 ::( ′a::{order ,zero}) matrix ))

apply (auto simp add : le-matrix-def )apply (drule-tac j =ja+(nat j ) and i=ia+(nat i) in spec2 )apply (auto)done

lemma move-matrix-zero-le[simp]: 0 <= j =⇒ 0 <= i =⇒ (0 <= move-matrixA j i) = ((0 ::( ′a::{order ,zero}) matrix ) <= A)

apply (auto simp add : le-matrix-def )apply (drule-tac j =ja+(nat j ) and i=ia+(nat i) in spec2 )apply (auto)done

lemma move-matrix-le-move-matrix-iff [simp]: 0 <= j =⇒ 0 <= i =⇒ (move-matrixA j i <= move-matrix B j i) = (A <= (B ::( ′a::{order ,zero}) matrix ))

apply (auto simp add : le-matrix-def )apply (drule-tac j =ja+(nat j ) and i=ia+(nat i) in spec2 )

32

Page 33: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

apply (auto)done

instantiation matrix :: ({lattice, zero}) latticebegin

definition inf = combine-matrix inf

definition sup = combine-matrix sup

instanceby standard (auto simp add : le-infI le-matrix-def inf-matrix-def sup-matrix-def )

end

instantiation matrix :: ({plus, zero}) plusbegin

definitionplus-matrix-def : A + B = combine-matrix (+) A B

instance ..

end

instantiation matrix :: ({uminus, zero}) uminusbegin

definitionminus-matrix-def : − A = apply-matrix uminus A

instance ..

end

instantiation matrix :: ({minus, zero}) minusbegin

definitiondiff-matrix-def : A − B = combine-matrix (−) A B

instance ..

end

instantiation matrix :: ({plus, times, zero}) timesbegin

definition

33

Page 34: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

times-matrix-def : A ∗ B = mult-matrix ((∗)) (+) A B

instance ..

end

instantiation matrix :: ({lattice, uminus, zero}) absbegin

definitionabs-matrix-def : |A :: ′a matrix | = sup A (− A)

instance ..

end

instance matrix :: (monoid-add) monoid-addproof

fix A B C :: ′a matrixshow A + B + C = A + (B + C )

apply (simp add : plus-matrix-def )apply (rule combine-matrix-assoc[simplified associative-def , THEN spec, THEN

spec, THEN spec])apply (simp-all add : add .assoc)done

show 0 + A = Aapply (simp add : plus-matrix-def )apply (rule combine-matrix-zero-l-neutral [simplified zero-l-neutral-def , THEN

spec])apply (simp)done

show A + 0 = Aapply (simp add : plus-matrix-def )apply (rule combine-matrix-zero-r-neutral [simplified zero-r-neutral-def , THEN

spec])apply (simp)done

qed

instance matrix :: (comm-monoid-add) comm-monoid-addproof

fix A B :: ′a matrixshow A + B = B + A

apply (simp add : plus-matrix-def )apply (rule combine-matrix-commute[simplified commutative-def , THEN spec,

THEN spec])apply (simp-all add : add .commute)done

show 0 + A = A

34

Page 35: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

apply (simp add : plus-matrix-def )apply (rule combine-matrix-zero-l-neutral [simplified zero-l-neutral-def , THEN

spec])apply (simp)done

qed

instance matrix :: (group-add) group-addproof

fix A B :: ′a matrixshow − A + A = 0

by (simp add : plus-matrix-def minus-matrix-def Rep-matrix-inject [symmetric]ext)

show A + − B = A − Bby (simp add : plus-matrix-def diff-matrix-def minus-matrix-def Rep-matrix-inject

[symmetric] ext)qed

instance matrix :: (ab-group-add) ab-group-addproof

fix A B :: ′a matrixshow − A + A = 0

by (simp add : plus-matrix-def minus-matrix-def Rep-matrix-inject [symmetric]ext)

show A − B = A + − Bby (simp add : plus-matrix-def diff-matrix-def minus-matrix-def Rep-matrix-inject [symmetric]

ext)qed

instance matrix :: (ordered-ab-group-add) ordered-ab-group-addproof

fix A B C :: ′a matrixassume A <= Bthen show C + A <= C + B

apply (simp add : plus-matrix-def )apply (rule le-left-combine-matrix )apply (simp-all)done

qed

instance matrix :: (lattice-ab-group-add) semilattice-inf-ab-group-add ..instance matrix :: (lattice-ab-group-add) semilattice-sup-ab-group-add ..

instance matrix :: (semiring-0 ) semiring-0proof

fix A B C :: ′a matrixshow A ∗ B ∗ C = A ∗ (B ∗ C )

apply (simp add : times-matrix-def )apply (rule mult-matrix-assoc)

35

Page 36: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

apply (simp-all add : associative-def algebra-simps)done

show (A + B) ∗ C = A ∗ C + B ∗ Capply (simp add : times-matrix-def plus-matrix-def )

apply (rule l-distributive-matrix [simplified l-distributive-def , THEN spec, THENspec, THEN spec])

apply (simp-all add : associative-def commutative-def algebra-simps)done

show A ∗ (B + C ) = A ∗ B + A ∗ Capply (simp add : times-matrix-def plus-matrix-def )

apply (rule r-distributive-matrix [simplified r-distributive-def , THEN spec, THENspec, THEN spec])

apply (simp-all add : associative-def commutative-def algebra-simps)done

show 0 ∗ A = 0 by (simp add : times-matrix-def )show A ∗ 0 = 0 by (simp add : times-matrix-def )

qed

instance matrix :: (ring) ring ..

instance matrix :: (ordered-ring) ordered-ringproof

fix A B C :: ′a matrixassume a: A ≤ Bassume b: 0 ≤ Cfrom a b show C ∗ A ≤ C ∗ B

apply (simp add : times-matrix-def )apply (rule le-left-mult)apply (simp-all add : add-mono mult-left-mono)done

from a b show A ∗ C ≤ B ∗ Capply (simp add : times-matrix-def )apply (rule le-right-mult)apply (simp-all add : add-mono mult-right-mono)done

qed

instance matrix :: (lattice-ring) lattice-ringproof

fix A B C :: ( ′a :: lattice-ring) matrixshow |A| = sup A (−A)

by (simp add : abs-matrix-def )qed

lemma Rep-matrix-add [simp]:Rep-matrix ((a::( ′a::monoid-add)matrix )+b) j i = (Rep-matrix a j i) + (Rep-matrix

b j i)by (simp add : plus-matrix-def )

36

Page 37: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

lemma Rep-matrix-mult : Rep-matrix ((a::( ′a::semiring-0 ) matrix ) ∗ b) j i =foldseq (+) (% k . (Rep-matrix a j k) ∗ (Rep-matrix b k i)) (max (ncols a) (nrows

b))apply (simp add : times-matrix-def )apply (simp add : Rep-mult-matrix )done

lemma apply-matrix-add : ∀ x y . f (x+y) = (f x ) + (f y) =⇒ f 0 = (0 :: ′a)=⇒ apply-matrix f ((a::( ′a::monoid-add) matrix ) + b) = (apply-matrix f a) +

(apply-matrix f b)apply (subst Rep-matrix-inject [symmetric])apply (rule ext)+apply (simp)done

lemma singleton-matrix-add : singleton-matrix j i ((a::-::monoid-add)+b) = (singleton-matrixj i a) + (singleton-matrix j i b)apply (subst Rep-matrix-inject [symmetric])apply (rule ext)+apply (simp)done

lemma nrows-mult : nrows ((A::( ′a::semiring-0 ) matrix ) ∗ B) <= nrows Aby (simp add : times-matrix-def mult-nrows)

lemma ncols-mult : ncols ((A::( ′a::semiring-0 ) matrix ) ∗ B) <= ncols Bby (simp add : times-matrix-def mult-ncols)

definitionone-matrix :: nat ⇒ ( ′a::{zero,one}) matrix whereone-matrix n = Abs-matrix (% j i . if j = i & j < n then 1 else 0 )

lemma Rep-one-matrix [simp]: Rep-matrix (one-matrix n) j i = (if (j = i & j <n) then 1 else 0 )apply (simp add : one-matrix-def )apply (simplesubst RepAbs-matrix )apply (rule exI [of - n], simp add : if-split)+by (simp add : if-split)

lemma nrows-one-matrix [simp]: nrows ((one-matrix n) :: ( ′a::zero-neq-one)matrix )= n (is ?r = -)proof −

have ?r <= n by (simp add : nrows-le)moreover have n <= ?r by (simp add :le-nrows, arith)ultimately show ?r = n by simp

qed

lemma ncols-one-matrix [simp]: ncols ((one-matrix n) :: ( ′a::zero-neq-one)matrix )= n (is ?r = -)

37

Page 38: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

proof −have ?r <= n by (simp add : ncols-le)moreover have n <= ?r by (simp add : le-ncols, arith)ultimately show ?r = n by simp

qed

lemma one-matrix-mult-right [simp]: ncols A <= n =⇒ (A::( ′a::{semiring-1}) ma-trix ) ∗ (one-matrix n) = Aapply (subst Rep-matrix-inject [THEN sym])apply (rule ext)+apply (simp add : times-matrix-def Rep-mult-matrix )apply (rule-tac j1 =xa in ssubst [OF foldseq-almostzero])apply (simp-all)by (simp add : ncols)

lemma one-matrix-mult-left [simp]: nrows A <= n =⇒ (one-matrix n) ∗ A =(A::( ′a::semiring-1 ) matrix )apply (subst Rep-matrix-inject [THEN sym])apply (rule ext)+apply (simp add : times-matrix-def Rep-mult-matrix )apply (rule-tac j1 =x in ssubst [OF foldseq-almostzero])apply (simp-all)by (simp add : nrows)

lemma transpose-matrix-mult : transpose-matrix ((A::( ′a::comm-ring) matrix )∗B)= (transpose-matrix B) ∗ (transpose-matrix A)apply (simp add : times-matrix-def )apply (subst transpose-mult-matrix )apply (simp-all add : mult .commute)done

lemma transpose-matrix-add : transpose-matrix ((A::( ′a::monoid-add) matrix )+B)= transpose-matrix A + transpose-matrix Bby (simp add : plus-matrix-def transpose-combine-matrix )

lemma transpose-matrix-diff : transpose-matrix ((A::( ′a::group-add) matrix )−B)= transpose-matrix A − transpose-matrix Bby (simp add : diff-matrix-def transpose-combine-matrix )

lemma transpose-matrix-minus: transpose-matrix (−(A::( ′a::group-add) matrix ))= − transpose-matrix (A:: ′a matrix )by (simp add : minus-matrix-def transpose-apply-matrix )

definition right-inverse-matrix :: ( ′a::{ring-1}) matrix ⇒ ′a matrix ⇒ bool whereright-inverse-matrix A X == (A ∗ X = one-matrix (max (nrows A) (ncols X )))∧ nrows X ≤ ncols A

definition left-inverse-matrix :: ( ′a::{ring-1}) matrix ⇒ ′a matrix ⇒ bool whereleft-inverse-matrix A X == (X ∗ A = one-matrix (max (nrows X ) (ncols A))) ∧

38

Page 39: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

ncols X ≤ nrows A

definition inverse-matrix :: ( ′a::{ring-1}) matrix ⇒ ′a matrix ⇒ bool whereinverse-matrix A X == (right-inverse-matrix A X ) ∧ (left-inverse-matrix A X )

lemma right-inverse-matrix-dim: right-inverse-matrix A X =⇒ nrows A = ncolsXapply (insert ncols-mult [of A X ], insert nrows-mult [of A X ])by (simp add : right-inverse-matrix-def )

lemma left-inverse-matrix-dim: left-inverse-matrix A Y =⇒ ncols A = nrows Yapply (insert ncols-mult [of Y A], insert nrows-mult [of Y A])by (simp add : left-inverse-matrix-def )

lemma left-right-inverse-matrix-unique:assumes left-inverse-matrix A Y right-inverse-matrix A Xshows X = Y

proof −have Y = Y ∗ one-matrix (nrows A)

apply (subst one-matrix-mult-right)using assmsapply (simp-all add : left-inverse-matrix-def )done

also have . . . = Y ∗ (A ∗ X )apply (insert assms)apply (frule right-inverse-matrix-dim)by (simp add : right-inverse-matrix-def )

also have . . . = (Y ∗ A) ∗ X by (simp add : mult .assoc)also have . . . = X

apply (insert assms)apply (frule left-inverse-matrix-dim)

apply (simp-all add : left-inverse-matrix-def right-inverse-matrix-def one-matrix-mult-left)done

ultimately show X = Y by (simp)qed

lemma inverse-matrix-inject : [[ inverse-matrix A X ; inverse-matrix A Y ]] =⇒ X= Y

by (auto simp add : inverse-matrix-def left-right-inverse-matrix-unique)

lemma one-matrix-inverse: inverse-matrix (one-matrix n) (one-matrix n)by (simp add : inverse-matrix-def left-inverse-matrix-def right-inverse-matrix-def )

lemma zero-imp-mult-zero: (a:: ′a::semiring-0 ) = 0 | b = 0 =⇒ a ∗ b = 0by auto

lemma Rep-matrix-zero-imp-mult-zero:∀ j i k . (Rep-matrix A j k = 0 ) | (Rep-matrix B k i) = 0 =⇒ A ∗ B =

(0 ::( ′a::lattice-ring) matrix )

39

Page 40: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

apply (subst Rep-matrix-inject [symmetric])apply (rule ext)+apply (auto simp add : Rep-matrix-mult foldseq-zero zero-imp-mult-zero)done

lemma add-nrows: nrows (A::( ′a::monoid-add) matrix ) <= u =⇒ nrows B <= u=⇒ nrows (A + B) <= uapply (simp add : plus-matrix-def )apply (rule combine-nrows)apply (simp-all)done

lemma move-matrix-row-mult : move-matrix ((A::( ′a::semiring-0 ) matrix ) ∗ B) j0 = (move-matrix A j 0 ) ∗ Bapply (subst Rep-matrix-inject [symmetric])apply (rule ext)+apply (auto simp add : Rep-matrix-mult foldseq-zero)apply (rule-tac foldseq-zerotail [symmetric])apply (auto simp add : nrows zero-imp-mult-zero max2 )apply (rule order-trans)apply (rule ncols-move-matrix-le)apply (simp add : max1 )done

lemma move-matrix-col-mult : move-matrix ((A::( ′a::semiring-0 ) matrix ) ∗ B) 0i = A ∗ (move-matrix B 0 i)apply (subst Rep-matrix-inject [symmetric])apply (rule ext)+apply (auto simp add : Rep-matrix-mult foldseq-zero)apply (rule-tac foldseq-zerotail [symmetric])apply (auto simp add : ncols zero-imp-mult-zero max1 )apply (rule order-trans)apply (rule nrows-move-matrix-le)apply (simp add : max2 )done

lemma move-matrix-add : ((move-matrix (A + B) j i)::(( ′a::monoid-add) matrix ))= (move-matrix A j i) + (move-matrix B j i)apply (subst Rep-matrix-inject [symmetric])apply (rule ext)+apply (simp)done

lemma move-matrix-mult : move-matrix ((A::( ′a::semiring-0 ) matrix )∗B) j i =(move-matrix A j 0 ) ∗ (move-matrix B 0 i)by (simp add : move-matrix-ortho[of A∗B ] move-matrix-col-mult move-matrix-row-mult)

definition scalar-mult :: ( ′a::ring) ⇒ ′a matrix ⇒ ′a matrix wherescalar-mult a m == apply-matrix ((∗) a) m

40

Page 41: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

lemma scalar-mult-zero[simp]: scalar-mult y 0 = 0by (simp add : scalar-mult-def )

lemma scalar-mult-add : scalar-mult y (a+b) = (scalar-mult y a) + (scalar-mult yb)by (simp add : scalar-mult-def apply-matrix-add algebra-simps)

lemma Rep-scalar-mult [simp]: Rep-matrix (scalar-mult y a) j i = y ∗ (Rep-matrixa j i)by (simp add : scalar-mult-def )

lemma scalar-mult-singleton[simp]: scalar-mult y (singleton-matrix j i x ) = singleton-matrixj i (y ∗ x )apply (subst Rep-matrix-inject [symmetric])apply (rule ext)+apply (auto)done

lemma Rep-minus[simp]: Rep-matrix (−(A::-::group-add)) x y = − (Rep-matrixA x y)by (simp add : minus-matrix-def )

lemma Rep-abs[simp]: Rep-matrix |A::-::lattice-ab-group-add | x y = |Rep-matrixA x y |by (simp add : abs-lattice sup-matrix-def )

end

theory SparseMatriximports Matrixbegin

type-synonym ′a spvec = (nat ∗ ′a) listtype-synonym ′a spmat = ′a spvec spvec

definition sparse-row-vector :: ( ′a::ab-group-add) spvec ⇒ ′a matrixwhere sparse-row-vector arr = foldl (% m x . m + (singleton-matrix 0 (fst x )

(snd x ))) 0 arr

definition sparse-row-matrix :: ( ′a::ab-group-add) spmat ⇒ ′a matrixwhere sparse-row-matrix arr = foldl (% m r . m + (move-matrix (sparse-row-vector

(snd r)) (int (fst r)) 0 )) 0 arr

code-datatype sparse-row-vector sparse-row-matrix

lemma sparse-row-vector-empty [simp]: sparse-row-vector [] = 0by (simp add : sparse-row-vector-def )

41

Page 42: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

lemma sparse-row-matrix-empty [simp]: sparse-row-matrix [] = 0by (simp add : sparse-row-matrix-def )

lemmas [code] = sparse-row-vector-empty [symmetric]

lemma foldl-distrstart : ∀ a x y . (f (g x y) a = g x (f y a)) =⇒ (foldl f (g x y) l= g x (foldl f y l))

by (induct l arbitrary : x y , auto)

lemma sparse-row-vector-cons[simp]:sparse-row-vector (a # arr) = (singleton-matrix 0 (fst a) (snd a)) + (sparse-row-vector

arr)apply (induct arr)apply (auto simp add : sparse-row-vector-def )apply (simp add : foldl-distrstart [of λm x . m + singleton-matrix 0 (fst x ) (snd

x ) λx m. singleton-matrix 0 (fst x ) (snd x ) + m])done

lemma sparse-row-vector-append [simp]:sparse-row-vector (a @ b) = (sparse-row-vector a) + (sparse-row-vector b)by (induct a) auto

lemma nrows-spvec[simp]: nrows (sparse-row-vector x ) <= (Suc 0 )apply (induct x )apply (simp-all add : add-nrows)done

lemma sparse-row-matrix-cons: sparse-row-matrix (a#arr) = ((move-matrix (sparse-row-vector(snd a)) (int (fst a)) 0 )) + sparse-row-matrix arr

apply (induct arr)apply (auto simp add : sparse-row-matrix-def )apply (simp add : foldl-distrstart [of λm x . m + (move-matrix (sparse-row-vector

(snd x )) (int (fst x )) 0 )% a m. (move-matrix (sparse-row-vector (snd a)) (int (fst a)) 0 ) + m])

done

lemma sparse-row-matrix-append : sparse-row-matrix (arr@brr) = (sparse-row-matrixarr) + (sparse-row-matrix brr)

apply (induct arr)apply (auto simp add : sparse-row-matrix-cons)done

primrec sorted-spvec :: ′a spvec ⇒ boolwhere

sorted-spvec [] = True| sorted-spvec-step: sorted-spvec (a#as) = (case as of [] ⇒ True | b#bs ⇒ ((fst a< fst b) & (sorted-spvec as)))

42

Page 43: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

primrec sorted-spmat :: ′a spmat ⇒ boolwhere

sorted-spmat [] = True| sorted-spmat (a#as) = ((sorted-spvec (snd a)) & (sorted-spmat as))

declare sorted-spvec.simps [simp del ]

lemma sorted-spvec-empty [simp]: sorted-spvec [] = Trueby (simp add : sorted-spvec.simps)

lemma sorted-spvec-cons1 : sorted-spvec (a#as) =⇒ sorted-spvec asapply (induct as)apply (auto simp add : sorted-spvec.simps)done

lemma sorted-spvec-cons2 : sorted-spvec (a#b#t) =⇒ sorted-spvec (a#t)apply (induct t)apply (auto simp add : sorted-spvec.simps)done

lemma sorted-spvec-cons3 : sorted-spvec(a#b#t) =⇒ fst a < fst bapply (auto simp add : sorted-spvec.simps)done

lemma sorted-sparse-row-vector-zero[rule-format ]: m <= n =⇒ sorted-spvec ((n,a)#arr)−→ Rep-matrix (sparse-row-vector arr) j m = 0apply (induct arr)apply (auto)apply (frule sorted-spvec-cons2 ,simp)+apply (frule sorted-spvec-cons3 , simp)done

lemma sorted-sparse-row-matrix-zero[rule-format ]: m <= n =⇒ sorted-spvec ((n,a)#arr)−→ Rep-matrix (sparse-row-matrix arr) m j = 0

apply (induct arr)apply (auto)apply (frule sorted-spvec-cons2 , simp)apply (frule sorted-spvec-cons3 , simp)apply (simp add : sparse-row-matrix-cons)done

primrec minus-spvec :: ( ′a::ab-group-add) spvec ⇒ ′a spvecwhere

minus-spvec [] = []| minus-spvec (a#as) = (fst a, −(snd a))#(minus-spvec as)

primrec abs-spvec :: ( ′a::lattice-ab-group-add-abs) spvec ⇒ ′a spvecwhere

abs-spvec [] = []

43

Page 44: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

| abs-spvec (a#as) = (fst a, |snd a|)#(abs-spvec as)

lemma sparse-row-vector-minus:sparse-row-vector (minus-spvec v) = − (sparse-row-vector v)apply (induct v)apply (simp-all add : sparse-row-vector-cons)apply (simp add : Rep-matrix-inject [symmetric])apply (rule ext)+apply simpdone

instance matrix :: (lattice-ab-group-add-abs) lattice-ab-group-add-absapply standardunfolding abs-matrix-defapply ruledone

lemma sparse-row-vector-abs:sorted-spvec (v :: ′a::lattice-ring spvec) =⇒ sparse-row-vector (abs-spvec v) =|sparse-row-vector v |

apply (induct v)apply simp-allapply (frule-tac sorted-spvec-cons1 , simp)apply (simp only : Rep-matrix-inject [symmetric])apply (rule ext)+apply autoapply (subgoal-tac Rep-matrix (sparse-row-vector v) 0 a = 0 )apply (simp)apply (rule sorted-sparse-row-vector-zero)apply autodone

lemma sorted-spvec-minus-spvec:sorted-spvec v =⇒ sorted-spvec (minus-spvec v)apply (induct v)apply (simp)apply (frule sorted-spvec-cons1 , simp)apply (simp add : sorted-spvec.simps split :list .split-asm)done

lemma sorted-spvec-abs-spvec:sorted-spvec v =⇒ sorted-spvec (abs-spvec v)apply (induct v)apply (simp)apply (frule sorted-spvec-cons1 , simp)apply (simp add : sorted-spvec.simps split :list .split-asm)done

44

Page 45: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

definition smult-spvec y = map (% a. (fst a, y ∗ snd a))

lemma smult-spvec-empty [simp]: smult-spvec y [] = []by (simp add : smult-spvec-def )

lemma smult-spvec-cons: smult-spvec y (a#arr) = (fst a, y ∗ (snd a)) # (smult-spvecy arr)

by (simp add : smult-spvec-def )

fun addmult-spvec :: ( ′a::ring) ⇒ ′a spvec ⇒ ′a spvec ⇒ ′a spvecwhere

addmult-spvec y arr [] = arr| addmult-spvec y [] brr = smult-spvec y brr| addmult-spvec y ((i ,a)#arr) ((j ,b)#brr) = (

if i < j then ((i ,a)#(addmult-spvec y arr ((j ,b)#brr)))else (if (j < i) then ((j , y ∗ b)#(addmult-spvec y ((i ,a)#arr) brr))else ((i , a + y∗b)#(addmult-spvec y arr brr))))

lemma addmult-spvec-empty1 [simp]: addmult-spvec y [] a = smult-spvec y aby (induct a) auto

lemma addmult-spvec-empty2 [simp]: addmult-spvec y a [] = aby (induct a) auto

lemma sparse-row-vector-map: (∀ x y . f (x+y) = (f x ) + (f y)) =⇒ (f :: ′a⇒( ′a::lattice-ring))0 = 0 =⇒sparse-row-vector (map (% x . (fst x , f (snd x ))) a) = apply-matrix f (sparse-row-vector

a)apply (induct a)apply (simp-all add : apply-matrix-add)done

lemma sparse-row-vector-smult : sparse-row-vector (smult-spvec y a) = scalar-multy (sparse-row-vector a)

apply (induct a)apply (simp-all add : smult-spvec-cons scalar-mult-add)done

lemma sparse-row-vector-addmult-spvec: sparse-row-vector (addmult-spvec (y :: ′a::lattice-ring)a b) =

(sparse-row-vector a) + (scalar-mult y (sparse-row-vector b))apply (induct y a b rule: addmult-spvec.induct)apply (simp add : scalar-mult-add smult-spvec-cons sparse-row-vector-smult singleton-matrix-add)+done

lemma sorted-smult-spvec: sorted-spvec a =⇒ sorted-spvec (smult-spvec y a)apply (auto simp add : smult-spvec-def )apply (induct a)

45

Page 46: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

apply (auto simp add : sorted-spvec.simps split :list .split-asm)done

lemma sorted-spvec-addmult-spvec-helper : [[sorted-spvec (addmult-spvec y ((a, b)# arr) brr); aa < a; sorted-spvec ((a, b) # arr);

sorted-spvec ((aa, ba) # brr)]] =⇒ sorted-spvec ((aa, y ∗ ba) # addmult-spvec y((a, b) # arr) brr)

apply (induct brr)apply (auto simp add : sorted-spvec.simps)done

lemma sorted-spvec-addmult-spvec-helper2 :[[sorted-spvec (addmult-spvec y arr ((aa, ba) # brr)); a < aa; sorted-spvec ((a, b)

# arr); sorted-spvec ((aa, ba) # brr)]]=⇒ sorted-spvec ((a, b) # addmult-spvec y arr ((aa, ba) # brr))

apply (induct arr)apply (auto simp add : smult-spvec-def sorted-spvec.simps)done

lemma sorted-spvec-addmult-spvec-helper3 [rule-format ]:sorted-spvec (addmult-spvec y arr brr) −→ sorted-spvec ((aa, b) # arr) −→

sorted-spvec ((aa, ba) # brr)−→ sorted-spvec ((aa, b + y ∗ ba) # (addmult-spvec y arr brr))

apply (induct y arr brr rule: addmult-spvec.induct)apply (simp-all add : sorted-spvec.simps smult-spvec-def split :list .split)done

lemma sorted-addmult-spvec: sorted-spvec a =⇒ sorted-spvec b =⇒ sorted-spvec(addmult-spvec y a b)

apply (induct y a b rule: addmult-spvec.induct)apply (simp-all add : sorted-smult-spvec)apply (rule conjI , intro strip)apply (case-tac ∼(i < j ))apply (simp-all)apply (frule-tac as=brr in sorted-spvec-cons1 )apply (simp add : sorted-spvec-addmult-spvec-helper)apply (intro strip | rule conjI )+apply (frule-tac as=arr in sorted-spvec-cons1 )apply (simp add : sorted-spvec-addmult-spvec-helper2 )apply (intro strip)apply (frule-tac as=arr in sorted-spvec-cons1 )apply (frule-tac as=brr in sorted-spvec-cons1 )apply (simp)apply (simp-all add : sorted-spvec-addmult-spvec-helper3 )done

fun mult-spvec-spmat :: ( ′a::lattice-ring) spvec ⇒ ′a spvec ⇒ ′a spmat ⇒ ′a spvecwhere

mult-spvec-spmat c [] brr = c

46

Page 47: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

| mult-spvec-spmat c arr [] = c| mult-spvec-spmat c ((i ,a)#arr) ((j ,b)#brr) = (

if (i < j ) then mult-spvec-spmat c arr ((j ,b)#brr)else if (j < i) then mult-spvec-spmat c ((i ,a)#arr) brrelse mult-spvec-spmat (addmult-spvec a c b) arr brr)

lemma sparse-row-mult-spvec-spmat [rule-format ]: sorted-spvec (a::( ′a::lattice-ring)spvec) −→ sorted-spvec B −→sparse-row-vector (mult-spvec-spmat c a B) = (sparse-row-vector c) + (sparse-row-vector

a) ∗ (sparse-row-matrix B)proof −

have comp-1 : !! a b. a < b =⇒ Suc 0 <= nat ((int b)−(int a)) by arithhave not-iff : !! a b. a = b =⇒ (∼ a) = (∼ b) by simphave max-helper : !! a b. ∼ (a <= max (Suc a) b) =⇒ False

by arith{

fix afix vassume a:a < nrows(sparse-row-vector v)have b:nrows(sparse-row-vector v) <= 1 by simpnote dummy = less-le-trans[of a nrows (sparse-row-vector v) 1 , OF a b]then have a = 0 by simp

}note nrows-helper = thisshow ?thesis

apply (induct c a B rule: mult-spvec-spmat .induct)apply simp+apply (rule conjI )apply (intro strip)apply (frule-tac as=brr in sorted-spvec-cons1 )apply (simp add : algebra-simps sparse-row-matrix-cons)apply (simplesubst Rep-matrix-zero-imp-mult-zero)apply (simp)apply (rule disjI2 )apply (intro strip)apply (subst nrows)apply (rule order-trans[of - 1 ])apply (simp add : comp-1 )+apply (subst Rep-matrix-zero-imp-mult-zero)apply (intro strip)apply (case-tac k <= j )

apply (rule-tac m1 = k and n1 = i and a1 = a in ssubst [OF sorted-sparse-row-vector-zero])apply (simp-all)apply (rule disjI2 )apply (rule nrows)apply (rule order-trans[of - 1 ])apply (simp-all add : comp-1 )

apply (intro strip | rule conjI )+

47

Page 48: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

apply (frule-tac as=arr in sorted-spvec-cons1 )apply (simp add : algebra-simps)apply (subst Rep-matrix-zero-imp-mult-zero)apply (simp)apply (rule disjI2 )apply (intro strip)apply (simp add : sparse-row-matrix-cons)apply (case-tac i <= j )apply (erule sorted-sparse-row-matrix-zero)apply (simp-all)apply (intro strip)apply (case-tac i=j )apply (simp-all)apply (frule-tac as=arr in sorted-spvec-cons1 )apply (frule-tac as=brr in sorted-spvec-cons1 )

apply (simp add : sparse-row-matrix-cons algebra-simps sparse-row-vector-addmult-spvec)apply (rule-tac B1 = sparse-row-matrix brr in ssubst [OF Rep-matrix-zero-imp-mult-zero])apply (auto)apply (rule sorted-sparse-row-matrix-zero)apply (simp-all)

apply (rule-tac A1 = sparse-row-vector arr in ssubst [OF Rep-matrix-zero-imp-mult-zero])apply (auto)

apply (rule-tac m=k and n = j and a = a and arr=arr in sorted-sparse-row-vector-zero)apply (simp-all)apply (drule nrows-notzero)apply (drule nrows-helper)apply (arith)

apply (subst Rep-matrix-inject [symmetric])apply (rule ext)+apply (simp)apply (subst Rep-matrix-mult)apply (rule-tac j1 =j in ssubst [OF foldseq-almostzero])apply (simp-all)apply (intro strip, rule conjI )apply (intro strip)apply (drule-tac max-helper)apply (simp)apply (auto)apply (rule zero-imp-mult-zero)apply (rule disjI2 )apply (rule nrows)apply (rule order-trans[of - 1 ])apply (simp)apply (simp)done

qed

lemma sorted-mult-spvec-spmat [rule-format ]:

48

Page 49: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

sorted-spvec (c::( ′a::lattice-ring) spvec) −→ sorted-spmat B −→ sorted-spvec(mult-spvec-spmat c a B)

apply (induct c a B rule: mult-spvec-spmat .induct)apply (simp-all add : sorted-addmult-spvec)done

primrec mult-spmat :: ( ′a::lattice-ring) spmat ⇒ ′a spmat ⇒ ′a spmatwhere

mult-spmat [] A = []| mult-spmat (a#as) A = (fst a, mult-spvec-spmat [] (snd a) A)#(mult-spmat asA)

lemma sparse-row-mult-spmat :sorted-spmat A =⇒ sorted-spvec B =⇒sparse-row-matrix (mult-spmat A B) = (sparse-row-matrix A) ∗ (sparse-row-matrix

B)apply (induct A)apply (auto simp add : sparse-row-matrix-cons sparse-row-mult-spvec-spmat algebra-simps

move-matrix-mult)done

lemma sorted-spvec-mult-spmat [rule-format ]:sorted-spvec (A::( ′a::lattice-ring) spmat) −→ sorted-spvec (mult-spmat A B)apply (induct A)apply (auto)apply (drule sorted-spvec-cons1 , simp)apply (case-tac A)apply (auto simp add : sorted-spvec.simps)done

lemma sorted-spmat-mult-spmat :sorted-spmat (B ::( ′a::lattice-ring) spmat) =⇒ sorted-spmat (mult-spmat A B)apply (induct A)apply (auto simp add : sorted-mult-spvec-spmat)done

fun add-spvec :: ( ′a::lattice-ab-group-add) spvec ⇒ ′a spvec ⇒ ′a spvecwhere

add-spvec arr [] = arr| add-spvec [] brr = brr| add-spvec ((i ,a)#arr) ((j ,b)#brr) = (

if i < j then (i ,a)#(add-spvec arr ((j ,b)#brr))else if (j < i) then (j ,b) # add-spvec ((i ,a)#arr) brrelse (i , a+b) # add-spvec arr brr)

lemma add-spvec-empty1 [simp]: add-spvec [] a = aby (cases a, auto)

49

Page 50: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

lemma sparse-row-vector-add : sparse-row-vector (add-spvec a b) = (sparse-row-vectora) + (sparse-row-vector b)

apply (induct a b rule: add-spvec.induct)apply (simp-all add : singleton-matrix-add)done

fun add-spmat :: ( ′a::lattice-ab-group-add) spmat ⇒ ′a spmat ⇒ ′a spmatwhere

add-spmat [] bs = bs| add-spmat as [] = as| add-spmat ((i ,a)#as) ((j ,b)#bs) = (

if i < j then(i ,a) # add-spmat as ((j ,b)#bs)

else if j < i then(j ,b) # add-spmat ((i ,a)#as) bs

else(i , add-spvec a b) # add-spmat as bs)

lemma add-spmat-Nil2 [simp]: add-spmat as [] = asby(cases as) auto

lemma sparse-row-add-spmat : sparse-row-matrix (add-spmat A B) = (sparse-row-matrixA) + (sparse-row-matrix B)

apply (induct A B rule: add-spmat .induct)apply (auto simp add : sparse-row-matrix-cons sparse-row-vector-add move-matrix-add)done

lemmas [code] = sparse-row-add-spmat [symmetric]lemmas [code] = sparse-row-vector-add [symmetric]

lemma sorted-add-spvec-helper1 [rule-format ]: add-spvec ((a,b)#arr) brr = (ab,bb) # list −→ (ab = a | (brr 6= [] & ab = fst (hd brr)))

proof −have (∀ x ab a. x = (a,b)#arr −→ add-spvec x brr = (ab, bb) # list −→ (ab

= a | (ab = fst (hd brr))))by (induct brr rule: add-spvec.induct) (auto split :if-splits)

then show ?thesisby (case-tac brr , auto)

qed

lemma sorted-add-spmat-helper1 [rule-format ]: add-spmat ((a,b)#arr) brr = (ab,bb) # list −→ (ab = a | (brr 6= [] & ab = fst (hd brr)))

proof −have (∀ x ab a. x = (a,b)#arr −→ add-spmat x brr = (ab, bb) # list −→ (ab

= a | (ab = fst (hd brr))))by (rule add-spmat .induct) (auto split :if-splits)

then show ?thesis

50

Page 51: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

by (case-tac brr , auto)qed

lemma sorted-add-spvec-helper : add-spvec arr brr = (ab, bb) # list =⇒ ((arr 6=[] & ab = fst (hd arr)) | (brr 6= [] & ab = fst (hd brr)))

apply (induct arr brr rule: add-spvec.induct)apply (auto split :if-splits)done

lemma sorted-add-spmat-helper : add-spmat arr brr = (ab, bb) # list =⇒ ((arr 6=[] & ab = fst (hd arr)) | (brr 6= [] & ab = fst (hd brr)))

apply (induct arr brr rule: add-spmat .induct)apply (auto split :if-splits)done

lemma add-spvec-commute: add-spvec a b = add-spvec b aby (induct a b rule: add-spvec.induct) auto

lemma add-spmat-commute: add-spmat a b = add-spmat b aapply (induct a b rule: add-spmat .induct)apply (simp-all add : add-spvec-commute)done

lemma sorted-add-spvec-helper2 : add-spvec ((a,b)#arr) brr = (ab, bb) # list =⇒aa < a =⇒ sorted-spvec ((aa, ba) # brr) =⇒ aa < ab

apply (drule sorted-add-spvec-helper1 )apply (auto)apply (case-tac brr)apply (simp-all)apply (drule-tac sorted-spvec-cons3 )apply (simp)done

lemma sorted-add-spmat-helper2 : add-spmat ((a,b)#arr) brr = (ab, bb) # list=⇒ aa < a =⇒ sorted-spvec ((aa, ba) # brr) =⇒ aa < ab

apply (drule sorted-add-spmat-helper1 )apply (auto)apply (case-tac brr)apply (simp-all)apply (drule-tac sorted-spvec-cons3 )apply (simp)done

lemma sorted-spvec-add-spvec[rule-format ]: sorted-spvec a −→ sorted-spvec b −→sorted-spvec (add-spvec a b)

apply (induct a b rule: add-spvec.induct)apply (simp-all)apply (rule conjI )apply (clarsimp)

51

Page 52: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

apply (frule-tac as=brr in sorted-spvec-cons1 )apply (simp)apply (subst sorted-spvec-step)apply (clarsimp simp: sorted-add-spvec-helper2 split : list .split)apply (clarify)apply (rule conjI )apply (clarify)apply (frule-tac as=arr in sorted-spvec-cons1 , simp)apply (subst sorted-spvec-step)apply (clarsimp simp: sorted-add-spvec-helper2 add-spvec-commute split : list .split)apply (clarify)apply (frule-tac as=arr in sorted-spvec-cons1 )apply (frule-tac as=brr in sorted-spvec-cons1 )apply (simp)apply (subst sorted-spvec-step)apply (simp split : list .split)apply (clarsimp)apply (drule-tac sorted-add-spvec-helper)apply (auto simp: neq-Nil-conv)apply (drule sorted-spvec-cons3 )apply (simp)apply (drule sorted-spvec-cons3 )apply (simp)done

lemma sorted-spvec-add-spmat [rule-format ]: sorted-spvec A −→ sorted-spvec B−→ sorted-spvec (add-spmat A B)

apply (induct A B rule: add-spmat .induct)apply (simp-all)apply (rule conjI )apply (intro strip)apply (simp)apply (frule-tac as=bs in sorted-spvec-cons1 )apply (simp)apply (subst sorted-spvec-step)apply (simp split : list .split)apply (clarify , simp)apply (simp add : sorted-add-spmat-helper2 )apply (clarify)apply (rule conjI )apply (clarify)apply (frule-tac as=as in sorted-spvec-cons1 , simp)apply (subst sorted-spvec-step)apply (clarsimp simp: sorted-add-spmat-helper2 add-spmat-commute split : list .split)apply (clarsimp)apply (frule-tac as=as in sorted-spvec-cons1 )apply (frule-tac as=bs in sorted-spvec-cons1 )apply (simp)apply (subst sorted-spvec-step)

52

Page 53: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

apply (simp split : list .split)apply (clarify , simp)apply (drule-tac sorted-add-spmat-helper)apply (auto simp:neq-Nil-conv)apply (drule sorted-spvec-cons3 )apply (simp)apply (drule sorted-spvec-cons3 )apply (simp)done

lemma sorted-spmat-add-spmat [rule-format ]: sorted-spmat A =⇒ sorted-spmat B=⇒ sorted-spmat (add-spmat A B)

apply (induct A B rule: add-spmat .induct)apply (simp-all add : sorted-spvec-add-spvec)done

fun le-spvec :: ( ′a::lattice-ab-group-add) spvec ⇒ ′a spvec ⇒ boolwhere

le-spvec [] [] = True| le-spvec ((-,a)#as) [] = (a <= 0 & le-spvec as [])| le-spvec [] ((-,b)#bs) = (0 <= b & le-spvec [] bs)| le-spvec ((i ,a)#as) ((j ,b)#bs) = (

if (i < j ) then a <= 0 & le-spvec as ((j ,b)#bs)else if (j < i) then 0 <= b & le-spvec ((i ,a)#as) bselse a <= b & le-spvec as bs)

fun le-spmat :: ( ′a::lattice-ab-group-add) spmat ⇒ ′a spmat ⇒ boolwhere

le-spmat [] [] = True| le-spmat ((i ,a)#as) [] = (le-spvec a [] & le-spmat as [])| le-spmat [] ((j ,b)#bs) = (le-spvec [] b & le-spmat [] bs)| le-spmat ((i ,a)#as) ((j ,b)#bs) = (

if i < j then (le-spvec a [] & le-spmat as ((j ,b)#bs))else if j < i then (le-spvec [] b & le-spmat ((i ,a)#as) bs)else (le-spvec a b & le-spmat as bs))

definition disj-matrices :: ( ′a::zero) matrix ⇒ ′a matrix ⇒ bool wheredisj-matrices A B ←→(∀ j i . (Rep-matrix A j i 6= 0 ) −→ (Rep-matrix B j i = 0 )) & (∀ j i . (Rep-matrix

B j i 6= 0 ) −→ (Rep-matrix A j i = 0 ))

declare [[simp-depth-limit = 6 ]]

lemma disj-matrices-contr1 : disj-matrices A B =⇒ Rep-matrix A j i 6= 0 =⇒Rep-matrix B j i = 0

by (simp add : disj-matrices-def )

53

Page 54: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

lemma disj-matrices-contr2 : disj-matrices A B =⇒ Rep-matrix B j i 6= 0 =⇒Rep-matrix A j i = 0

by (simp add : disj-matrices-def )

lemma disj-matrices-add : disj-matrices A B =⇒ disj-matrices C D =⇒ disj-matricesA D =⇒ disj-matrices B C =⇒

(A + B <= C + D) = (A <= C & B <= (D ::( ′a::lattice-ab-group-add) matrix ))apply (auto)apply (simp (no-asm-use) only : le-matrix-def disj-matrices-def )apply (intro strip)apply (erule conjE )+apply (drule-tac j =j and i=i in spec2 )+apply (case-tac Rep-matrix B j i = 0 )apply (case-tac Rep-matrix D j i = 0 )apply (simp-all)apply (simp (no-asm-use) only : le-matrix-def disj-matrices-def )apply (intro strip)apply (erule conjE )+apply (drule-tac j =j and i=i in spec2 )+apply (case-tac Rep-matrix A j i = 0 )apply (case-tac Rep-matrix C j i = 0 )apply (simp-all)apply (erule add-mono)apply (assumption)done

lemma disj-matrices-zero1 [simp]: disj-matrices 0 Bby (simp add : disj-matrices-def )

lemma disj-matrices-zero2 [simp]: disj-matrices A 0by (simp add : disj-matrices-def )

lemma disj-matrices-commute: disj-matrices A B = disj-matrices B Aby (auto simp add : disj-matrices-def )

lemma disj-matrices-add-le-zero: disj-matrices A B =⇒(A + B <= 0 ) = (A <= 0 & (B ::( ′a::lattice-ab-group-add) matrix ) <= 0 )

by (rule disj-matrices-add [of A B 0 0 , simplified ])

lemma disj-matrices-add-zero-le: disj-matrices A B =⇒(0 <= A + B) = (0 <= A & 0 <= (B ::( ′a::lattice-ab-group-add) matrix ))

by (rule disj-matrices-add [of 0 0 A B , simplified ])

lemma disj-matrices-add-x-le: disj-matrices A B =⇒ disj-matrices B C =⇒(A <= B + C ) = (A <= C & 0 <= (B ::( ′a::lattice-ab-group-add) matrix ))

by (auto simp add : disj-matrices-add [of 0 A B C , simplified ])

lemma disj-matrices-add-le-x : disj-matrices A B =⇒ disj-matrices B C =⇒

54

Page 55: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

(B + A <= C ) = (A <= C & (B ::( ′a::lattice-ab-group-add) matrix ) <= 0 )by (auto simp add : disj-matrices-add [of B A 0 C ,simplified ] disj-matrices-commute)

lemma disj-sparse-row-singleton: i <= j =⇒ sorted-spvec((j ,y)#v) =⇒ disj-matrices(sparse-row-vector v) (singleton-matrix 0 i x )

apply (simp add : disj-matrices-def )apply (rule conjI )apply (rule neg-imp)apply (simp)apply (intro strip)apply (rule sorted-sparse-row-vector-zero)apply (simp-all)apply (intro strip)apply (rule sorted-sparse-row-vector-zero)apply (simp-all)done

lemma disj-matrices-x-add : disj-matrices A B =⇒ disj-matrices A C =⇒ disj-matrices(A::( ′a::lattice-ab-group-add) matrix ) (B+C )

apply (simp add : disj-matrices-def )apply (auto)apply (drule-tac j =j and i=i in spec2 )+apply (case-tac Rep-matrix B j i = 0 )apply (case-tac Rep-matrix C j i = 0 )apply (simp-all)done

lemma disj-matrices-add-x : disj-matrices A B =⇒ disj-matrices A C =⇒ disj-matrices(B+C ) (A::( ′a::lattice-ab-group-add) matrix )

by (simp add : disj-matrices-x-add disj-matrices-commute)

lemma disj-singleton-matrices[simp]: disj-matrices (singleton-matrix j i x ) (singleton-matrixu v y) = (j 6= u | i 6= v | x = 0 | y = 0 )

by (auto simp add : disj-matrices-def )

lemma disj-move-sparse-vec-mat [simplified disj-matrices-commute]:j <= a =⇒ sorted-spvec((a,c)#as) =⇒ disj-matrices (move-matrix (sparse-row-vector

b) (int j ) i) (sparse-row-matrix as)apply (auto simp add : disj-matrices-def )apply (drule nrows-notzero)apply (drule less-le-trans[OF - nrows-spvec])apply (subgoal-tac ja = j )apply (simp add : sorted-sparse-row-matrix-zero)apply (arith)apply (rule nrows)apply (rule order-trans[of - 1 -])apply (simp)apply (case-tac nat (int ja − int j ) = 0 )apply (case-tac ja = j )

55

Page 56: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

apply (simp add : sorted-sparse-row-matrix-zero)apply arith+done

lemma disj-move-sparse-row-vector-twice:j 6= u =⇒ disj-matrices (move-matrix (sparse-row-vector a) j i) (move-matrix

(sparse-row-vector b) u v)apply (auto simp add : disj-matrices-def )apply (rule nrows, rule order-trans[of - 1 ], simp, drule nrows-notzero, drule

less-le-trans[OF - nrows-spvec], arith)+done

lemma le-spvec-iff-sparse-row-le[rule-format ]: (sorted-spvec a) −→ (sorted-spvecb) −→ (le-spvec a b) = (sparse-row-vector a <= sparse-row-vector b)

apply (induct a b rule: le-spvec.induct)apply (simp-all add : sorted-spvec-cons1 disj-matrices-add-le-zero disj-matrices-add-zero-le

disj-sparse-row-singleton[OF order-refl ] disj-matrices-commute)apply (rule conjI , intro strip)apply (simp add : sorted-spvec-cons1 )apply (subst disj-matrices-add-x-le)apply (simp add : disj-sparse-row-singleton[OF less-imp-le] disj-matrices-x-add

disj-matrices-commute)apply (simp add : disj-sparse-row-singleton[OF order-refl ] disj-matrices-commute)apply (simp, blast)apply (intro strip, rule conjI , intro strip)apply (simp add : sorted-spvec-cons1 )apply (subst disj-matrices-add-le-x )apply (simp-all add : disj-sparse-row-singleton[OF order-refl ] disj-sparse-row-singleton[OF

less-imp-le] disj-matrices-commute disj-matrices-x-add)apply (blast)apply (intro strip)apply (simp add : sorted-spvec-cons1 )apply (case-tac a=b, simp-all)apply (subst disj-matrices-add)apply (simp-all add : disj-sparse-row-singleton[OF order-refl ] disj-matrices-commute)done

lemma le-spvec-empty2-sparse-row [rule-format ]: sorted-spvec b −→ le-spvec b [] =(sparse-row-vector b <= 0 )

apply (induct b)apply (simp-all add : sorted-spvec-cons1 )apply (intro strip)apply (subst disj-matrices-add-le-zero)apply (auto simp add : disj-matrices-commute disj-sparse-row-singleton[OF order-refl ]

sorted-spvec-cons1 )done

lemma le-spvec-empty1-sparse-row [rule-format ]: (sorted-spvec b) −→ (le-spvec []

56

Page 57: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

b = (0 <= sparse-row-vector b))apply (induct b)apply (simp-all add : sorted-spvec-cons1 )apply (intro strip)apply (subst disj-matrices-add-zero-le)apply (auto simp add : disj-matrices-commute disj-sparse-row-singleton[OF order-refl ]

sorted-spvec-cons1 )done

lemma le-spmat-iff-sparse-row-le[rule-format ]: (sorted-spvec A) −→ (sorted-spmatA) −→ (sorted-spvec B) −→ (sorted-spmat B) −→

le-spmat A B = (sparse-row-matrix A <= sparse-row-matrix B)apply (induct A B rule: le-spmat .induct)apply (simp add : sparse-row-matrix-cons disj-matrices-add-le-zero disj-matrices-add-zero-le

disj-move-sparse-vec-mat [OF order-refl ]disj-matrices-commute sorted-spvec-cons1 le-spvec-empty2-sparse-row le-spvec-empty1-sparse-row)+

apply (rule conjI , intro strip)apply (simp add : sorted-spvec-cons1 )apply (subst disj-matrices-add-x-le)apply (rule disj-matrices-add-x )apply (simp add : disj-move-sparse-row-vector-twice)apply (simp add : disj-move-sparse-vec-mat [OF less-imp-le] disj-matrices-commute)apply (simp add : disj-move-sparse-vec-mat [OF order-refl ] disj-matrices-commute)apply (simp, blast)apply (intro strip, rule conjI , intro strip)apply (simp add : sorted-spvec-cons1 )apply (subst disj-matrices-add-le-x )apply (simp add : disj-move-sparse-vec-mat [OF order-refl ])apply (rule disj-matrices-x-add)apply (simp add : disj-move-sparse-row-vector-twice)apply (simp add : disj-move-sparse-vec-mat [OF less-imp-le] disj-matrices-commute)apply (simp, blast)apply (intro strip)apply (case-tac i=j )apply (simp-all)apply (subst disj-matrices-add)apply (simp-all add : disj-matrices-commute disj-move-sparse-vec-mat [OF order-refl ])apply (simp add : sorted-spvec-cons1 le-spvec-iff-sparse-row-le)done

declare [[simp-depth-limit = 999 ]]

primrec abs-spmat :: ( ′a::lattice-ring) spmat ⇒ ′a spmatwhere

abs-spmat [] = []| abs-spmat (a#as) = (fst a, abs-spvec (snd a))#(abs-spmat as)

primrec minus-spmat :: ( ′a::lattice-ring) spmat ⇒ ′a spmat

57

Page 58: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

whereminus-spmat [] = []| minus-spmat (a#as) = (fst a, minus-spvec (snd a))#(minus-spmat as)

lemma sparse-row-matrix-minus:sparse-row-matrix (minus-spmat A) = − (sparse-row-matrix A)apply (induct A)apply (simp-all add : sparse-row-vector-minus sparse-row-matrix-cons)apply (subst Rep-matrix-inject [symmetric])apply (rule ext)+apply simpdone

lemma Rep-sparse-row-vector-zero: x 6= 0 =⇒ Rep-matrix (sparse-row-vector v)x y = 0proof −

assume x :x 6= 0have r :nrows (sparse-row-vector v) <= Suc 0 by (rule nrows-spvec)show ?thesis

apply (rule nrows)apply (subgoal-tac Suc 0 <= x )apply (insert r)apply (simp only :)apply (insert x )apply arithdone

qed

lemma sparse-row-matrix-abs:sorted-spvec A =⇒ sorted-spmat A =⇒ sparse-row-matrix (abs-spmat A) =

|sparse-row-matrix A|apply (induct A)apply (simp-all add : sparse-row-vector-abs sparse-row-matrix-cons)apply (frule-tac sorted-spvec-cons1 , simp)apply (simplesubst Rep-matrix-inject [symmetric])apply (rule ext)+apply autoapply (case-tac x=a)apply (simp)apply (simplesubst sorted-sparse-row-matrix-zero)apply autoapply (simplesubst Rep-sparse-row-vector-zero)apply simp-alldone

lemma sorted-spvec-minus-spmat : sorted-spvec A =⇒ sorted-spvec (minus-spmatA)

apply (induct A)apply (simp)

58

Page 59: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

apply (frule sorted-spvec-cons1 , simp)apply (simp add : sorted-spvec.simps split :list .split-asm)done

lemma sorted-spvec-abs-spmat : sorted-spvec A =⇒ sorted-spvec (abs-spmat A)apply (induct A)apply (simp)apply (frule sorted-spvec-cons1 , simp)apply (simp add : sorted-spvec.simps split :list .split-asm)done

lemma sorted-spmat-minus-spmat : sorted-spmat A =⇒ sorted-spmat (minus-spmatA)

apply (induct A)apply (simp-all add : sorted-spvec-minus-spvec)done

lemma sorted-spmat-abs-spmat : sorted-spmat A =⇒ sorted-spmat (abs-spmat A)apply (induct A)apply (simp-all add : sorted-spvec-abs-spvec)done

definition diff-spmat :: ( ′a::lattice-ring) spmat ⇒ ′a spmat ⇒ ′a spmatwhere diff-spmat A B = add-spmat A (minus-spmat B)

lemma sorted-spmat-diff-spmat : sorted-spmat A =⇒ sorted-spmat B =⇒ sorted-spmat(diff-spmat A B)by (simp add : diff-spmat-def sorted-spmat-minus-spmat sorted-spmat-add-spmat)

lemma sorted-spvec-diff-spmat : sorted-spvec A =⇒ sorted-spvec B =⇒ sorted-spvec(diff-spmat A B)

by (simp add : diff-spmat-def sorted-spvec-minus-spmat sorted-spvec-add-spmat)

lemma sparse-row-diff-spmat : sparse-row-matrix (diff-spmat A B ) = (sparse-row-matrixA) − (sparse-row-matrix B)

by (simp add : diff-spmat-def sparse-row-add-spmat sparse-row-matrix-minus)

definition sorted-sparse-matrix :: ′a spmat ⇒ boolwhere sorted-sparse-matrix A ←→ sorted-spvec A & sorted-spmat A

lemma sorted-sparse-matrix-imp-spvec: sorted-sparse-matrix A =⇒ sorted-spvec Aby (simp add : sorted-sparse-matrix-def )

lemma sorted-sparse-matrix-imp-spmat : sorted-sparse-matrix A =⇒ sorted-spmatA

by (simp add : sorted-sparse-matrix-def )

lemmas sorted-sp-simps =sorted-spvec.simps

59

Page 60: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

sorted-spmat .simpssorted-sparse-matrix-def

lemma bool1 : (¬ True) = False by blastlemma bool2 : (¬ False) = True by blastlemma bool3 : ((P ::bool) ∧ True) = P by blastlemma bool4 : (True ∧ (P ::bool)) = P by blastlemma bool5 : ((P ::bool) ∧ False) = False by blastlemma bool6 : (False ∧ (P ::bool)) = False by blastlemma bool7 : ((P ::bool) ∨ True) = True by blastlemma bool8 : (True ∨ (P ::bool)) = True by blastlemma bool9 : ((P ::bool) ∨ False) = P by blastlemma bool10 : (False ∨ (P ::bool)) = P by blastlemmas boolarith = bool1 bool2 bool3 bool4 bool5 bool6 bool7 bool8 bool9 bool10

lemma if-case-eq : (if b then x else y) = (case b of True => x | False => y) bysimp

primrec pprt-spvec :: ( ′a::{lattice-ab-group-add}) spvec ⇒ ′a spvecwhere

pprt-spvec [] = []| pprt-spvec (a#as) = (fst a, pprt (snd a)) # (pprt-spvec as)

primrec nprt-spvec :: ( ′a::{lattice-ab-group-add}) spvec ⇒ ′a spvecwhere

nprt-spvec [] = []| nprt-spvec (a#as) = (fst a, nprt (snd a)) # (nprt-spvec as)

primrec pprt-spmat :: ( ′a::{lattice-ab-group-add}) spmat ⇒ ′a spmatwhere

pprt-spmat [] = []| pprt-spmat (a#as) = (fst a, pprt-spvec (snd a))#(pprt-spmat as)

primrec nprt-spmat :: ( ′a::{lattice-ab-group-add}) spmat ⇒ ′a spmatwhere

nprt-spmat [] = []| nprt-spmat (a#as) = (fst a, nprt-spvec (snd a))#(nprt-spmat as)

lemma pprt-add : disj-matrices A (B ::(-::lattice-ring) matrix ) =⇒ pprt (A+B) =pprt A + pprt B

apply (simp add : pprt-def sup-matrix-def )apply (simp add : Rep-matrix-inject [symmetric])apply (rule ext)+apply simpapply (case-tac Rep-matrix A x xa 6= 0 )apply (simp-all add : disj-matrices-contr1 )done

60

Page 61: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

lemma nprt-add : disj-matrices A (B ::(-::lattice-ring) matrix ) =⇒ nprt (A+B) =nprt A + nprt B

apply (simp add : nprt-def inf-matrix-def )apply (simp add : Rep-matrix-inject [symmetric])apply (rule ext)+apply simpapply (case-tac Rep-matrix A x xa 6= 0 )apply (simp-all add : disj-matrices-contr1 )done

lemma pprt-singleton[simp]: pprt (singleton-matrix j i (x ::-::lattice-ring)) = singleton-matrixj i (pprt x )

apply (simp add : pprt-def sup-matrix-def )apply (simp add : Rep-matrix-inject [symmetric])apply (rule ext)+apply simpdone

lemma nprt-singleton[simp]: nprt (singleton-matrix j i (x ::-::lattice-ring)) = singleton-matrixj i (nprt x )

apply (simp add : nprt-def inf-matrix-def )apply (simp add : Rep-matrix-inject [symmetric])apply (rule ext)+apply simpdone

lemma less-imp-le: a < b =⇒ a <= (b::-::order) by (simp add : less-def )

lemma sparse-row-vector-pprt : sorted-spvec (v :: ′a::lattice-ring spvec) =⇒ sparse-row-vector(pprt-spvec v) = pprt (sparse-row-vector v)

apply (induct v)apply (simp-all)apply (frule sorted-spvec-cons1 , auto)apply (subst pprt-add)apply (subst disj-matrices-commute)apply (rule disj-sparse-row-singleton)apply autodone

lemma sparse-row-vector-nprt : sorted-spvec (v :: ′a::lattice-ring spvec) =⇒ sparse-row-vector(nprt-spvec v) = nprt (sparse-row-vector v)

apply (induct v)apply (simp-all)apply (frule sorted-spvec-cons1 , auto)apply (subst nprt-add)apply (subst disj-matrices-commute)apply (rule disj-sparse-row-singleton)apply autodone

61

Page 62: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

lemma pprt-move-matrix : pprt (move-matrix (A::( ′a::lattice-ring) matrix ) j i) =move-matrix (pprt A) j i

apply (simp add : pprt-def )apply (simp add : sup-matrix-def )apply (simp add : Rep-matrix-inject [symmetric])apply (rule ext)+apply (simp)done

lemma nprt-move-matrix : nprt (move-matrix (A::( ′a::lattice-ring) matrix ) j i) =move-matrix (nprt A) j i

apply (simp add : nprt-def )apply (simp add : inf-matrix-def )apply (simp add : Rep-matrix-inject [symmetric])apply (rule ext)+apply (simp)done

lemma sparse-row-matrix-pprt : sorted-spvec (m :: ′a::lattice-ring spmat) =⇒ sorted-spmatm =⇒ sparse-row-matrix (pprt-spmat m) = pprt (sparse-row-matrix m)

apply (induct m)apply simpapply simpapply (frule sorted-spvec-cons1 )apply (simp add : sparse-row-matrix-cons sparse-row-vector-pprt)apply (subst pprt-add)apply (subst disj-matrices-commute)apply (rule disj-move-sparse-vec-mat)apply autoapply (simp add : sorted-spvec.simps)apply (simp split : list .split)apply autoapply (simp add : pprt-move-matrix )done

lemma sparse-row-matrix-nprt : sorted-spvec (m :: ′a::lattice-ring spmat) =⇒ sorted-spmatm =⇒ sparse-row-matrix (nprt-spmat m) = nprt (sparse-row-matrix m)

apply (induct m)apply simpapply simpapply (frule sorted-spvec-cons1 )apply (simp add : sparse-row-matrix-cons sparse-row-vector-nprt)apply (subst nprt-add)apply (subst disj-matrices-commute)apply (rule disj-move-sparse-vec-mat)apply autoapply (simp add : sorted-spvec.simps)

62

Page 63: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

apply (simp split : list .split)apply autoapply (simp add : nprt-move-matrix )done

lemma sorted-pprt-spvec: sorted-spvec v =⇒ sorted-spvec (pprt-spvec v)apply (induct v)apply (simp)apply (frule sorted-spvec-cons1 )apply simpapply (simp add : sorted-spvec.simps split :list .split-asm)done

lemma sorted-nprt-spvec: sorted-spvec v =⇒ sorted-spvec (nprt-spvec v)apply (induct v)apply (simp)apply (frule sorted-spvec-cons1 )apply simpapply (simp add : sorted-spvec.simps split :list .split-asm)done

lemma sorted-spvec-pprt-spmat : sorted-spvec m =⇒ sorted-spvec (pprt-spmat m)apply (induct m)apply (simp)apply (frule sorted-spvec-cons1 )apply simpapply (simp add : sorted-spvec.simps split :list .split-asm)done

lemma sorted-spvec-nprt-spmat : sorted-spvec m =⇒ sorted-spvec (nprt-spmat m)apply (induct m)apply (simp)apply (frule sorted-spvec-cons1 )apply simpapply (simp add : sorted-spvec.simps split :list .split-asm)done

lemma sorted-spmat-pprt-spmat : sorted-spmat m =⇒ sorted-spmat (pprt-spmatm)

apply (induct m)apply (simp-all add : sorted-pprt-spvec)done

lemma sorted-spmat-nprt-spmat : sorted-spmat m =⇒ sorted-spmat (nprt-spmatm)

apply (induct m)apply (simp-all add : sorted-nprt-spvec)done

63

Page 64: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

definition mult-est-spmat :: ( ′a::lattice-ring) spmat ⇒ ′a spmat ⇒ ′a spmat ⇒′a spmat ⇒ ′a spmat where

mult-est-spmat r1 r2 s1 s2 =add-spmat (mult-spmat (pprt-spmat s2 ) (pprt-spmat r2 )) (add-spmat (mult-spmat

(pprt-spmat s1 ) (nprt-spmat r2 ))(add-spmat (mult-spmat (nprt-spmat s2 ) (pprt-spmat r1 )) (mult-spmat (nprt-spmat

s1 ) (nprt-spmat r1 ))))

lemmas sparse-row-matrix-op-simps =sorted-sparse-matrix-imp-spmat sorted-sparse-matrix-imp-spvecsparse-row-add-spmat sorted-spvec-add-spmat sorted-spmat-add-spmatsparse-row-diff-spmat sorted-spvec-diff-spmat sorted-spmat-diff-spmatsparse-row-matrix-minus sorted-spvec-minus-spmat sorted-spmat-minus-spmatsparse-row-mult-spmat sorted-spvec-mult-spmat sorted-spmat-mult-spmatsparse-row-matrix-abs sorted-spvec-abs-spmat sorted-spmat-abs-spmatle-spmat-iff-sparse-row-lesparse-row-matrix-pprt sorted-spvec-pprt-spmat sorted-spmat-pprt-spmatsparse-row-matrix-nprt sorted-spvec-nprt-spmat sorted-spmat-nprt-spmat

lemmas sparse-row-matrix-arith-simps =mult-spmat .simps mult-spvec-spmat .simpsaddmult-spvec.simpssmult-spvec-empty smult-spvec-consadd-spmat .simps add-spvec.simpsminus-spmat .simps minus-spvec.simpsabs-spmat .simps abs-spvec.simpsdiff-spmat-defle-spmat .simps le-spvec.simpspprt-spmat .simps pprt-spvec.simpsnprt-spmat .simps nprt-spvec.simpsmult-est-spmat-def

end

theory LPimports Main HOL−Library .Lattice-Algebrasbegin

lemma le-add-right-mono:assumesa <= b + (c:: ′a::ordered-ab-group-add)c <= dshows a <= b + dapply (rule-tac order-trans[where y = b+c])apply (simp-all add : assms)

64

Page 65: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

done

lemma linprog-dual-estimate:assumesA ∗ x ≤ (b:: ′a::lattice-ring)0 ≤ y|A − A ′| ≤ δ-Ab ≤ b ′

|c − c ′| ≤ δ-c|x | ≤ rshowsc ∗ x ≤ y ∗ b ′ + (y ∗ δ-A + |y ∗ A ′ − c ′| + δ-c) ∗ r

proof −from assms have 1 : y ∗ b <= y ∗ b ′ by (simp add : mult-left-mono)from assms have 2 : y ∗ (A ∗ x ) <= y ∗ b by (simp add : mult-left-mono)have 3 : y ∗ (A ∗ x ) = c ∗ x + (y ∗ (A − A ′) + (y ∗ A ′ − c ′) + (c ′−c)) ∗ x

by (simp add : algebra-simps)from 1 2 3 have 4 : c ∗ x + (y ∗ (A − A ′) + (y ∗ A ′ − c ′) + (c ′−c)) ∗ x <=

y ∗ b ′ by simphave 5 : c ∗ x <= y ∗ b ′ + |(y ∗ (A − A ′) + (y ∗ A ′ − c ′) + (c ′−c)) ∗ x |

by (simp only : 4 estimate-by-abs)have 6 : |(y ∗ (A − A ′) + (y ∗ A ′ − c ′) + (c ′−c)) ∗ x | <= |y ∗ (A − A ′) + (y∗ A ′ − c ′) + (c ′−c)| ∗ |x |

by (simp add : abs-le-mult)have 7 : (|y ∗ (A − A ′) + (y ∗ A ′ − c ′) + (c ′−c)|) ∗ |x | <= (|y ∗ (A−A ′) +

(y∗A ′−c ′)| + |c ′ − c|) ∗ |x |by(rule abs-triangle-ineq [THEN mult-right-mono]) simp

have 8 : (|y ∗ (A−A ′) + (y∗A ′−c ′)| + |c ′ − c|) ∗ |x | <= (|y ∗ (A−A ′)| +|y∗A ′−c ′| + |c ′ − c|) ∗ |x |

by (simp add : abs-triangle-ineq mult-right-mono)have 9 : (|y ∗ (A−A ′)| + |y∗A ′−c ′| + |c ′−c|) ∗ |x | <= (|y | ∗ |A−A ′| + |y∗A ′−c ′|

+ |c ′−c|) ∗ |x |by (simp add : abs-le-mult mult-right-mono)

have 10 : c ′−c = −(c−c ′) by (simp add : algebra-simps)have 11 : |c ′−c| = |c−c ′|

by (subst 10 , subst abs-minus-cancel , simp)have 12 : (|y | ∗ |A−A ′| + |y∗A ′−c ′| + |c ′−c|) ∗ |x | <= (|y | ∗ |A−A ′| + |y∗A ′−c ′|

+ δ-c) ∗ |x |by (simp add : 11 assms mult-right-mono)

have 13 : (|y | ∗ |A−A ′| + |y∗A ′−c ′| + δ-c) ∗ |x | <= (|y | ∗ δ-A + |y∗A ′−c ′| +δ-c) ∗ |x |

by (simp add : assms mult-right-mono mult-left-mono)have r : (|y | ∗ δ-A + |y∗A ′−c ′| + δ-c) ∗ |x | <= (|y | ∗ δ-A + |y∗A ′−c ′| + δ-c)∗ r

apply (rule mult-left-mono)apply (simp add : assms)apply (rule-tac add-mono[of 0 :: ′a - 0 , simplified ])+apply (rule mult-left-mono[of 0 δ-A, simplified ])apply (simp-all)

65

Page 66: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

apply (rule order-trans[where y=|A−A ′|], simp-all add : assms)apply (rule order-trans[where y=|c−c ′|], simp-all add : assms)done

from 6 7 8 9 12 13 r have 14 : |(y ∗ (A − A ′) + (y ∗ A ′ − c ′) + (c ′−c)) ∗ x |<= (|y | ∗ δ-A + |y∗A ′−c ′| + δ-c) ∗ r

by (simp)show ?thesis

apply (rule le-add-right-mono[of - - |(y ∗ (A − A ′) + (y ∗ A ′ − c ′) + (c ′−c))∗ x |])

apply (simp-all only : 5 14 [simplified abs-of-nonneg [of y , simplified assms]])done

qed

lemma le-ge-imp-abs-diff-1 :assumesA1 <= (A:: ′a::lattice-ring)A <= A2shows |A−A1 | <= A2−A1

proof −have 0 <= A − A1proof −

from assms add-right-mono [of A1 A − A1 ] show ?thesis by simpqedthen have |A−A1 | = A−A1 by (rule abs-of-nonneg)with assms show |A−A1 | <= (A2−A1 ) by simp

qed

lemma mult-le-prts:assumesa1 <= (a:: ′a::lattice-ring)a <= a2b1 <= bb <= b2showsa ∗ b <= pprt a2 ∗ pprt b2 + pprt a1 ∗ nprt b2 + nprt a2 ∗ pprt b1 + nprt a1∗ nprt b1proof −

have a ∗ b = (pprt a + nprt a) ∗ (pprt b + nprt b)apply (subst prts[symmetric])+apply simpdone

then have a ∗ b = pprt a ∗ pprt b + pprt a ∗ nprt b + nprt a ∗ pprt b + nprta ∗ nprt b

by (simp add : algebra-simps)moreover have pprt a ∗ pprt b <= pprt a2 ∗ pprt b2

by (simp-all add : assms mult-mono)moreover have pprt a ∗ nprt b <= pprt a1 ∗ nprt b2proof −

have pprt a ∗ nprt b <= pprt a ∗ nprt b2

66

Page 67: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

by (simp add : mult-left-mono assms)moreover have pprt a ∗ nprt b2 <= pprt a1 ∗ nprt b2

by (simp add : mult-right-mono-neg assms)ultimately show ?thesis

by simpqedmoreover have nprt a ∗ pprt b <= nprt a2 ∗ pprt b1proof −

have nprt a ∗ pprt b <= nprt a2 ∗ pprt bby (simp add : mult-right-mono assms)

moreover have nprt a2 ∗ pprt b <= nprt a2 ∗ pprt b1by (simp add : mult-left-mono-neg assms)

ultimately show ?thesisby simp

qedmoreover have nprt a ∗ nprt b <= nprt a1 ∗ nprt b1proof −

have nprt a ∗ nprt b <= nprt a ∗ nprt b1by (simp add : mult-left-mono-neg assms)

moreover have nprt a ∗ nprt b1 <= nprt a1 ∗ nprt b1by (simp add : mult-right-mono-neg assms)

ultimately show ?thesisby simp

qedultimately show ?thesis

by − (rule add-mono | simp)+qed

lemma mult-le-dual-prts:assumesA ∗ x ≤ (b:: ′a::lattice-ring)0 ≤ yA1 ≤ AA ≤ A2c1 ≤ cc ≤ c2r1 ≤ xx ≤ r2showsc ∗ x ≤ y ∗ b + (let s1 = c1 − y ∗ A2 ; s2 = c2 − y ∗ A1 in pprt s2 ∗ pprt r2

+ pprt s1 ∗ nprt r2 + nprt s2 ∗ pprt r1 + nprt s1 ∗ nprt r1 )(is - <= - + ?C )

proof −from assms have y ∗ (A ∗ x ) <= y ∗ b by (simp add : mult-left-mono)moreover have y ∗ (A ∗ x ) = c ∗ x + (y ∗ A − c) ∗ x by (simp add :

algebra-simps)ultimately have c ∗ x + (y ∗ A − c) ∗ x <= y ∗ b by simpthen have c ∗ x <= y ∗ b − (y ∗ A − c) ∗ x by (simp add : le-diff-eq)then have cx : c ∗ x <= y ∗ b + (c − y ∗ A) ∗ x by (simp add : algebra-simps)

67

Page 68: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

have s2 : c − y ∗ A <= c2 − y ∗ A1by (simp add : assms add-mono mult-left-mono algebra-simps)

have s1 : c1 − y ∗ A2 <= c − y ∗ Aby (simp add : assms add-mono mult-left-mono algebra-simps)

have prts: (c − y ∗ A) ∗ x <= ?Capply (simp add : Let-def )apply (rule mult-le-prts)apply (simp-all add : assms s1 s2 )done

then have y ∗ b + (c − y ∗ A) ∗ x <= y ∗ b + ?Cby simp

with cx show ?thesisby(simp only :)

qed

end

1 Floating Point Representation of the Reals

theory ComputeFloatimports Complex-Main HOL−Library .Lattice-Algebrasbegin

ML-file 〈∼∼/src/Tools/float .ML〉

definition int-of-real :: real ⇒ intwhere int-of-real x = (SOME y . real-of-int y = x )

definition real-is-int :: real ⇒ boolwhere real-is-int x = (∃ (u::int). x = real-of-int u)

lemma real-is-int-def2 : real-is-int x = (x = real-of-int (int-of-real x ))by (auto simp add : real-is-int-def int-of-real-def )

lemma real-is-int-real [simp]: real-is-int (real-of-int (x ::int))by (auto simp add : real-is-int-def int-of-real-def )

lemma int-of-real-real [simp]: int-of-real (real-of-int x ) = xby (simp add : int-of-real-def )

lemma real-int-of-real [simp]: real-is-int x =⇒ real-of-int (int-of-real x ) = xby (auto simp add : int-of-real-def real-is-int-def )

lemma real-is-int-add-int-of-real : real-is-int a =⇒ real-is-int b =⇒ (int-of-real(a+b)) = (int-of-real a) + (int-of-real b)by (auto simp add : int-of-real-def real-is-int-def )

68

Page 69: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

lemma real-is-int-add [simp]: real-is-int a =⇒ real-is-int b =⇒ real-is-int (a+b)apply (subst real-is-int-def2 )apply (simp add : real-is-int-add-int-of-real real-int-of-real)done

lemma int-of-real-sub: real-is-int a =⇒ real-is-int b =⇒ (int-of-real (a−b)) =(int-of-real a) − (int-of-real b)by (auto simp add : int-of-real-def real-is-int-def )

lemma real-is-int-sub[simp]: real-is-int a =⇒ real-is-int b =⇒ real-is-int (a−b)apply (subst real-is-int-def2 )apply (simp add : int-of-real-sub real-int-of-real)done

lemma real-is-int-rep: real-is-int x =⇒ ∃ !(a::int). real-of-int a = xby (auto simp add : real-is-int-def )

lemma int-of-real-mult :assumes real-is-int a real-is-int bshows (int-of-real (a∗b)) = (int-of-real a) ∗ (int-of-real b)using assmsby (auto simp add : real-is-int-def of-int-mult [symmetric]

simp del : of-int-mult)

lemma real-is-int-mult [simp]: real-is-int a =⇒ real-is-int b =⇒ real-is-int (a∗b)apply (subst real-is-int-def2 )apply (simp add : int-of-real-mult)done

lemma real-is-int-0 [simp]: real-is-int (0 ::real)by (simp add : real-is-int-def int-of-real-def )

lemma real-is-int-1 [simp]: real-is-int (1 ::real)proof −

have real-is-int (1 ::real) = real-is-int(real-of-int (1 ::int)) by autoalso have . . . = True by (simp only : real-is-int-real)ultimately show ?thesis by auto

qed

lemma real-is-int-n1 : real-is-int (−1 ::real)proof −

have real-is-int (−1 ::real) = real-is-int(real-of-int (−1 ::int)) by autoalso have . . . = True by (simp only : real-is-int-real)ultimately show ?thesis by auto

qed

lemma real-is-int-numeral [simp]: real-is-int (numeral x )by (auto simp: real-is-int-def intro!: exI [of - numeral x ])

69

Page 70: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

lemma real-is-int-neg-numeral [simp]: real-is-int (− numeral x )by (auto simp: real-is-int-def intro!: exI [of - − numeral x ])

lemma int-of-real-0 [simp]: int-of-real (0 ::real) = (0 ::int)by (simp add : int-of-real-def )

lemma int-of-real-1 [simp]: int-of-real (1 ::real) = (1 ::int)proof −

have 1 : (1 ::real) = real-of-int (1 ::int) by autoshow ?thesis by (simp only : 1 int-of-real-real)

qed

lemma int-of-real-numeral [simp]: int-of-real (numeral b) = numeral bunfolding int-of-real-def by simp

lemma int-of-real-neg-numeral [simp]: int-of-real (− numeral b) = − numeral bunfolding int-of-real-defby (metis int-of-real-def int-of-real-real of-int-minus of-int-of-nat-eq of-nat-numeral)

lemma int-div-zdiv : int (a div b) = (int a) div (int b)by (rule zdiv-int)

lemma int-mod-zmod : int (a mod b) = (int a) mod (int b)by (rule zmod-int)

lemma abs-div-2-less: a 6= 0 =⇒ a 6= −1 =⇒ |(a::int) div 2 | < |a|by arith

lemma norm-0-1 : (1 ::-::numeral) = Numeral1by auto

lemma add-left-zero: 0 + a = (a:: ′a::comm-monoid-add)by simp

lemma add-right-zero: a + 0 = (a:: ′a::comm-monoid-add)by simp

lemma mult-left-one: 1 ∗ a = (a:: ′a::semiring-1 )by simp

lemma mult-right-one: a ∗ 1 = (a:: ′a::semiring-1 )by simp

lemma int-pow-0 : (a::int)ˆ0 = 1by simp

lemma int-pow-1 : (a::int)ˆ(Numeral1 ) = aby simp

70

Page 71: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

lemma one-eq-Numeral1-nring : (1 :: ′a::numeral) = Numeral1by simp

lemma one-eq-Numeral1-nat : (1 ::nat) = Numeral1by simp

lemma zpower-Pls: (z ::int)ˆ0 = Numeral1by simp

lemma fst-cong : a=a ′ =⇒ fst (a,b) = fst (a ′,b)by simp

lemma snd-cong : b=b ′ =⇒ snd (a,b) = snd (a,b ′)by simp

lemma lift-bool : x =⇒ x=Trueby simp

lemma nlift-bool : ∼x =⇒ x=Falseby simp

lemma not-false-eq-true: (∼ False) = True by simp

lemma not-true-eq-false: (∼ True) = False by simp

lemmas powerarith = nat-numeral power-numeral-evenpower-numeral-odd zpower-Pls

definition float :: (int × int) ⇒ real wherefloat = (λ(a, b). real-of-int a ∗ 2 powr real-of-int b)

lemma float-add-l0 : float (0 , e) + x = xby (simp add : float-def )

lemma float-add-r0 : x + float (0 , e) = xby (simp add : float-def )

lemma float-add :float (a1 , e1 ) + float (a2 , e2 ) =(if e1<=e2 then float (a1 +a2∗2ˆ(nat(e2−e1 )), e1 ) else float (a1∗2ˆ(nat (e1−e2 ))+a2 ,

e2 ))by (simp add : float-def algebra-simps powr-realpow [symmetric] powr-diff )

lemma float-mult-l0 : float (0 , e) ∗ x = float (0 , 0 )by (simp add : float-def )

lemma float-mult-r0 : x ∗ float (0 , e) = float (0 , 0 )by (simp add : float-def )

71

Page 72: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

lemma float-mult :float (a1 , e1 ) ∗ float (a2 , e2 ) = (float (a1 ∗ a2 , e1 + e2 ))by (simp add : float-def powr-add)

lemma float-minus:− (float (a,b)) = float (−a, b)by (simp add : float-def )

lemma zero-le-float :(0 <= float (a,b)) = (0 <= a)by (simp add : float-def zero-le-mult-iff )

lemma float-le-zero:(float (a,b) <= 0 ) = (a <= 0 )by (simp add : float-def mult-le-0-iff )

lemma float-abs:|float (a,b)| = (if 0 <= a then (float (a,b)) else (float (−a,b)))by (simp add : float-def abs-if mult-less-0-iff not-less)

lemma float-zero:float (0 , b) = 0by (simp add : float-def )

lemma float-pprt :pprt (float (a, b)) = (if 0 <= a then (float (a,b)) else (float (0 , b)))by (auto simp add : zero-le-float float-le-zero float-zero)

lemma float-nprt :nprt (float (a, b)) = (if 0 <= a then (float (0 ,b)) else (float (a, b)))by (auto simp add : zero-le-float float-le-zero float-zero)

definition lbound :: real ⇒ realwhere lbound x = min 0 x

definition ubound :: real ⇒ realwhere ubound x = max 0 x

lemma lbound : lbound x ≤ xby (simp add : lbound-def )

lemma ubound : x ≤ ubound xby (simp add : ubound-def )

lemma pprt-lbound : pprt (lbound x ) = float (0 , 0 )by (auto simp: float-def lbound-def )

lemma nprt-ubound : nprt (ubound x ) = float (0 , 0 )

72

Page 73: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

by (auto simp: float-def ubound-def )

lemmas floatarith[simplified norm-0-1 ] = float-add float-add-l0 float-add-r0 float-multfloat-mult-l0 float-mult-r0

float-minus float-abs zero-le-float float-pprt float-nprt pprt-lbound nprt-ubound

lemmas arith = arith-simps rel-simps diff-nat-numeral nat-0nat-neg-numeral powerarith floatarith not-false-eq-true not-true-eq-false

ML-file 〈float-arith.ML〉

end

theory Compute-Oracle imports HOL.HOLbegin

ML-file 〈am.ML〉

ML-file 〈am-compiler .ML〉

ML-file 〈am-interpreter .ML〉

ML-file 〈am-ghc.ML〉

ML-file 〈am-sml .ML〉

ML-file 〈report .ML〉

ML-file 〈compute.ML〉

ML-file 〈linker .ML〉

endtheory ComputeHOLimports Complex-Main Compute-Oracle/Compute-Oraclebegin

lemma Trueprop-eq-eq : Trueprop X == (X == True) by (simp add : atomize-eq)lemma meta-eq-trivial : x == y =⇒ x == y by simplemma meta-eq-imp-eq : x == y =⇒ x = y by autolemma eq-trivial : x = y =⇒ x = y by autolemma bool-to-true: x :: bool =⇒ x == True by simplemma transmeta-1 : x = y =⇒ y == z =⇒ x = z by simplemma transmeta-2 : x == y =⇒ y = z =⇒ x = z by simplemma transmeta-3 : x == y =⇒ y == z =⇒ x = z by simp

lemma If-True: If True = (λ x y . x ) by ((rule ext)+,auto)lemma If-False: If False = (λ x y . y) by ((rule ext)+, auto)

lemmas compute-if = If-True If-False

73

Page 74: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

lemma bool1 : (¬ True) = False by blastlemma bool2 : (¬ False) = True by blastlemma bool3 : (P ∧ True) = P by blastlemma bool4 : (True ∧ P) = P by blastlemma bool5 : (P ∧ False) = False by blastlemma bool6 : (False ∧ P) = False by blastlemma bool7 : (P ∨ True) = True by blastlemma bool8 : (True ∨ P) = True by blastlemma bool9 : (P ∨ False) = P by blastlemma bool10 : (False ∨ P) = P by blastlemma bool11 : (True −→ P) = P by blastlemma bool12 : (P −→ True) = True by blastlemma bool13 : (True −→ P) = P by blastlemma bool14 : (P −→ False) = (¬ P) by blastlemma bool15 : (False −→ P) = True by blastlemma bool16 : (False = False) = True by blastlemma bool17 : (True = True) = True by blastlemma bool18 : (False = True) = False by blastlemma bool19 : (True = False) = False by blast

lemmas compute-bool = bool1 bool2 bool3 bool4 bool5 bool6 bool7 bool8 bool9 bool10bool11 bool12 bool13 bool14 bool15 bool16 bool17 bool18 bool19

lemma compute-fst : fst (x ,y) = x by simplemma compute-snd : snd (x ,y) = y by simplemma compute-pair-eq : ((a, b) = (c, d)) = (a = c ∧ b = d) by auto

lemma case-prod-simp: case-prod f (x ,y) = f x y by simp

lemmas compute-pair = compute-fst compute-snd compute-pair-eq case-prod-simp

lemma compute-the: the (Some x ) = x by simplemma compute-None-Some-eq : (None = Some x ) = False by autolemma compute-Some-None-eq : (Some x = None) = False by autolemma compute-None-None-eq : (None = None) = True by autolemma compute-Some-Some-eq : (Some x = Some y) = (x = y) by auto

definition case-option-compute :: ′b option ⇒ ′a ⇒ ( ′b ⇒ ′a) ⇒ ′awhere case-option-compute opt a f = case-option a f opt

lemma case-option-compute: case-option = (λ a f opt . case-option-compute opt af )

74

Page 75: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

by (simp add : case-option-compute-def )

lemma case-option-compute-None: case-option-compute None = (λ a f . a)apply (rule ext)+apply (simp add : case-option-compute-def )done

lemma case-option-compute-Some: case-option-compute (Some x ) = (λ a f . f x )apply (rule ext)+apply (simp add : case-option-compute-def )done

lemmas compute-case-option = case-option-compute case-option-compute-None case-option-compute-Some

lemmas compute-option = compute-the compute-None-Some-eq compute-Some-None-eqcompute-None-None-eq compute-Some-Some-eq compute-case-option

lemma length-cons:length (x#xs) = 1 + (length xs)by simp

lemma length-nil : length [] = 0by simp

lemmas compute-list-length = length-nil length-cons

definition case-list-compute :: ′b list ⇒ ′a ⇒ ( ′b ⇒ ′b list ⇒ ′a) ⇒ ′awhere case-list-compute l a f = case-list a f l

lemma case-list-compute: case-list = (λ (a:: ′a) f (l :: ′b list). case-list-compute l af )

apply (rule ext)+apply (simp add : case-list-compute-def )done

lemma case-list-compute-empty : case-list-compute ([]:: ′b list) = (λ (a:: ′a) f . a)apply (rule ext)+apply (simp add : case-list-compute-def )done

lemma case-list-compute-cons: case-list-compute (u#v) = (λ (a:: ′a) f . (f (u:: ′b)v))

apply (rule ext)+apply (simp add : case-list-compute-def )done

75

Page 76: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

lemmas compute-case-list = case-list-compute case-list-compute-empty case-list-compute-cons

lemma compute-list-nth: ((x#xs) ! n) = (if n = 0 then x else (xs ! (n − 1 )))by (cases n, auto)

lemmas compute-list = compute-case-list compute-list-length compute-list-nth

lemmas compute-let = Let-def

lemmas compute-hol = compute-if compute-bool compute-pair compute-option compute-listcompute-let

ML 〈

signature ComputeHOL =sig

val prep-thms : thm list −> thm listval to-meta-eq : thm −> thmval to-hol-eq : thm −> thmval symmetric : thm −> thmval trans : thm −> thm −> thm

end

structure ComputeHOL : ComputeHOL =struct

localfun lhs-of eq = fst (Thm.dest-equals (Thm.cprop-of eq));infun rewrite-conv [] ct = raise CTERM (rewrite-conv , [ct ])| rewrite-conv (eq :: eqs) ct =

Thm.instantiate (Thm.match (lhs-of eq , ct)) eqhandle Pattern.MATCH => rewrite-conv eqs ct ;

end

val convert-conditions = Conv .fconv-rule (Conv .prems-conv ∼1 (Conv .try-conv(rewrite-conv [@{thm Trueprop-eq-eq}])))

val eq-th = @{thm HOL.eq-reflection}

76

Page 77: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

val meta-eq-trivial = @{thm ComputeHOL.meta-eq-trivial}val bool-to-true = @{thm ComputeHOL.bool-to-true}

fun to-meta-eq th = eq-th OF [th] handle THM - => meta-eq-trivial OF [th] handleTHM - => bool-to-true OF [th]

fun to-hol-eq th = @{thm meta-eq-imp-eq} OF [th] handle THM - => @{thmeq-trivial} OF [th]

fun prep-thms ths = map (convert-conditions o to-meta-eq) ths

fun symmetric th = @{thm HOL.sym} OF [th] handle THM - => @{thm Pure.symmetric}OF [th]

localval trans-HOL = @{thm HOL.trans}val trans-HOL-1 = @{thm ComputeHOL.transmeta-1}val trans-HOL-2 = @{thm ComputeHOL.transmeta-2}val trans-HOL-3 = @{thm ComputeHOL.transmeta-3}fun tr [] th1 th2 = trans-HOL OF [th1 , th2 ]| tr (t ::ts) th1 th2 = (t OF [th1 , th2 ] handle THM - => tr ts th1 th2 )

infun trans th1 th2 = tr [trans-HOL, trans-HOL-1 , trans-HOL-2 , trans-HOL-3 ]

th1 th2end

end〉

endtheory ComputeNumeralimports ComputeHOL ComputeFloatbegin

lemmas biteq = eq-num-simps

lemmas bitless = less-num-simps

lemmas bitle = le-num-simps

lemmas bitadd = add-num-simps

lemmas bitmul = mult-num-simps

77

Page 78: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

lemmas bitarith = arith-simps

lemmas natnorm = one-eq-Numeral1-nat

fun natfac :: nat ⇒ natwhere natfac n = (if n = 0 then 1 else n ∗ (natfac (n − 1 )))

lemmas compute-natarith =arith-simps rel-simpsdiff-nat-numeral nat-numeral nat-0 nat-neg-numeralnumeral-One [symmetric]numeral-1-eq-Suc-0 [symmetric]Suc-numeral natfac.simps

lemmas number-norm = numeral-One[symmetric]

lemmas compute-numberarith =arith-simps rel-simps number-norm

lemmas compute-num-conversions =of-nat-numeral of-nat-0nat-numeral nat-0 nat-neg-numeralof-int-numeral of-int-neg-numeral of-int-0

lemmas zpowerarith = power-numeral-even power-numeral-odd zpower-Pls int-pow-1

lemmas compute-div-mod = div-0 mod-0 div-by-0 mod-by-0 div-by-1 mod-by-1one-div-numeral one-mod-numeral minus-one-div-numeral minus-one-mod-numeralone-div-minus-numeral one-mod-minus-numeralnumeral-div-numeral numeral-mod-numeral minus-numeral-div-numeral minus-numeral-mod-numeralnumeral-div-minus-numeral numeral-mod-minus-numeraldiv-minus-minus mod-minus-minus Divides.adjust-div-eq of-bool-eq one-neq-zeronumeral-neq-zero neg-equal-0-iff-equal arith-simps arith-special divmod-trivialdivmod-steps divmod-cancel divmod-step-eq fst-conv snd-conv numeral-Onecase-prod-beta rel-simps Divides.adjust-mod-def div-minus1-right mod-minus1-rightminus-minus numeral-times-numeral mult-zero-right mult-1-right

lemma even-0-int : even (0 ::int) = Trueby simp

lemma even-One-int : even (numeral Num.One :: int) = Falseby simp

78

Page 79: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

lemma even-Bit0-int : even (numeral (Num.Bit0 x ) :: int) = Trueby (simp only : even-numeral)

lemma even-Bit1-int : even (numeral (Num.Bit1 x ) :: int) = Falseby (simp only : odd-numeral)

lemmas compute-even = even-0-int even-One-int even-Bit0-int even-Bit1-int

lemmas compute-numeral = compute-if compute-let compute-pair compute-boolcompute-natarith compute-numberarith max-def min-def

compute-num-conversions zpowerarith compute-div-mod compute-even

end

theory Cpleximports SparseMatrix LP ComputeFloat ComputeNumeralbegin

ML-file 〈Cplex-tools.ML〉

ML-file 〈CplexMatrixConverter .ML〉

ML-file 〈FloatSparseMatrixBuilder .ML〉

ML-file 〈fspmlp.ML〉

lemma spm-mult-le-dual-prts:assumessorted-sparse-matrix A1sorted-sparse-matrix A2sorted-sparse-matrix c1sorted-sparse-matrix c2sorted-sparse-matrix ysorted-sparse-matrix r1sorted-sparse-matrix r2sorted-spvec ble-spmat [] ysparse-row-matrix A1 ≤ AA ≤ sparse-row-matrix A2sparse-row-matrix c1 ≤ cc ≤ sparse-row-matrix c2sparse-row-matrix r1 ≤ xx ≤ sparse-row-matrix r2A ∗ x ≤ sparse-row-matrix (b::( ′a::lattice-ring) spmat)showsc ∗ x ≤ sparse-row-matrix (add-spmat (mult-spmat y b)(let s1 = diff-spmat c1 (mult-spmat y A2 ); s2 = diff-spmat c2 (mult-spmat y

A1 ) inadd-spmat (mult-spmat (pprt-spmat s2 ) (pprt-spmat r2 )) (add-spmat (mult-spmat

(pprt-spmat s1 ) (nprt-spmat r2 ))(add-spmat (mult-spmat (nprt-spmat s2 ) (pprt-spmat r1 )) (mult-spmat (nprt-spmat

79

Page 80: Matrix - isabelle.in.tum.deisabelle.in.tum.de/library/HOL/HOL-Matrix_LP/document.pdfMatrix Steven Obua June 9, 2019 theory Matrix imports Main HOL Library:Lattice-Algebras begin type-synonym

s1 ) (nprt-spmat r1 ))))))apply (simp add : Let-def )apply (insert assms)apply (simp add : sparse-row-matrix-op-simps algebra-simps)apply (rule mult-le-dual-prts[where A=A, simplified Let-def algebra-simps])apply (auto)done

lemma spm-mult-le-dual-prts-no-let :assumessorted-sparse-matrix A1sorted-sparse-matrix A2sorted-sparse-matrix c1sorted-sparse-matrix c2sorted-sparse-matrix ysorted-sparse-matrix r1sorted-sparse-matrix r2sorted-spvec ble-spmat [] ysparse-row-matrix A1 ≤ AA ≤ sparse-row-matrix A2sparse-row-matrix c1 ≤ cc ≤ sparse-row-matrix c2sparse-row-matrix r1 ≤ xx ≤ sparse-row-matrix r2A ∗ x ≤ sparse-row-matrix (b::( ′a::lattice-ring) spmat)showsc ∗ x ≤ sparse-row-matrix (add-spmat (mult-spmat y b)(mult-est-spmat r1 r2 (diff-spmat c1 (mult-spmat y A2 )) (diff-spmat c2 (mult-spmat

y A1 ))))by (simp add : assms mult-est-spmat-def spm-mult-le-dual-prts[where A=A, sim-

plified Let-def ])

ML-file 〈matrixlp.ML〉

end

80