Amortized Complexity Verified

Size: px
Start display at page:

Download "Amortized Complexity Verified"

Transcription

1 Amortized Complexity Verified Tobias Nipkow August 28, 2014 Abstract A framework for the analysis of the amortized complexity of (functional) data structures is formalized in Isabelle/HOL and applied to a number of standard examples and to two non-trivial ones: skew heaps and splay trees. In the same spirit we show that the move-to-front algorithm for the list update problem is 2-competitive. A preliminary paper describing this research has been presented at the Isabelle Workshop Contents 1 Amortized Complexity Binary Counter Dynamic tables: insert only Stack with multipop Queue Dynamic tables: insert and delete Amortized Analysis of Skew Heaps Amortized Analysis Splay Tree Analysis Time Wlog in tree Analysis of splay Optimal Interpretation Overall analysis List Inversion 14 5 Online Algorithm Move-to-Front 17 1

2 theory Amor imports Complex Main begin 1 Amortized Complexity For the background and most of the examples see some algorithms textbook, for example, [1]. locale amor = fixes init :: s fixes nxt :: o s s fixes inv :: s bool fixes t :: o s real fixes Φ :: s real fixes U :: o s real assumes inv init: inv init assumes inv nxt: inv s = inv(nxt f s) assumes ppos: inv s = Φ s 0 assumes p0 : Φ init = 0 assumes U : inv s = t f s + Φ(nxt f s) Φ s U f s begin fun state :: (nat o) nat s where state f 0 = init state f (Suc n) = nxt (f n) (state f n) lemma inv state: inv(state f n) definition a :: (nat o) nat real where a f i = t (f i) (state f i) + Φ(state f (i+1 )) Φ(state f i) lemma aeq: ( i<n. t (f i) (state f i)) = ( i<n. a f i) Φ(state f n) corollary ta: ( i<n. t (f i) (state f i)) ( i<n. a f i) lemma aa1 : a f i U (f i) (state f i) lemma ub: ( i<n. t (f i) (state f i)) ( i<n. U (f i) (state f i)) 2

