VerifyThis @ FM 2012, problem 2
Computes prefix sums of an array, in-place
Authors: Claude Marché / François Bobot
Topics: Array Data Structure / Trees / Ghost code
Tools: Why3
References: VerifyThis @ FM2012
see also the index (by topic, by tool, by reference, by year)
download ZIP archive
The VerifyThis competition at FM2012: Problem 2
see this web page
module PrefixSumRec use int.Int use int.ComputerDivision use int.Power use map.Map use array.Array use array.ArraySum
Preliminary lemmas on division by 2 and power of 2
lemma Div_mod_2: forall x:int. x >= 0 -> x >= 2 * div x 2 >= x-1
Needed for the proof of phase1_frame and phase1_frame2
predicate is_power_of_2 (x:int) = exists k:int. (k >= 0 /\ x = power 2 k)
The step must be a power of 2
(* needed *) lemma is_power_of_2_1: forall x:int. is_power_of_2 x -> x > 1 -> 2 * div x 2 = x
Modeling the "upsweep" phase
function go_left (left right:int) : int = let space = right - left in left - div space 2
Shortcuts
function go_right (left right:int) : int = let space = right - left in right - div space 2
Description in a purely logic way the effect of the first phase "upsweep" of the algorithm. The second phase "downsweep" then traverses the array in the same way than the first phase. Hence, the inductive nature of this definition is not an issue.
inductive phase1 (left:int) (right:int) (a0: array int) (a: array int) = | Leaf : forall left right:int, a0 a : array int. right = left + 1 -> a[left] = a0[left] -> phase1 left right a0 a | Node : forall left right:int, a0 a : array int. right > left + 1 -> phase1 (go_left left right) left a0 a -> phase1 (go_right left right) right a0 a -> a[left] = sum a0 (left-(right-left)+1) (left+1) -> phase1 left right a0 a
Frame properties of the "phase1" inductive
let rec lemma phase1_frame (left right:int) (a0 a a' : array int) : unit requires { forall i:int. left-(right-left) < i < right -> a[i] = a'[i]} requires { phase1 left right a0 a } variant { right-left } ensures { phase1 left right a0 a' } = if right > left + 1 then begin phase1_frame (go_left left right) left a0 a a'; phase1_frame (go_right left right) right a0 a a' end
frame lemma for "phase1" on fourth argument. needed to prove both upsweep, downsweep and compute_sums
let rec lemma phase1_frame2 (left right:int) (a0 a0' a : array int) : unit requires { forall i:int. left-(right-left) < i < right -> a0[i] = a0'[i]} requires { phase1 left right a0 a } variant { right-left } ensures { phase1 left right a0' a } = if right > left + 1 then begin phase1_frame2 (go_left left right) left a0 a0' a; phase1_frame2 (go_right left right) right a0 a0' a end
frame lemma for "phase1" on third argument. needed to prove upsweep and compute_sums
The upsweep phase
First function: modify partially the table and compute some intermediate partial sums
let rec upsweep (left right: int) (a: array int) requires { 0 <= left < right < a.length } requires { -1 <= left - (right - left) } requires { is_power_of_2 (right - left) } variant { right - left } ensures { phase1 left right (old a) a } ensures { let space = right - left in a[right] = sum (old a) (left-space+1) (right+1) /\ (forall i: int. i <= left-space -> a[i] = (old a)[i]) /\ (forall i: int. i > right -> a[i] = (old a)[i]) } = let space = right - left in if right > left+1 then begin upsweep (left - div space 2) left a; upsweep (right - div space 2) right a; assert { phase1 (left - div space 2) left (old a) a }; assert { phase1 (right - div space 2) right (old a) a }; assert { a[left] = sum (old a) (left-(right-left)+1) (left+1) }; assert { a[right] = sum (old a) (left+1) (right+1) } end; a[right] <- a[left] + a[right]; assert { right > left+1 -> phase1 (left - div space 2) left (old a) a }; assert { right > left+1 -> phase1 (right - div space 2) right (old a) a }
The downsweep phase
predicate partial_sum (left:int) (right:int) (a0 a : array int) = forall i : int. (left-(right-left)) < i <= right -> a[i] = sum a0 0 i
The property we ultimately want to prove
let rec downsweep (left right: int) (ghost a0 : array int) (a : array int) requires { 0 <= left < right < a.length } requires { -1 <= left - (right - left) } requires { is_power_of_2 (right - left) } requires { a[right] = sum a0 0 (left-(right-left) + 1) } requires { phase1 left right a0 a } variant { right - left } ensures { partial_sum left right a0 a } ensures { forall i: int. i <= left-(right-left) -> a[i] = (old a)[i] } ensures { forall i: int. i > right -> a[i] = (old a)[i] } = let tmp = a[right] in assert { a[right] = sum a0 0 (left-(right-left) + 1) }; assert { a[left] = sum a0 (left-(right-left)+1) (left+1) }; a[right] <- a[right] + a[left]; a[left] <- tmp; assert { a[right] = sum a0 0 (left + 1) }; if right > left+1 then let space = right - left in assert { phase1 (go_left left right) left a0 (old a) }; assert { phase1 (go_right left right) right a0 (old a) }; assert { phase1 (go_right left right) right a0 a }; downsweep (left - div space 2) left a0 a; assert { phase1 (go_right left right) right a0 a }; downsweep (right - div space 2) right a0 a; assert { partial_sum (left - div space 2) left a0 a }; assert { partial_sum (right - div space 2) right a0 a }
Second function: complete the partial using the remaining intial value and the partial sum already computed
The main procedure
let compute_sums a requires { a.length >= 2 } requires { is_power_of_2 a.length } ensures { forall i : int. 0 <= i < a.length -> a[i] = sum (old a) 0 i } = let a0 = ghost (copy a) in let l = a.length in let left = div l 2 - 1 in let right = l - 1 in upsweep left right a; (* needed for the precondition of downsweep *) assert { phase1 left right a0 a }; a[right] <- 0; downsweep left right a0 a; (* needed to prove the post-condition *) assert { forall i : int. left-(right-left) < i <= right -> a[i] = sum a0 0 i }
A simple test
let test_harness () = let a = make 8 0 in (* needed for the precondition of compute_sums *) assert { power 2 3 = a.length }; a[0] <- 3; a[1] <- 1; a[2] <- 7; a[3] <- 0; a[4] <- 4; a[5] <- 1; a[6] <- 6; a[7] <- 3; compute_sums a; assert { a[0] = 0 }; assert { a[1] = 3 }; assert { a[2] = 4 }; assert { a[3] = 11 }; assert { a[4] = 11 }; assert { a[5] = 15 }; assert { a[6] = 16 }; assert { a[7] = 22 } exception BenchFailure let bench () raises { BenchFailure -> true } = let a = make 8 0 in (* needed for the precondition of compute_sums *) assert { power 2 3 = a.length }; a[0] <- 3; a[1] <- 1; a[2] <- 7; a[3] <- 0; a[4] <- 4; a[5] <- 1; a[6] <- 6; a[7] <- 3; compute_sums a; if a[0] <> 0 then raise BenchFailure; if a[1] <> 3 then raise BenchFailure; if a[2] <> 4 then raise BenchFailure; if a[3] <> 11 then raise BenchFailure; if a[4] <> 11 then raise BenchFailure; if a[5] <> 15 then raise BenchFailure; if a[6] <> 16 then raise BenchFailure; if a[7] <> 22 then raise BenchFailure; a end
Why3 Proof Results for Project "verifythis_PrefixSumRec"
Theory "verifythis_PrefixSumRec.PrefixSumRec": fully verified
Obligations | Alt-Ergo 2.0.0 | Alt-Ergo 2.3.0 | CVC3 2.4.1 | CVC4 1.4 | CVC4 1.5 | CVC4 1.6 | Z3 3.2 | Z3 4.3.1 | Z3 4.3.2 | |
Div_mod_2 | 0.01 | --- | 0.02 | 0.03 | --- | --- | 0.02 | --- | 0.02 | |
is_power_of_2_1 | 0.01 | --- | 0.03 | 0.02 | --- | --- | --- | --- | --- | |
VC for phase1_frame | --- | --- | --- | --- | --- | --- | --- | --- | --- | |
split_goal_right | ||||||||||
variant decrease | 0.02 | --- | --- | 0.02 | --- | --- | --- | --- | --- | |
precondition | 0.02 | --- | --- | 0.03 | --- | --- | --- | --- | 0.01 | |
precondition | 0.01 | --- | --- | 0.02 | --- | --- | --- | --- | 0.02 | |
variant decrease | 0.01 | --- | --- | 0.02 | --- | --- | --- | --- | 0.01 | |
precondition | 0.01 | --- | --- | 0.03 | --- | --- | --- | --- | 0.01 | |
precondition | --- | --- | 0.04 | --- | --- | --- | --- | --- | --- | |
postcondition | --- | 0.24 | --- | --- | --- | --- | --- | --- | --- | |
VC for phase1_frame2 | --- | --- | --- | --- | --- | --- | --- | --- | --- | |
split_goal_right | ||||||||||
variant decrease | 0.02 | --- | --- | 0.04 | --- | --- | --- | --- | 0.01 | |
precondition | --- | --- | 0.03 | --- | --- | --- | 0.02 | --- | --- | |
precondition | 0.01 | --- | --- | 0.02 | --- | --- | --- | --- | 0.01 | |
variant decrease | 0.01 | --- | --- | 0.02 | --- | --- | --- | --- | --- | |
precondition | 0.02 | --- | --- | 0.03 | --- | --- | --- | --- | 0.01 | |
precondition | 0.02 | --- | --- | 0.03 | --- | --- | --- | --- | --- | |
postcondition | --- | --- | --- | --- | --- | --- | --- | 0.04 | --- | |
VC for upsweep | --- | --- | --- | --- | --- | --- | --- | --- | --- | |
split_goal_right | ||||||||||
precondition | 0.04 | --- | 0.12 | 0.05 | --- | --- | --- | --- | --- | |
variant decrease | 0.02 | --- | 0.01 | --- | --- | --- | 0.09 | --- | 0.11 | |
precondition | 0.10 | --- | 0.01 | 0.05 | --- | --- | --- | --- | --- | |
precondition | 0.01 | --- | 0.00 | 0.02 | --- | --- | 0.00 | --- | 0.00 | |
precondition | 0.01 | --- | 0.01 | 0.01 | --- | --- | 8.62 | --- | --- | |
precondition | --- | --- | --- | --- | 0.03 | --- | --- | --- | --- | |
variant decrease | --- | --- | 0.02 | --- | --- | --- | 0.06 | --- | --- | |
precondition | --- | --- | 0.02 | --- | --- | --- | --- | --- | --- | |
precondition | 0.03 | --- | 0.06 | 0.04 | --- | --- | --- | --- | --- | |
precondition | 0.02 | --- | 0.18 | 0.02 | --- | --- | 10.72 | --- | --- | |
assertion | 0.02 | --- | --- | --- | --- | --- | --- | --- | --- | |
assertion | 0.02 | --- | --- | --- | --- | --- | --- | --- | --- | |
assertion | 0.00 | --- | 0.01 | --- | --- | --- | 4.00 | --- | 0.03 | |
assertion | --- | --- | 2.61 | --- | --- | --- | --- | --- | --- | |
index in array bounds | 0.02 | --- | 0.02 | 0.02 | --- | --- | 0.02 | --- | 0.02 | |
index in array bounds | 0.02 | --- | 0.02 | 0.02 | --- | --- | 0.02 | --- | 0.03 | |
index in array bounds | 0.01 | --- | --- | 0.07 | --- | --- | --- | --- | --- | |
assertion | 0.01 | --- | 0.34 | 0.02 | --- | --- | 0.15 | --- | 0.22 | |
assertion | 0.01 | --- | 0.14 | 0.12 | --- | --- | 0.01 | --- | 0.03 | |
postcondition | --- | --- | 0.24 | --- | --- | --- | 0.06 | --- | --- | |
postcondition | 1.13 | --- | --- | 0.11 | --- | --- | --- | --- | --- | |
index in array bounds | 0.04 | --- | --- | 0.09 | --- | --- | --- | --- | --- | |
index in array bounds | 0.01 | --- | 0.02 | 0.02 | --- | --- | --- | --- | --- | |
index in array bounds | 0.02 | --- | 0.02 | 0.01 | --- | --- | 0.02 | --- | 0.05 | |
assertion | 0.00 | --- | 0.04 | 0.07 | --- | --- | --- | --- | --- | |
assertion | 0.00 | --- | 0.00 | 0.01 | --- | --- | 0.02 | --- | 0.02 | |
postcondition | --- | 0.01 | --- | --- | --- | --- | --- | --- | --- | |
postcondition | 0.85 | --- | 0.01 | 0.02 | --- | --- | --- | --- | --- | |
VC for downsweep | --- | --- | --- | --- | --- | --- | --- | --- | --- | |
split_goal_right | ||||||||||
index in array bounds | 0.01 | --- | 0.03 | 0.01 | --- | --- | 0.02 | --- | 0.04 | |
assertion | 0.03 | --- | 0.02 | 0.02 | --- | --- | --- | --- | --- | |
assertion | --- | --- | --- | 0.06 | --- | --- | 0.05 | --- | --- | |
index in array bounds | 0.01 | --- | 0.02 | 0.02 | --- | --- | 0.03 | --- | 0.03 | |
index in array bounds | 0.00 | --- | 0.02 | 0.07 | --- | --- | --- | --- | --- | |
index in array bounds | 0.04 | --- | 0.02 | 0.02 | --- | --- | --- | --- | --- | |
index in array bounds | 0.08 | --- | --- | 0.02 | --- | --- | 0.02 | --- | --- | |
assertion | 0.33 | --- | --- | 0.02 | --- | --- | 4.22 | --- | --- | |
assertion | 0.01 | --- | 0.02 | 0.01 | --- | --- | 0.00 | --- | 0.00 | |
assertion | 0.03 | --- | 0.03 | 0.03 | --- | --- | 0.04 | --- | 0.04 | |
assertion | 0.01 | --- | 3.42 | 0.01 | --- | --- | 6.42 | --- | 0.69 | |
precondition | 0.02 | --- | 0.03 | 0.03 | --- | --- | --- | --- | --- | |
variant decrease | 0.06 | --- | --- | --- | --- | --- | --- | --- | --- | |
precondition | 0.10 | --- | 0.09 | 0.06 | --- | --- | --- | --- | --- | |
precondition | 0.11 | --- | 0.02 | 0.03 | --- | --- | 0.02 | --- | --- | |
precondition | 0.21 | --- | --- | --- | --- | 0.14 | 4.91 | --- | --- | |
precondition | --- | --- | --- | 0.02 | --- | --- | 0.04 | --- | --- | |
precondition | --- | --- | --- | --- | --- | --- | --- | 0.75 | --- | |
assertion | --- | --- | 0.08 | --- | --- | --- | 0.12 | --- | --- | |
precondition | 0.02 | --- | 0.12 | --- | --- | --- | --- | --- | --- | |
variant decrease | 0.11 | --- | 0.03 | 0.04 | --- | --- | 0.08 | --- | 0.11 | |
precondition | 0.02 | --- | 0.06 | 0.02 | --- | --- | 0.02 | --- | 0.03 | |
precondition | --- | --- | --- | --- | 0.02 | --- | --- | --- | --- | |
precondition | 0.03 | --- | --- | 0.13 | --- | --- | 7.04 | --- | --- | |
precondition | 0.03 | --- | 0.04 | 0.02 | --- | --- | 0.02 | --- | 0.03 | |
precondition | 0.02 | --- | 0.02 | --- | --- | --- | 0.04 | --- | --- | |
assertion | --- | --- | --- | --- | 0.02 | --- | --- | --- | --- | |
assertion | --- | --- | 0.02 | 0.02 | --- | --- | --- | --- | --- | |
postcondition | 0.04 | --- | 0.68 | 0.04 | --- | --- | --- | --- | --- | |
postcondition | --- | --- | --- | 2.63 | --- | --- | --- | --- | --- | |
postcondition | 0.02 | --- | --- | --- | --- | --- | --- | --- | --- | |
postcondition | 0.01 | --- | 0.36 | 0.18 | --- | --- | --- | --- | --- | |
postcondition | 0.00 | --- | 0.06 | 0.01 | --- | --- | 0.01 | --- | 0.12 | |
postcondition | 0.12 | --- | --- | 0.09 | --- | --- | --- | --- | --- | |
VC for compute_sums | --- | --- | --- | --- | --- | --- | --- | --- | --- | |
split_goal_right | ||||||||||
precondition | --- | --- | 0.01 | 0.02 | --- | --- | --- | --- | --- | |
precondition | 0.02 | --- | 0.02 | 0.01 | --- | --- | 1.13 | --- | 0.03 | |
precondition | 0.02 | --- | 0.03 | 0.02 | --- | --- | 0.02 | --- | 0.03 | |
precondition | --- | --- | --- | 0.04 | --- | --- | --- | --- | --- | |
assertion | 0.02 | --- | --- | 0.04 | --- | --- | --- | --- | --- | |
index in array bounds | 0.02 | --- | 0.02 | 0.01 | --- | --- | 0.03 | --- | 0.03 | |
precondition | --- | --- | 0.03 | 0.04 | --- | --- | 0.02 | --- | --- | |
precondition | --- | --- | --- | --- | 0.02 | --- | --- | --- | --- | |
precondition | --- | --- | 0.06 | 0.08 | --- | --- | --- | --- | --- | |
precondition | 0.09 | --- | 0.03 | 0.04 | --- | --- | --- | --- | --- | |
precondition | 0.02 | --- | 0.01 | 0.01 | --- | --- | 0.68 | --- | --- | |
assertion | 0.02 | --- | 0.03 | 0.02 | --- | --- | 0.07 | --- | 0.05 | |
postcondition | --- | --- | --- | 7.95 | --- | --- | --- | --- | --- | |
VC for test_harness | --- | --- | --- | --- | --- | --- | --- | --- | --- | |
split_goal_right | ||||||||||
array creation size | 0.02 | --- | 0.01 | 0.01 | --- | --- | 0.00 | --- | 0.00 | |
assertion | 0.04 | --- | 0.03 | 0.03 | --- | --- | 0.05 | --- | --- | |
index in array bounds | 0.03 | --- | 0.02 | 0.01 | --- | --- | 0.02 | --- | 0.04 | |
index in array bounds | 0.01 | --- | 0.01 | 0.02 | --- | --- | 0.00 | --- | 0.00 | |
index in array bounds | 0.01 | --- | 0.01 | 0.01 | --- | --- | 0.00 | --- | 0.00 | |
index in array bounds | 0.01 | --- | 0.01 | 0.02 | --- | --- | 0.00 | --- | 0.00 | |
index in array bounds | 0.04 | --- | 0.03 | 0.03 | --- | --- | --- | --- | --- | |
index in array bounds | 0.02 | --- | 0.03 | 0.02 | --- | --- | 0.02 | --- | 0.03 | |
index in array bounds | 0.01 | --- | 0.02 | 0.01 | --- | --- | 0.00 | --- | 0.00 | |
index in array bounds | 0.01 | --- | 0.01 | 0.01 | --- | --- | 0.00 | --- | 0.00 | |
precondition | 0.04 | --- | 0.02 | 0.05 | --- | --- | --- | --- | --- | |
precondition | 0.04 | --- | 0.02 | 0.03 | --- | --- | --- | --- | --- | |
assertion | 0.04 | --- | 0.02 | 0.03 | --- | --- | --- | --- | --- | |
assertion | 0.02 | --- | 0.01 | 0.00 | --- | --- | 0.18 | --- | 0.94 | |
assertion | 0.07 | --- | 0.02 | 0.04 | --- | --- | 0.32 | --- | 5.31 | |
assertion | 0.04 | --- | 0.03 | 0.06 | --- | --- | 0.33 | --- | 5.27 | |
assertion | 0.04 | --- | 0.03 | 0.04 | --- | --- | 0.24 | --- | 4.98 | |
assertion | 0.01 | --- | 0.01 | 0.12 | --- | --- | 0.34 | --- | 5.28 | |
assertion | 0.02 | --- | 0.02 | 0.13 | --- | --- | 0.34 | --- | 5.33 | |
assertion | 0.00 | --- | 0.02 | 0.15 | --- | --- | 0.18 | --- | 0.84 | |
VC for bench | 3.47 | --- | --- | --- | --- | --- | --- | --- | --- |