3 end 1.1 Binary Counter fun incr where incr [] = [True] incr (False#bs) = True # bs incr (True#bs) = False # incr bs fun t incr :: bool list real where t incr [] = 1 t incr (False#bs) = 1 t incr (True#bs) = t incr bs + 1 definition p incr :: bool list real (Φ incr ) where Φ incr bs = length(filter id bs) lemma a incr: t incr bs + Φ incr (incr bs) Φ incr bs = 2 interpretation incr: amor where init = [] and nxt = %. incr and inv = λ. True and t = λ. t incr and Φ = Φ incr and U = λ. 2 thm incr.ub 1.2 Dynamic tables: insert only fun t ins :: nat nat real where t ins (n,l) = (if n<l then 1 else n+1 ) interpretation ins: amor where init = (0 ::nat,0 ::nat) and nxt = λ (n,l). (n+1, if n<l then l else if l=0 then 1 else 2 l) and inv = λ(n,l). if l=0 then n=0 else n l l < 2 n and t = λ. t ins and Φ = λ(n,l). 2 n l and U = λ. 3 locale table insert = fixes a :: real fixes c :: real assumes c1 [arith]: c > 1 3

4 assumes ac2 : a c/(c 1 ) begin lemma ac: a 1 /(c 1 ) lemma a0 [arith]: a>0 definition b = 1 /(c 1 ) lemma b0 [arith]: b > 0 fun ins :: nat nat nat nat where ins(n,l) = (n+1, if n<l then l else if l=0 then 1 else nat(ceiling(c l))) fun pins :: nat nat => real where pins(n,l) = a n b l interpretation ins: amor where init = (0,0 ) and nxt = %. ins and inv = λ(n,l). if l=0 then n=0 else n l (b/a) l n and t = λ. t ins and Φ = pins and U = λ. a + 1 thm ins.ub end 1.3 Stack with multipop datatype a op stk = Push a Pop nat fun nxt stk :: a op stk a list a list where nxt stk (Push x) xs = x # xs nxt stk (Pop n) xs = drop n xs fun t stk :: a op stk a list real where t stk (Push x) xs = 1 t stk (Pop n) xs = min n (length xs) interpretation stack: amor 4

5 where init = [] and nxt = nxt stk and inv = λ. True and t = t stk and Φ = length and U = λf. case f of Push 0 2 Pop 1.4 Queue See, for example, the book by Okasaki [3]. datatype a op q = Enq a Deq type synonym a queue = a list a list fun nxt q :: a op q a queue a queue where nxt q (Enq x) (xs,ys) = (x#xs,ys) nxt q Deq (xs,ys) = (if ys = [] then ([], tl(rev xs)) else (xs,tl ys)) fun t q :: a op q a queue real where t q (Enq x) (xs,ys) = 1 t q Deq (xs,ys) = (if ys = [] then length xs else 0 ) interpretation queue: amor where init = ([],[]) and nxt = nxt q and inv = λ. True and t = t q and Φ = λ(xs,ys). length xs and U = λf. 2 fun balance :: a queue a queue where balance(xs,ys) = (if size xs size ys then (xs,ys) else ([], rev xs)) fun nxt q2 :: a op q a queue a queue where nxt q2 (Enq a) (xs,ys) = balance (a#xs,ys) nxt q2 Deq (xs,ys) = balance (xs, tl ys) fun t q2 :: a op q a queue real where t q2 (Enq ) (xs,ys) = 1 + (if size xs + 1 size ys then 0 else size xs size ys) t q2 Deq (xs,ys) = (if size xs size ys 1 then 0 else size xs + (size ys 1 )) interpretation queue2 : amor where init = ([],[]) and nxt = nxt q2 and inv = λ(xs,ys). size xs size ys 5

6 and t = t q2 and Φ = λ(xs,ys). 2 size xs and U = λf. case f of Enq 3 Deq Dynamic tables: insert and delete datatype op tb = Ins Del fun nxt tb :: op tb nat nat nat nat where nxt tb Ins (n,l) = (n+1, if n<l then l else if l=0 then 1 else 2 l) nxt tb Del (n,l) = (n 1, if n=1 then 0 else if 4 (n 1 )<l then l div 2 else l) fun t tb :: op tb nat nat real where t tb Ins (n,l) = (if n<l then 1 else n+1 ) t tb Del (n,l) = (if n=1 then 1 else if 4 (n 1 )<l then n else 1 ) interpretation tb: amor where init = (0,0 ) and nxt = nxt tb and inv = λ(n,l). if l=0 then n=0 else n l l 4 n and t = t tb and Φ = (λ(n,l). if 2 n < l then l/2 n else 2 n l) and U = λf. case f of Ins 3 Del 2 end theory Skew Heap Analysis imports../skew Heap/Skew Heap Amor begin 2 Amortized Analysis of Skew Heaps The following proof is a simplified version of the one by Kaldewaij and Schoenmakers [2]. fun rheavy :: a heap bool where rheavy(node l r) = (size l < size r) fun rpath :: a heap a heap list where rpath Leaf = [] rpath (Node l a r) = Node l a r # rpath r fun lpath :: a heap a heap list where lpath Leaf = [] lpath (Node l a r) = Node l a r # lpath l 6

7 fun G where G h = length(filter rheavy (lpath h)) fun D where D h = length(filter (λp. rheavy p) (rpath h)) lemma Gexp: 2 ˆ G h size h + 1 corollary Glog: G h log 2 (size h + 1 ) lemma Dexp: 2 ˆ D h size h + 1 corollary Dlog: D h log 2 (size h + 1 ) function t meld :: a::linorder heap a heap nat where t meld Leaf h = 1 t meld h Leaf = 1 t meld (Node l1 a1 r1 ) (Node l2 a2 r2 ) = (if a1 a2 then t meld (Node l2 a2 r2 ) r1 else t meld (Node l1 a1 r1 ) r2 ) + 1 termination function t meld 2 :: a::linorder heap a heap nat where t meld 2 Leaf Leaf = 1 t meld 2 Leaf (Node l2 a2 r2 ) = t meld 2 r2 Leaf + 1 t meld 2 (Node l1 a1 r1 ) Leaf = t meld 2 r1 Leaf + 1 t meld 2 (Node l1 a1 r1 ) (Node l2 a2 r2 ) = (if a1 a2 then t meld 2 (Node l2 a2 r2 ) r1 else t meld 2 (Node l1 a1 r1 ) r2 ) + 1 termination 2.1 Amortized Analysis fun Φ :: a heap nat where Φ Leaf = 0 7

8 Φ (Node l r) = Φ l + Φ r + (if size l < size r then 1 else 0 ) corollary amor eq: t meld 2 t1 t2 + Φ(meld2 t1 t2 ) Φ t1 Φ t2 = G(meld2 t1 t2 ) + D t1 + D t2 + 1 lemma plus log le 2log plus: assumes [arith]: x 1 y 1 b > 1 shows log b x + log b y 2 log b (x + y) lemma amor le: t meld t1 t2 + Φ (meld t1 t2 ) Φ t1 Φ t2 G(meld t1 t2 ) + D t1 + D t2 + 1 lemma a meld ub: t meld t1 t2 + Φ(meld t1 t2 ) Φ t1 Φ t2 3 log 2 (size t1 + size t2 + 2 ) + 1 (is?l ) datatype a op pq = Insert a Delmin fun nxt pq :: a::linorder op pq a heap a heap where nxt pq (Insert a) h = Skew Heap.insert a h nxt pq Delmin h = del min h fun t pq :: a::linorder op pq a heap nat where t pq (Insert a) h = t meld (Node Leaf a Leaf ) h + 1 t pq Delmin h = (case h of Leaf 1 Node t1 a t2 t meld t1 t2 + 1 ) interpretation pq: amor where init = Leaf and nxt = nxt pq and inv = λ. True and t = t pq and Φ = Φ and U = λf h. case f of Insert 3 log 2 (size h + 3 ) + 2 Delmin 3 log 2 (size h + 3 ) + 4 end theory Splay Tree Analysis imports../splay Tree/Splay Tree 8

9 Amor /src/hol/library/sum of Squares begin 3 Splay Tree Analysis This analysis follows Schoenmakers [4]. 3.1 Time fun t splay :: a::linorder a tree nat where t splay a Leaf = 0 t splay a (Node l b r) = (if a=b then 0 else if a < b then case l of Leaf 0 Node ll c lr (if a=c then 0 else if a < c then if ll = Leaf then 0 else t splay a ll + 1 else if lr = Leaf then 0 else t splay a lr + 1 ) else case r of Leaf 0 Node rl c rr (if a=c then 0 else if a < c then if rl = Leaf then 0 else t splay a rl + 1 else if rr = Leaf then 0 else t splay a rr + 1 )) lemma t splay simps[simp]: t splay a (Node l a r) = 0 a<b = t splay a (Node Leaf b r) = 0 a<b = t splay a (Node (Node ll a lr) b r) = 0 a<b = a<c = t splay a (Node (Node ll c lr) b r) = (if ll = Leaf then 0 else t splay a ll + 1 ) a<b = c<a = t splay a (Node (Node ll c lr) b r) = (if lr = Leaf then 0 else t splay a lr + 1 ) b<a = t splay a (Node l b Leaf ) = 0 b<a = t splay a (Node l b (Node rl a rr)) = 0 b<a = a<c = t splay a (Node l b (Node rl c rr)) = (if rl=leaf then 0 else t splay a rl + 1 ) b<a = c<a = t splay a (Node l b (Node rl c rr)) = (if rr=leaf then 0 else t splay a rr + 1 ) 9

10 declare t splay.simps(2 )[simp del] 3.2 Wlog in tree lemma ex in set tree: t Leaf = bst t = a set tree t. splay a t = splay a t t splay a t = t splay a t fun t splay max :: a::linorder tree nat where t splay max Leaf = 0 t splay max (Node l b Leaf ) = 0 t splay max (Node l b (Node rl c Leaf )) = 0 t splay max (Node l b (Node rl c rr)) = t splay max rr Analysis of splay locale Splay Analysis = fixes α :: real and β :: real assumes a1 [arith]: α > 1 assumes A1 : [0 x; 0 y; 0 z ] = (2 +x+y) (2 +y+z) powr β (2 +x+y) powr β (3 +x+y+z) assumes A2 : [0 l ; 0 r ; 0 lr; 0 r ] = α (2 +l +r ) (2 +lr+r) powr β (3 +lr+r +r) powr β (2 +l +r ) powr β (3 +l +lr+r ) powr β (l +lr+r +r+4 ) assumes A3 : [0 l ; 0 r ; 0 ll; 0 r ] = α (2 +l +r ) (2 +l +ll) powr β (2 +r +r) powr β (2 +l +r ) powr β (3 +l +ll+r ) powr β (l +ll+r +r+4 ) begin lemma nl2 : [ ll 0 ; lr 0 ; r 0 ] = log α (2 + ll + lr) + β log α (2 + lr + r) β log α (2 + ll + lr) + log α (3 + ll + lr + r) definition ϕ :: a tree a tree real where ϕ t1 t2 = β log α (size t1 + size t2 + 2 ) fun Φ :: a tree real where Φ Leaf = 0 Φ (Node l r) = Φ l + Φ r + ϕ l r 10

11 definition A :: a::linorder a tree real where A a t = t splay a t + Φ(splay a t) Φ t lemma A simps[simp]: A a (Node l a r) = 0 a<b = A a (Node (Node ll a lr) b r) = ϕ lr r ϕ lr ll b<a = A a (Node l b (Node rl a rr)) = ϕ rl l ϕ rr rl lemma A simp3 : [ a<b; b<c; bst ll; a set tree ll ] = A a (Node (Node ll b lr) c r) = (case splay a ll of Node lll llr A a ll + ϕ llr (Node lr c r) + ϕ lr r ϕ ll lr ϕ lll llr + 1 ) lemma A simp3 : [ b<a; c<b; bst rr; a set tree rr ] = A a (Node l c (Node rl b rr)) = (case splay a rr of Node rrl rrr A a rr + ϕ rrl (Node l c rl) + ϕ l rl ϕ rl rr ϕ rrl rrr + 1 ) lemma A simp4 : [ b<a; a<c; bst lr; a set tree lr ] = A a (Node (Node ll b lr) c r) = (case splay a lr of Node lrl lrr A a lr + ϕ ll lrl + ϕ lrr r ϕ ll lr ϕ lrl lrr + 1 ) lemma A simp4 : [ c<a; a<b; bst rl; a set tree rl ] = A a (Node l c (Node rl b rr)) = (case splay a rl of Node rll rlr A a rl + ϕ rr rlr + ϕ rll l ϕ rl rr ϕ rll rlr + 1 ) lemma A ub: [ bst t; Node la a ra : subtrees t ] = A a t log α ((size t + 1 )/(size la + size ra + 2 )) lemma A ub2 : assumes bst t a : set tree t shows A a t log α ((size t + 1 )/2 ) lemma A ub3 : assumes bst t shows A a t log α (size t + 1 ) 11

12 definition Am :: a::linorder tree real where Am t = t splay max t + Φ(splay max t) Φ t lemma Am simp3 : [ c<b; bst rr; rr Leaf ] = Am (Node l c (Node rl b rr)) = (case splay max rr of Node rrl rrr Am rr + ϕ rrl (Node l c rl) + ϕ l rl ϕ rl rr ϕ rrl rrr + 1 ) lemma Am ub: [ bst t; t Leaf ] = Am t log α ((size t + 1 )/2 ) lemma Am ub3 : assumes bst t shows Am t log α (size t + 1 ) end 3.4 Optimal Interpretation lemma mult root eq root: n>0 = y 0 = root n x y = root n (x (y ˆ n)) lemma mult root eq root2 : n>0 = y 0 = y root n x = root n ((y ˆ n) x) lemma powr inverse numeral: 0 < x = x powr (1 / numeral n) = root (numeral n) x lemmas root simps = mult root eq root mult root eq root2 powr inverse numeral lemma nl31 : [ (l ::real) 0 ; r 0 ; lr 0 ; r 0 ] = 4 (2 + l + r ) (2 + lr + r) (l + lr + r + r + 4 )ˆ2 lemma nl32 : assumes (l ::real) 0 r 0 lr 0 r 0 shows 4 (2 + l + r ) (2 + lr + r) (3 + lr + r + r) (l + lr + r + r + 4 )ˆ3 12

13 lemma nl3 : assumes (l ::real) 0 r 0 lr 0 r 0 shows 4 (2 + l + r )ˆ2 (2 + lr + r) (3 + lr + r + r) (3 + l + lr + r ) (l + lr + r + r + 4 )ˆ3 lemma nl41 : assumes (l ::real) 0 r 0 ll 0 r 0 shows 4 (2 + l + ll) (2 + r + r) (l + ll + r + r + 4 )ˆ2 lemma nl42 : assumes (l ::real) 0 r 0 ll 0 r 0 shows 4 (2 + l + r ) (2 + l + ll) (2 + r + r) (l + ll + r + r + 4 )ˆ3 lemma nl4 : assumes (l ::real) 0 r 0 ll 0 r 0 shows 4 (2 + l + r )ˆ2 (2 + l + ll) (2 + r + r) (3 + l + ll + r ) (l + ll + r + r + 4 )ˆ3 lemma cancel: x>(0 ::real) = c x ˆ 2 y z u v = c x ˆ 3 y z x u v interpretation S34 : Splay Analysis root /3 lemma log4 log2 : log 4 x = log 2 x / 2 declare log base root[simp] real of nat Suc[simp] lemma A34 ub: assumes bst t shows S34.A a t (3 /2 ) log 2 (size t + 1 ) lemma Am34 ub: assumes bst t shows S34.Am t (3 /2 ) log 2 (size t + 1 ) 13

14 3.5 Overall analysis fun t delete :: a::linorder a tree nat where t delete a Leaf = 0 t delete a t = t splay a t + (case splay a t of Node l x r if x=a then case l of Leaf 0 t splay max l else 0 ) datatype a op st = Splay a Insert a Delete a fun nxt st :: a::linorder op st a tree a tree where nxt st (Splay a) t = splay a t nxt st (Insert a) t = Splay Tree.insert a t nxt st (Delete a) t = Splay Tree.delete a t fun t st :: a::linorder op st a tree real where t st (Splay a) t = t splay a t t st (Insert a) t = t splay a t t st (Delete a) t = t delete a t interpretation ST : amor where init = Leaf and nxt = nxt st and inv = bst and t = t st and Φ = S34.Φ and U = λf t. case f of Splay (3 /2 ) log 2 (size t + 1 ) Insert 2 log 2 (size t + 1 ) + 1 /2 Delete 3 log 2 (size t + 1 ) end theory Inversion imports../list Index/List Index begin 4 List Inversion abbreviation dist perm xs ys distinct xs distinct ys set xs = set ys definition before in :: a a a list bool (( </ / in ) [55,55,55 ] 55 ) where x < y in xs = (index xs x < index xs y y set xs) definition Inv :: a list a list ( a a) set where 14

15 Inv xs ys = {(x,y). x < y in xs y < x in ys} lemma before in setd1 : x < y in xs = x : set xs lemma before in setd2 : x < y in xs = y : set xs lemma not before in: x : set xs = y : set xs = x < y in xs y < x in xs x=y lemma no before ini [simp]: x < y in xs = ( y < x in xs) = True lemma finite Invs[simp]: finite(inv xs ys) lemma Inv id[simp]: Inv xs xs = {} lemma card Inv sym: card(inv xs ys) = card(inv ys xs) lemma Inv tri ineq: dist perm xs ys = dist perm ys zs = Inv xs zs Inv xs ys Un Inv ys zs lemma card Inv tri ineq: dist perm xs ys = dist perm ys zs = card (Inv xs zs) card(inv xs ys) + card (Inv ys zs) definition mtf x xs = (if x set xs then x # (take (index xs x) drop (index xs x + 1 ) xs else xs) lemma mtf id[simp]: x / set xs = mtf x xs = xs lemma before in mtf : assumes z set xs shows x < y in mtf z xs 15

16 (y z (if x=z then y set xs else x < y in xs)) lemma Inv mtf : set xs = set ys = z : set ys = Inv xs (mtf z ys) = Inv xs ys {(x,z) x. x < z in xs x < z in ys} {(z,x) x. z < x in xs x < z in ys} lemma set mtf [simp]: set(mtf x xs) = set xs lemma length mtf [simp]: size (mtf x xs) = size xs lemma distinct mtf [simp]: distinct (mtf x xs) = distinct xs definition swapsuc n xs = xs[n := xs!suc n, Suc n := xs!n] lemma index swapsuc distinct: distinct xs = Suc n < length xs = index (swapsuc n xs) x = (if x = xs!n then Suc n else if x = xs!suc n then n else index xs x) lemma set swapsuc[simp]: Suc n < size xs = set(swapsuc n xs) = set xs lemma before in swapsuc: dist perm xs ys = Suc n < size xs = x < y in (swapsuc n xs) x < y in xs (x = xs!n y = xs!suc n) x = xs!suc n y = xs!n lemma Inv swap: assumes dist perm xs ys Suc n < size xs shows Inv xs (swapsuc n ys) = (if ys!n < ys!suc n in xs then Inv xs ys {(ys!n, ys!suc n)} else Inv xs ys {(ys!suc n, ys!n)}) end theory Move to Front 16

17 imports Complex Main Inversion begin 5 Online Algorithm Move-to-Front It was first proved by Sleator and Tarjan [5] that the Move-to-Front algorithm is 2-competitive. lemma sum telescope: fixes f :: nat a::linordered ab group add shows ( i=0..<k. f (Suc i) f i) = f k f 0 lemma potential: fixes t :: nat a::linordered ab group add and p :: nat a assumes p0 : p 0 = 0 and ppos: n. p n 0 and ub: n. t n + p(suc n) p n u n shows ( i<n. t i) ( i<n. u i) abbreviation before x xs {y. y < x in xs} abbreviation after x xs {y. x < y in xs} lemma finite before[simp]: finite (before x xs) lemma finite after[simp]: finite (after x xs) lemma before conv take: x : set xs = before x xs = set(take (index xs x) xs) lemma card before: distinct xs = x : set xs = card (before x xs) = index xs x lemma card subset before: distinct xs = x : set xs = A before x xs = card(a) index xs x lemma before Un: set xs = set ys = x : set xs = before x ys = before x xs before x ys Un after x xs before x ys 17

18 lemma before disj : dist perm xs ys = x : set xs = before x xs before x ys (after x xs before x ys) = {} lemma phi diff aux: card (Inv xs ys {(y, x) y. y < x in xs y < x in ys} {(x, y) y. x < y in xs y < x in ys}) = card(inv xs ys) + card(before x xs before x ys) int(card(after x xs before x ys)) (is card(?i?b?a) = card?i + card?b int(card?a)) lemma foldr swap inv: i set sws. Suc i < length xs = set (foldr swapsuc sws xs) = set xs size(foldr swapsuc sws xs) = size xs distinct(foldr swapsuc sws xs) = distinct xs lemma card Inv foldr swapsuc: i set sws. Suc i < length xs = distinct xs = card (Inv xs (foldr swapsuc sws xs)) length sws locale MTF Off = fixes q :: nat a fixes init :: a list fixes sw off :: nat nat list assumes dist init[simp]: distinct init assumes sw off size: i set(sw off n). Suc i < size init begin fun s off :: nat a list where s off 0 = init s off (Suc n) = foldr swapsuc (sw off n) (s off n) lemma length s off [simp]: length(s off n) = length init lemma dist s off [simp]: distinct(s off n) 18

19 lemma set s off [simp]: set(s off n) = set init fun s mtf :: nat a list where s mtf 0 = init s mtf (Suc n) = mtf (q n) (s mtf n) definition t mtf :: nat int where t mtf n = index (s mtf n) (q n) + 1 definition t off :: nat int where t off n = index (s off n) (q n) size(sw off n) lemma length s mtf [simp]: length(s mtf n) = length init lemma dist s mtf [simp]: distinct(s mtf n) lemma set s mtf [simp]: set (s mtf n) = set init lemma dperm inv: dist perm (s off n) (s mtf n) definition Phi :: nat int (Φ) where Phi n = card(inv (s off n) (s mtf n)) lemma phi0 : Phi 0 = 0 lemma phi pos: Phi n 0 lemma mtf ub: t mtf n + Phi (Suc n) Phi n 2 t off n theorem t mtf t off : ( i<n. t mtf i) 2 ( i<n. t off i) end 19

20 end References [1] T. H. Cormen, C. E. Leiserson, and R. L. Rivest. Introduction to Algorithms. MIT Press, [2] A. Kaldewaij and B. Schoenmakers. The derivation of a tighter bound for top-down skew heaps. Information Processing Letters, 37: , [3] C. Okasaki. Purely Functional Data Structures. Cambridge University Press, [4] B. Schoenmakers. A systematic analysis of splaying. Information Processing Letters, 45:41 50, [5] D. D. Sleator and R. E. Tarjan. Amortized efficiency of list update and paging rules. Comm. ACM, 28(2): ,

Amortized Complexity Verified

Amortized Complexity Verified Amortized Complexity Verified Tobias Nipkow Technische Universität München Abstract A framework for the analysis of the amortized complexity of (functional) data structures is formalized in Isabelle/HOL

More information

Functional Data Structures

Functional Data Structures Functional Data Structures with Isabelle/HOL Tobias Nipkow Fakultät für Informatik Technische Universität München 2017-2-3 1 Part II Functional Data Structures 2 Chapter 1 Binary Trees 3 1 Binary Trees

More information

Verified Analysis of Functional Data Structures

Verified Analysis of Functional Data Structures Verified Analysis of Functional Data Structures with Isabelle/HOL Tobias Nipkow Fakultät für Informatik Technische Universität München 2017-8-6 1 Lecture 1 Introduction Lecture 2 Search Trees Lecture 3

More information

A Tight Lower Bound for Top-Down Skew Heaps

A Tight Lower Bound for Top-Down Skew Heaps A Tight Lower Bound for Top-Down Skew Heaps Berry Schoenmakers January, 1997 Abstract Previously, it was shown in a paper by Kaldewaij and Schoenmakers that for topdown skew heaps the amortized number

More information

Red-black Trees of smlnj Studienarbeit. Angelika Kimmig

Red-black Trees of smlnj Studienarbeit. Angelika Kimmig Red-black Trees of smlnj Studienarbeit Angelika Kimmig January 26, 2004 Contents 1 Introduction 1 2 Development of the Isabelle/HOL Theory RBT 2 2.1 Translation of Datatypes and Functions into HOL......

More information

Free Groups. Joachim Breitner. April 17, 2016

Free Groups. Joachim Breitner. April 17, 2016 Free Groups Joachim Breitner April 17, 2016 Abstract Free Groups are, in a sense, the most generic kind of group. They are defined over a set of generators with no additional relations in between them.

More information

A General Method for the Proof of Theorems on Tail-recursive Functions

A General Method for the Proof of Theorems on Tail-recursive Functions A General Method for the Proof of Theorems on Tail-recursive Functions Pasquale Noce Security Certification Specialist at Arjo Systems - Gep S.p.A. pasquale dot noce dot lavoro at gmail dot com pasquale

More information

Cauchy s Mean Theorem and the Cauchy-Schwarz Inequality. Benjamin Porter

Cauchy s Mean Theorem and the Cauchy-Schwarz Inequality. Benjamin Porter Cauchy s Mean Theorem and the Cauchy-Schwarz Inequality Benjamin Porter March 12, 2013 Contents 1 Cauchy s Mean Theorem 3 1.1 Abstract.............................. 3 1.2 Formal proof...........................

More information

Landau Symbols. Manuel Eberl. November 28, 2018

Landau Symbols. Manuel Eberl. November 28, 2018 Landau Symbols Manuel Eberl November 28, 2018 Contents 1 Sorting and grouping factors 1 2 Decision procedure for real functions 4 2.1 Eventual non-negativity/non-zeroness............. 4 2.2 Rewriting Landau

More information

Imperative Insertion Sort

Imperative Insertion Sort Imperative Insertion Sort Christian Sternagel April 17, 2016 Contents 1 Looping Constructs for Imperative HOL 1 1.1 While Loops............................ 1 1.2 For Loops.............................

More information

CS213d Data Structures and Algorithms

CS213d Data Structures and Algorithms CS21d Data Structures and Algorithms Heaps and their Applications Milind Sohoni IIT Bombay and IIT Dharwad March 22, 2017 1 / 18 What did I see today? March 22, 2017 2 / 18 Heap-Trees A tree T of height

More information

AMORTIZED ANALYSIS. binary counter multipop stack dynamic table. Lecture slides by Kevin Wayne. Last updated on 1/24/17 11:31 AM

AMORTIZED ANALYSIS. binary counter multipop stack dynamic table. Lecture slides by Kevin Wayne. Last updated on 1/24/17 11:31 AM AMORTIZED ANALYSIS binary counter multipop stack dynamic table Lecture slides by Kevin Wayne http://www.cs.princeton.edu/~wayne/kleinberg-tardos Last updated on 1/24/17 11:31 AM Amortized analysis Worst-case

More information

Examples for program extraction in Higher-Order Logic

Examples for program extraction in Higher-Order Logic Examples for program extraction in Higher-Order Logic Stefan Berghofer October 10, 2011 Contents 1 Auxiliary lemmas used in program extraction examples 1 2 Quotient and remainder 2 3 Greatest common divisor

More information

The Factorization Algorithm of Berlekamp and Zassenhaus

The Factorization Algorithm of Berlekamp and Zassenhaus The Factorization Algorithm of Berlekamp and Zassenhaus Jose Divasón Sebastiaan Joosten René Thiemann Akihisa Yamada November 10, 2018 Abstract We formalize the Berlekamp-Zassenhaus algorithm for factoring

More information

Amortized analysis. Amortized analysis

Amortized analysis. Amortized analysis In amortized analysis the goal is to bound the worst case time of a sequence of operations on a data-structure. If n operations take T (n) time (worst case), the amortized cost of an operation is T (n)/n.

More information

Chapter 4. Greedy Algorithms. Slides by Kevin Wayne. Copyright 2005 Pearson-Addison Wesley. All rights reserved.

Chapter 4. Greedy Algorithms. Slides by Kevin Wayne. Copyright 2005 Pearson-Addison Wesley. All rights reserved. Chapter 4 Greedy Algorithms Slides by Kevin Wayne. Copyright Pearson-Addison Wesley. All rights reserved. 4 4.1 Interval Scheduling Interval Scheduling Interval scheduling. Job j starts at s j and finishes

More information

Verified Analysis of Algorithms for the List Update Problem

Verified Analysis of Algorithms for the List Update Problem FAKULTÄT FÜR INFORMATIK DER TECHNISCHEN UNIVERSITÄT MÜNCHEN Master s Thesis in Informatics Verified Analysis of Algorithms for the List Update Problem Maximilian Paul Louis Haslbeck FAKULTÄT FÜR INFORMATIK

More information

past balancing schemes require maintenance of balance info at all times, are aggresive use work of searches to pay for work of rebalancing

past balancing schemes require maintenance of balance info at all times, are aggresive use work of searches to pay for work of rebalancing 1 Splay Trees Sleator and Tarjan, Self Adjusting Binary Search Trees JACM 32(3) 1985 The claim planning ahead. But matches previous idea of being lazy, letting potential build up, using it to pay for expensive

More information

Non-context-Free Languages. CS215, Lecture 5 c

Non-context-Free Languages. CS215, Lecture 5 c Non-context-Free Languages CS215, Lecture 5 c 2007 1 The Pumping Lemma Theorem. (Pumping Lemma) Let be context-free. There exists a positive integer divided into five pieces, Proof for for each, and..

More information

Lecture 5: Splay Trees

Lecture 5: Splay Trees 15-750: Graduate Algorithms February 3, 2016 Lecture 5: Splay Trees Lecturer: Gary Miller Scribe: Jiayi Li, Tianyi Yang 1 Introduction Recall from previous lecture that balanced binary search trees (BST)

More information

Reversing Lists Haskell: The craft of functional programming, by Simon Thompson, page 140

Reversing Lists Haskell: The craft of functional programming, by Simon Thompson, page 140 Reversing Lists Haskell: The craft of functional programming, by Simon Thompson, page 14 Reverse.hs rev :: [t] -> [t] rev [] = [] rev (e:l) = (rev l) ++ [e] 6 5 rev [1..n] n**2/2 4 3 2 1 2 4 6 8 1 1 Array

More information

Regular Expressions [1] Regular Expressions. Regular expressions can be seen as a system of notations for denoting ɛ-nfa

Regular Expressions [1] Regular Expressions. Regular expressions can be seen as a system of notations for denoting ɛ-nfa Regular Expressions [1] Regular Expressions Regular expressions can be seen as a system of notations for denoting ɛ-nfa They form an algebraic representation of ɛ-nfa algebraic : expressions with equations

More information

Fibonacci Heaps These lecture slides are adapted from CLRS, Chapter 20.

Fibonacci Heaps These lecture slides are adapted from CLRS, Chapter 20. Fibonacci Heaps These lecture slides are adapted from CLRS, Chapter 20. Princeton University COS 4 Theory of Algorithms Spring 2002 Kevin Wayne Priority Queues Operation mae-heap insert find- delete- union

More information

Computing N-th Roots using the Babylonian Method

Computing N-th Roots using the Babylonian Method Computing N-th Roots using the Babylonian Method René Thiemann May 27, 2015 Abstract We implement the Babylonian method [1] to compute n-th roots of numbers. We provide precise algorithms for naturals,

More information

Algorithms and Data Structures

Algorithms and Data Structures Algorithms and Data Structures Amortized Analysis Ulf Leser Two Examples Two Analysis Methods Dynamic Tables SOL - Analysis This lecture is not covered in [OW93] but, for instance, in [Cor03] Ulf Leser:

More information

Another way of saying this is that amortized analysis guarantees the average case performance of each operation in the worst case.

Another way of saying this is that amortized analysis guarantees the average case performance of each operation in the worst case. Amortized Analysis: CLRS Chapter 17 Last revised: August 30, 2006 1 In amortized analysis we try to analyze the time required by a sequence of operations. There are many situations in which, even though

More information

Lecture 2 14 February, 2007

Lecture 2 14 February, 2007 6.851: Advanced Data Structures Spring 2007 Prof. Erik Demaine Lecture 2 14 February, 2007 Scribes: Christopher Moh, Aditya Rathnam 1 Overview In the last lecture we covered linked lists and some results

More information

Computational Models - Lecture 4

Computational Models - Lecture 4 Computational Models - Lecture 4 Regular languages: The Myhill-Nerode Theorem Context-free Grammars Chomsky Normal Form Pumping Lemma for context free languages Non context-free languages: Examples Push

More information

Chapter 5 Data Structures Algorithm Theory WS 2017/18 Fabian Kuhn

Chapter 5 Data Structures Algorithm Theory WS 2017/18 Fabian Kuhn Chapter 5 Data Structures Algorithm Theory WS 2017/18 Fabian Kuhn Priority Queue / Heap Stores (key,data) pairs (like dictionary) But, different set of operations: Initialize-Heap: creates new empty heap

More information

Automated Reasoning. Lecture 9: Isar A Language for Structured Proofs

Automated Reasoning. Lecture 9: Isar A Language for Structured Proofs Automated Reasoning Lecture 9: Isar A Language for Structured Proofs Jacques Fleuriot jdf@inf.ed.ac.uk Acknowledgement: Tobias Nipkow kindly provided the slides for this lecture Apply scripts unreadable

More information

Analyse et Conception Formelle. Lesson 4. Proofs with a proof assistant

Analyse et Conception Formelle. Lesson 4. Proofs with a proof assistant Analyse et Conception Formelle Lesson 4 Proofs with a proof assistant T. Genet (ISTIC/IRISA) ACF-4 1 / 26 Prove logic formulas... to prove programs fun nth:: "nat => a list => a" where "nth 0 (x#_)=x"

More information

Much Ado about Two. By Sascha Böhme. February 22, 2013

Much Ado about Two. By Sascha Böhme. February 22, 2013 Much Ado about Two By Sascha Böhme February 22, 2013 Abstract This article is an Isabelle formalisation of a paper with the same. In a similar way as Knuth s 0-1-principle for sorting algorithms, that

More information

Splay Trees. CMSC 420: Lecture 8

Splay Trees. CMSC 420: Lecture 8 Splay Trees CMSC 420: Lecture 8 AVL Trees Nice Features: - Worst case O(log n) performance guarantee - Fairly simple to implement Problem though: - Have to maintain etra balance factor storage at each

More information

MP 5 Program Transition Systems and Linear Temporal Logic

MP 5 Program Transition Systems and Linear Temporal Logic MP 5 Program Transition Systems and Linear Temporal Logic CS 477 Spring 2018 Revision 1.0 Assigned April 10, 2018 Due April 17, 2018, 9:00 PM Extension extend48 hours (penalty 20% of total points possible)

More information

4.8 Huffman Codes. These lecture slides are supplied by Mathijs de Weerd

4.8 Huffman Codes. These lecture slides are supplied by Mathijs de Weerd 4.8 Huffman Codes These lecture slides are supplied by Mathijs de Weerd Data Compression Q. Given a text that uses 32 symbols (26 different letters, space, and some punctuation characters), how can we

More information

SAT Solver verification

SAT Solver verification SAT Solver verification By Filip Marić April 17, 2016 Abstract This document contains formall correctness proofs of modern SAT solvers. Two different approaches are used state-transition systems shallow

More information

Subresultants. Sebastiaan Joosten, René Thiemann and Akihisa Yamada. October 10, 2017

Subresultants. Sebastiaan Joosten, René Thiemann and Akihisa Yamada. October 10, 2017 Subresultants Sebastiaan Joosten, René Thiemann and Akihisa Yamada October 10, 2017 Abstract We formalize the theory of subresultants and the subresultant polynomial remainder sequence as described by

More information

Computational Models - Lecture 5 1

Computational Models - Lecture 5 1 Computational Models - Lecture 5 1 Handout Mode Iftach Haitner. Tel Aviv University. November 28, 2016 1 Based on frames by Benny Chor, Tel Aviv University, modifying frames by Maurice Herlihy, Brown University.

More information

The Hahn-Banach Theorem for Real Vector Spaces

The Hahn-Banach Theorem for Real Vector Spaces The Hahn-Banach Theorem for Real Vector Spaces Gertrud Bauer http://www.in.tum.de/ bauerg/ February 12, 2013 Abstract The Hahn-Banach Theorem is one of the most fundamental results in functional analysis.

More information

An Isabelle/HOL Formalization of the Textbook Proof of Huffman s Algorithm

An Isabelle/HOL Formalization of the Textbook Proof of Huffman s Algorithm An Isabelle/HOL Formalization of the Textbook Proof of Huffman s Algorithm Jasmin Christian Blanchette Institut für Informatik, Technische Universität München, Germany blanchette@in.tum.de March 12, 2013

More information

Note: In any grammar here, the meaning and usage of P (productions) is equivalent to R (rules).

Note: In any grammar here, the meaning and usage of P (productions) is equivalent to R (rules). Note: In any grammar here, the meaning and usage of P (productions) is equivalent to R (rules). 1a) G = ({R, S, T}, {0,1}, P, S) where P is: S R0R R R0R1R R1R0R T T 0T ε (S generates the first 0. R generates

More information

The Tortoise and the Hare Algorithm

The Tortoise and the Hare Algorithm The Tortoise and the Hare Algorithm Peter Gammie October 11, 2017 Abstract We formalize the Tortoise and Hare cycle-finding algorithm ascribed to Floyd by Knuth (1981, p7, exercise 6), and an improved

More information

Imperative Insertion Sort

Imperative Insertion Sort Imperative Insertion Sort Christian Sternagel October 11, 2017 Contents 1 Looping Constructs for Imperative HOL 1 1.1 While Loops............................ 1 1.2 For Loops.............................

More information

Code Generation for a Simple First-Order Prover

Code Generation for a Simple First-Order Prover Code Generation for a Simple First-Order Prover Jørgen Villadsen, Anders Schlichtkrull, and Andreas Halkjær From DTU Compute, Technical University of Denmark, 2800 Kongens Lyngby, Denmark Abstract. We

More information

Discrete Mathematics Review

Discrete Mathematics Review CS 1813 Discrete Mathematics Discrete Mathematics Review or Yes, the Final Will Be Comprehensive 1 Truth Tables for Logical Operators P Q P Q False False False P Q False P Q False P Q True P Q True P True

More information

Part 4 out of 5 DFA NFA REX. Automata & languages. A primer on the Theory of Computation. Last week, we showed the equivalence of DFA, NFA and REX

Part 4 out of 5 DFA NFA REX. Automata & languages. A primer on the Theory of Computation. Last week, we showed the equivalence of DFA, NFA and REX Automata & languages A primer on the Theory of Computation Laurent Vanbever www.vanbever.eu Part 4 out of 5 ETH Zürich (D-ITET) October, 12 2017 Last week, we showed the equivalence of DFA, NFA and REX

More information

Formalizing Push-Relabel Algorithms

Formalizing Push-Relabel Algorithms Formalizing Push-Relabel Algorithms Peter Lammich and S. Reza Sefidgar March 6, 2017 Abstract We present a formalization of push-relabel algorithms for computing the maximum flow in a network. We start

More information

Properties of Context-Free Languages

Properties of Context-Free Languages Properties of Context-Free Languages Seungjin Choi Department of Computer Science and Engineering Pohang University of Science and Technology 77 Cheongam-ro, Nam-gu, Pohang 37673, Korea seungjin@postech.ac.kr

More information

HOL: Well-founded and Primitive Recursion

HOL: Well-founded and Primitive Recursion Dipl.-Inf. Achim D. Brucker Dr. Burkhart Wolff Computer-supported Modeling and Reasoning http://www.infsec.ethz.ch/ education/permanent/csmr/ (rev. 16814) Submission date: HOL: Well-founded and Primitive

More information

AVL Trees. Properties Insertion. October 17, 2017 Cinda Heeren / Geoffrey Tien 1

AVL Trees. Properties Insertion. October 17, 2017 Cinda Heeren / Geoffrey Tien 1 AVL Trees Properties Insertion October 17, 2017 Cinda Heeren / Geoffrey Tien 1 BST rotation g g e h b h b f i a e i a c c f d d Rotations preserve the in-order BST property, while altering the tree structure

More information

Menger's Theorem. Christoph Dittmann August 16, Contents

Menger's Theorem. Christoph Dittmann August 16, Contents Menger's Theorem Christoph Dittmann isabelle@christoph-d.de August 16, 2018 We present a formalization of Menger's Theorem for directed and undirected graphs in Isabelle/HOL. This well-known result shows

More information

CSE548, AMS542: Analysis of Algorithms, Fall 2017 Date: Oct 26. Homework #2. ( Due: Nov 8 )

CSE548, AMS542: Analysis of Algorithms, Fall 2017 Date: Oct 26. Homework #2. ( Due: Nov 8 ) CSE548, AMS542: Analysis of Algorithms, Fall 2017 Date: Oct 26 Homework #2 ( Due: Nov 8 ) Task 1. [ 80 Points ] Average Case Analysis of Median-of-3 Quicksort Consider the median-of-3 quicksort algorithm

More information

The Divergence of the Prime Harmonic Series

The Divergence of the Prime Harmonic Series The Divergence of the Prime Harmonic Series Manuel Eberl April 17, 2016 Abstract In this work, we prove the lower bound ln(h n ) ln( 5 3 ) for the partial sum of the Prime Harmonic series and, based on

More information

Syntax and semantics of a GPU kernel programming language

Syntax and semantics of a GPU kernel programming language Syntax and semantics of a GPU kernel programming language John Wickerson April 17, 2016 Abstract This document accompanies the article The Design and Implementation of a Verification Technique for GPU

More information

Adding Relations in Multi-levels to an Organization Structure of a Complete Binary Tree Maximizing Total Shortening Path Length

Adding Relations in Multi-levels to an Organization Structure of a Complete Binary Tree Maximizing Total Shortening Path Length Adding Relations in Multi-levels to an Organization Structure of a Complete Binary Tree Maximizing Total Shortening Path Length Kiyoshi Sawada and Kazuyuki Amano Abstract This paper proposes a model of

More information

The Divergence of the Prime Harmonic Series

The Divergence of the Prime Harmonic Series The Divergence of the Prime Harmonic Series Manuel Eberl December 16, 2018 Abstract In this work, we prove the lower bound ln(h n ) ln( 5 3 ) for the partial sum of the Prime Harmonic series and, based

More information

8 Priority Queues. 8 Priority Queues. Prim s Minimum Spanning Tree Algorithm. Dijkstra s Shortest Path Algorithm

8 Priority Queues. 8 Priority Queues. Prim s Minimum Spanning Tree Algorithm. Dijkstra s Shortest Path Algorithm 8 Priority Queues 8 Priority Queues A Priority Queue S is a dynamic set data structure that supports the following operations: S. build(x 1,..., x n ): Creates a data-structure that contains just the elements

More information

Hall s Marriage Theorem

Hall s Marriage Theorem Hall s Marriage Theorem Dongchen Jiang and Tobias Nipkow February 16, 2013 Abstract A proof of Hall s Marriage Theorem due to Halmos and Vaughan [1]. theory Marriage imports Main begin theorem marriage-necessary:

More information

Amortized Analysis. DistributeMoney(n, k) 1 Each of n people gets $1. 2 for i = 1to k 3 do Give a dollar to a random person

Amortized Analysis. DistributeMoney(n, k) 1 Each of n people gets $1. 2 for i = 1to k 3 do Give a dollar to a random person Amortized Analysis DistributeMoney(n, k) 1 Each of n people gets $1. 2 for i = 1to k 3 do Give a dollar to a random person What is the maximum amount of money I can receive? Amortized Analysis DistributeMoney(n,

More information

A Formalisation of Lehmer s Primality Criterion

A Formalisation of Lehmer s Primality Criterion A Formalisation of Lehmer s Primality Criterion By Simon Wimmer and Lars Noschinski April 17, 2016 Abstract In 1927, Lehmer presented criterions for primality, based on the converse of Fermat s litte theorem

More information

Section 1 (closed-book) Total points 30

Section 1 (closed-book) Total points 30 CS 454 Theory of Computation Fall 2011 Section 1 (closed-book) Total points 30 1. Which of the following are true? (a) a PDA can always be converted to an equivalent PDA that at each step pops or pushes

More information

UNIT II REGULAR LANGUAGES

UNIT II REGULAR LANGUAGES 1 UNIT II REGULAR LANGUAGES Introduction: A regular expression is a way of describing a regular language. The various operations are closure, union and concatenation. We can also find the equivalent regular

More information

Rank-Nullity Theorem in Linear Algebra

Rank-Nullity Theorem in Linear Algebra Rank-Nullity Theorem in Linear Algebra By Jose Divasón and Jesús Aransay April 17, 2016 Abstract In this contribution, we present some formalizations based on the HOL-Multivariate-Analysis session of Isabelle.

More information

Roy L. Crole. Operational Semantics Abstract Machines and Correctness. University of Leicester, UK

Roy L. Crole. Operational Semantics Abstract Machines and Correctness. University of Leicester, UK Midlands Graduate School, University of Birmingham, April 2008 1 Operational Semantics Abstract Machines and Correctness Roy L. Crole University of Leicester, UK Midlands Graduate School, University of

More information

Accumulators. A Trivial Example in Oz. A Trivial Example in Prolog. MergeSort Example. Accumulators. Declarative Programming Techniques

Accumulators. A Trivial Example in Oz. A Trivial Example in Prolog. MergeSort Example. Accumulators. Declarative Programming Techniques Declarative Programming Techniques Accumulators, Difference Lists (VRH 3.4.3-3.4.4) Carlos Varela RPI Adapted with permission from: Seif Haridi KTH Peter Van Roy UCL September 13, 2007 Accumulators Accumulator

More information

Amortized Complexity Main Idea

Amortized Complexity Main Idea omp2711 S1 2006 Amortized Complexity Example 1 Amortized Complexity Main Idea Worst case analysis of run time complexity is often too pessimistic. Average case analysis may be difficult because (i) it

More information

Declarative Programming Techniques

Declarative Programming Techniques Declarative Programming Techniques Accumulators (CTM 3.4.3) Difference Lists (CTM 3.4.4) Carlos Varela RPI Adapted with permission from: Seif Haridi KTH Peter Van Roy UCL December 1, 2015 C. Varela; Adapted

More information

ENS Lyon Camp. Day 2. Basic group. Cartesian Tree. 26 October

ENS Lyon Camp. Day 2. Basic group. Cartesian Tree. 26 October ENS Lyon Camp. Day 2. Basic group. Cartesian Tree. 26 October Contents 1 Cartesian Tree. Definition. 1 2 Cartesian Tree. Construction 1 3 Cartesian Tree. Operations. 2 3.1 Split............................................

More information

EXAM. CS331 Compiler Design Spring Please read all instructions, including these, carefully

EXAM. CS331 Compiler Design Spring Please read all instructions, including these, carefully EXAM Please read all instructions, including these, carefully There are 7 questions on the exam, with multiple parts. You have 3 hours to work on the exam. The exam is open book, open notes. Please write

More information

Search Trees. Chapter 10. CSE 2011 Prof. J. Elder Last Updated: :52 AM

Search Trees. Chapter 10. CSE 2011 Prof. J. Elder Last Updated: :52 AM Search Trees Chapter 1 < 6 2 > 1 4 = 8 9-1 - Outline Ø Binary Search Trees Ø AVL Trees Ø Splay Trees - 2 - Outline Ø Binary Search Trees Ø AVL Trees Ø Splay Trees - 3 - Binary Search Trees Ø A binary search

More information

Computational Models - Lecture 4 1

Computational Models - Lecture 4 1 Computational Models - Lecture 4 1 Handout Mode Iftach Haitner and Yishay Mansour. Tel Aviv University. April 3/8, 2013 1 Based on frames by Benny Chor, Tel Aviv University, modifying frames by Maurice

More information

Declarative Programming Techniques

Declarative Programming Techniques Declarative Programming Techniques Accumulators and Difference Lists (CTM 3.4.3-3.4.4) Carlos Varela RPI Adapted with permission from: Seif Haridi KTH Peter Van Roy UCL February 12, 2015 C. Varela; Adapted

More information

Jacques Fleuriot. Automated Reasoning Rewrite Rules

Jacques Fleuriot. Automated Reasoning Rewrite Rules Automated Reasoning Rewrite Rules Jacques Fleuriot Lecture 8, page 1 Term Rewriting Rewriting is a technique for replacing terms in an expression with equivalent terms useful for simplification, e.g. given

More information

Data Structures 1 NTIN066

Data Structures 1 NTIN066 Data Structures 1 NTIN066 Jirka Fink Department of Theoretical Computer Science and Mathematical Logic Faculty of Mathematics and Physics Charles University in Prague Winter semester 2017/18 Last change

More information

CHAPTER 2 BOOLEAN ALGEBRA

CHAPTER 2 BOOLEAN ALGEBRA CHAPTER 2 BOOLEAN ALGEBRA This chapter in the book includes: Objectives Study Guide 2.1 Introduction 2.2 Basic Operations 2.3 Boolean Expressions and Truth Tables 2.4 Basic Theorems 2.5 Commutative, Associative,

More information

Fundamental Properties of Lambda-calculus

Fundamental Properties of Lambda-calculus Fundamental Properties of Lambda-calculus Tobias Nipkow Stefan Berghofer February 12, 2013 Contents 1 Basic definitions of Lambda-calculus 4 1.1 Lambda-terms in de Bruijn notation and substitution....

More information

Finite Automata Theory and Formal Languages TMV027/DIT321 LP4 2018

Finite Automata Theory and Formal Languages TMV027/DIT321 LP4 2018 Finite Automata Theory and Formal Languages TMV027/DIT321 LP4 2018 Lecture 4 Ana Bove March 23rd 2018 Recap: Formal Proofs How formal should a proof be? Depends on its purpose...... but should be convincing......

More information

Algorithm Design and Analysis

Algorithm Design and Analysis Algorithm Design and Analysis LECTURE 7 Greedy Graph Algorithms Shortest paths Minimum Spanning Tree Sofya Raskhodnikova 9/14/016 S. Raskhodnikova; based on slides by E. Demaine, C. Leiserson, A. Smith,

More information

CSE 505, Fall 2009, Midterm Examination 5 November Please do not turn the page until everyone is ready.

CSE 505, Fall 2009, Midterm Examination 5 November Please do not turn the page until everyone is ready. CSE 505, Fall 2009, Midterm Examination 5 November 2009 Please do not turn the page until everyone is ready Rules: The exam is closed-book, closed-note, except for one side of one 85x11in piece of paper

More information

EDA045F: Program Analysis LECTURE 10: TYPES 1. Christoph Reichenbach

EDA045F: Program Analysis LECTURE 10: TYPES 1. Christoph Reichenbach EDA045F: Program Analysis LECTURE 10: TYPES 1 Christoph Reichenbach In the last lecture... Performance Counters Challenges in Dynamic Performance Analysis Taint Analysis Binary Instrumentation 2 / 44 Types

More information

Prime and Semiprime Bi-ideals in Ordered Semigroups

Prime and Semiprime Bi-ideals in Ordered Semigroups International Journal of Algebra, Vol. 7, 2013, no. 17, 839-845 HIKARI Ltd, www.m-hikari.com http://dx.doi.org/10.12988/ija.2013.310105 Prime and Semiprime Bi-ideals in Ordered Semigroups R. Saritha Department

More information

Tutorial 4. Dynamic Set: Amortized Analysis

Tutorial 4. Dynamic Set: Amortized Analysis Tutorial 4 Dynamic Set: Amortized Analysis Review Binary tree Complete binary tree Full binary tree 2-tree P115 Unlike common binary tree, the base case is not an empty tree, but a external node Heap Binary

More information

CHAPTER 2 PROBABILITY

CHAPTER 2 PROBABILITY CHAPTER 2 PROBABILITY 2. a. S = {(X,X), (X,Y), (X,Z), (Y,X), (Y,Y), (Y,Z), (Z,X), (Z,Y), (Z,Z)} b. A = {(X,X), (Y,Y), (Z,Z)} c. B = {(X,X), (X,Z), (Z,X), (Z,Z)} 4.a. A = { RRR, LLL, SSS } b. B = { RLS,

More information

Connected Graphs and Spanning Trees

Connected Graphs and Spanning Trees Connected Graphs and Spanning Trees GAINA, Daniel January 9, 2014 1 / 22 Describing the problem I G = (V, E) - graph 1 V - set of vertices 2 E - (multi)set of edges Example: 1 2 3 4 5 6 7 8 9 V = {1,...,

More information

Kleene Algebra with Tests and Demonic Refinement Algebras

Kleene Algebra with Tests and Demonic Refinement Algebras Kleene Algebra with Tests and Demonic Refinement Algebras Alasdair Armstrong Victor B. F. Gomes Georg Struth August 28, 2014 Abstract We formalise Kleene algebra with tests (KAT) and demonic refinement

More information

Chapter 2: Switching Algebra and Logic Circuits

Chapter 2: Switching Algebra and Logic Circuits Chapter 2: Switching Algebra and Logic Circuits Formal Foundation of Digital Design In 1854 George Boole published An investigation into the Laws of Thoughts Algebraic system with two values 0 and 1 Used

More information

TAFL 1 (ECS-403) Unit- III. 3.1 Definition of CFG (Context Free Grammar) and problems. 3.2 Derivation. 3.3 Ambiguity in Grammar

TAFL 1 (ECS-403) Unit- III. 3.1 Definition of CFG (Context Free Grammar) and problems. 3.2 Derivation. 3.3 Ambiguity in Grammar TAFL 1 (ECS-403) Unit- III 3.1 Definition of CFG (Context Free Grammar) and problems 3.2 Derivation 3.3 Ambiguity in Grammar 3.3.1 Inherent Ambiguity 3.3.2 Ambiguous to Unambiguous CFG 3.4 Simplification

More information

Chapter 1 Basic (Elementary) Inequalities and Their Application

Chapter 1 Basic (Elementary) Inequalities and Their Application Chapter 1 Basic (Elementary) Inequalities and Their Application There are many trivial facts which are the basis for proving inequalities. Some of them are as follows: 1. If x y and y z then x z, for any

More information

The Tortoise and the Hare Algorithm

The Tortoise and the Hare Algorithm The Tortoise and the Hare Algorithm Peter Gammie April 17, 2016 Abstract We formalize the Tortoise and Hare cycle-finding algorithm ascribed to Floyd by Knuth (1981, p7, exercise 6), and an improved version

More information

Breadth First Search, Dijkstra s Algorithm for Shortest Paths

Breadth First Search, Dijkstra s Algorithm for Shortest Paths CS 374: Algorithms & Models of Computation, Spring 2017 Breadth First Search, Dijkstra s Algorithm for Shortest Paths Lecture 17 March 1, 2017 Chandra Chekuri (UIUC) CS374 1 Spring 2017 1 / 42 Part I Breadth

More information

Analysis of Algorithms. Outline 1 Introduction Basic Definitions Ordered Trees. Fibonacci Heaps. Andres Mendez-Vazquez. October 29, Notes.

Analysis of Algorithms. Outline 1 Introduction Basic Definitions Ordered Trees. Fibonacci Heaps. Andres Mendez-Vazquez. October 29, Notes. Analysis of Algorithms Fibonacci Heaps Andres Mendez-Vazquez October 29, 2015 1 / 119 Outline 1 Introduction Basic Definitions Ordered Trees 2 Binomial Trees Example 3 Fibonacci Heap Operations Fibonacci

More information

Search Trees. EECS 2011 Prof. J. Elder Last Updated: 24 March 2015

Search Trees. EECS 2011 Prof. J. Elder Last Updated: 24 March 2015 Search Trees < 6 2 > 1 4 = 8 9-1 - Outline Ø Binary Search Trees Ø AVL Trees Ø Splay Trees - 2 - Learning Outcomes Ø From this lecture, you should be able to: q Define the properties of a binary search

More information

Graduate Analysis of Algorithms Dr. Haim Levkowitz

Graduate Analysis of Algorithms Dr. Haim Levkowitz UMass Lowell Computer Science 9.53 Graduate Analysis of Algorithms Dr. Haim Levkowitz Fall 27 Lecture 5 Tuesday, 2 Oct 27 Amortized Analysis Overview Amortize: To pay off a debt, usually by periodic payments

More information

Incorrect reasoning about RL. Equivalence of NFA, DFA. Epsilon Closure. Proving equivalence. One direction is easy:

Incorrect reasoning about RL. Equivalence of NFA, DFA. Epsilon Closure. Proving equivalence. One direction is easy: Incorrect reasoning about RL Since L 1 = {w w=a n, n N}, L 2 = {w w = b n, n N} are regular, therefore L 1 L 2 = {w w=a n b n, n N} is regular If L 1 is a regular language, then L 2 = {w R w L 1 } is regular,

More information

A Simple Implementation Technique for Priority Search Queues

A Simple Implementation Technique for Priority Search Queues A Simple Implementation Technique for Priority Search Queues RALF HINZE Institute of Information and Computing Sciences Utrecht University Email: ralf@cs.uu.nl Homepage: http://www.cs.uu.nl/~ralf/ April,

More information

Positional Determinacy of Parity Games

Positional Determinacy of Parity Games Positional Determinacy of Parity Games Christoph Dittmann christoph.dittmann@tu-berlin.de April 17, 2016 We present a formalization of parity games (a two-player game on directed graphs) and a proof of

More information

Trees. A tree is a graph which is. (a) Connected and. (b) has no cycles (acyclic).

Trees. A tree is a graph which is. (a) Connected and. (b) has no cycles (acyclic). Trees A tree is a graph which is (a) Connected and (b) has no cycles (acyclic). 1 Lemma 1 Let the components of G be C 1, C 2,..., C r, Suppose e = (u, v) / E, u C i, v C j. (a) i = j ω(g + e) = ω(g).

More information

Algorithms. Jordi Planes. Escola Politècnica Superior Universitat de Lleida

Algorithms. Jordi Planes. Escola Politècnica Superior Universitat de Lleida Algorithms Jordi Planes Escola Politècnica Superior Universitat de Lleida 2016 Syllabus What s been done Formal specification Computational Cost Transformation recursion iteration Divide and conquer Sorting

More information

Polynomial Factorization

Polynomial Factorization Polynomial Factorization René Thiemann and Akihisa Yamada October 11, 2017 Abstract Based on existing libraries for polynomial interpolation and matrices, we formalized several factorization algorithms

More information

Chapter 6 Cardinal Numbers and Cardinal Arithmetic

Chapter 6 Cardinal Numbers and Cardinal Arithmetic Principles of Mathematics (Math 2450) A Ë@ Õæ Aë áöß @. X. @ 2014-2015 Èð B@ É Ë@ Chapter 6 Cardinal Numbers and Cardinal Arithmetic In this chapter we introduce the concept of cardinal numbers. The properties

More information