Since we are interested in using Coq for low-levels programs,
we choosed to realize the specification of log2 with
a while loop using shift
(multiplication by 2) and unshift
(quotient by 2, and parity check).
See the module shift
.
Definition log2_spec:=[n0:nat] {l:nat & {n0=(two_power l)}+ {(lt (two_power l) n0) /\ (lt n0 (two_power (S l)))}}.We want to prove the lemma:
log2_impl: (n0:nat)(lt O n0)->(log2_spec n0).
In order to do that, we used the imperative and while modules.
C
-like
syntax:
{ /* computation of the integer logarithm of n0 (base 2) */ n := n0; p:= O; b := true; while (n>O) { if (b) { /* n0=2^p * n */ if(n is even) { p:=p+1; n:=unshift(n); b:=true} else {p:=p+1; n:=unshift(n); b:=false} } else /* 2^p * < n0 < 2^p(n+1) /\ n>0 */ { p:=p+1; n:=unshift(n); b:=false} } return (p,b); }
n,p:nat
, and
b:bool
.
It is straightforward to define a Set
:
Record state:Set := mkstate { state_n:nat; state_p:nat; state_b:bool}.
ok
if
Inductive ok:state->Prop:= oktrue: (p:nat)n0=(two_power p)-> (ok (mkstate one p true)) | okfalse: (p:nat)(lt (two_power p) n0)-> (lt n0 (two_power (S p))) -> (ok (mkstate one p false)).
(n,p,b)
:
Local init:=[s:state](state_n s)=n0 /\ (state_p s)=O /\ (state_b s)=true.
n
reach the value 1:
Inductive final:state->Prop := mkfinal:(p:nat)(b:bool)(final (mkstate one p b)).
(n,p,b)
, the boolean b
expresses
the exactness of the logarithm to compute:
Inductive invar:state->Prop := exact:(n,p:nat) n0=(mult (two_power p) n)-> (lt O n)-> (invar (mkstate n p true)) |inexact:(n,p:nat) (lt (mult (two_power p)n) n0)-> (lt n0 (mult (two_power p) (S n)))-> (lt O n)-> (invar (mkstate n p false)).
ltof=[A:Set][f:A->nat][a,b:A](lt (f a) (f b))defined in the module ARITH/Wf_nat of the distribution. The theorem well_founded_ltof ensures us that the order
stat_order:=(ltof state state_n)is well founded.
while
loop is realized by:
Realizer [s:state] <bool>if (le_lt_eq_dec (S O) (state_n s)) then false else true
while
loop is realized by:
Realizer [s:state <state>if (state_b s) then <state>let (m:nat;b':bool)=(Unshift (state_n s)) in <state> if b' then (mkstate m (S (state_p s)) true) else (mkstate m (S (state_p s)) false) else <state>let (m:nat;b:bool)=(Unshift (state_n s)) in (mkstate m (S (state_p s)) false)
log2_spec
.
This is done with a
Realizer [a:state] <(prod nat bool)>Case (state_b a) of ((state_p a),true) ((state_p a),false) end
Program
,
and a lot of logical goals as usual.
return (p,b);
of the C-like program
is tranlated into:
<(prod nat bool)>Case (state_b a) of ((state_p a),true) ((state_p a),false) endand not into:
((state_p a),(state_b a))We intented to do this way, but we failed.
Inversion_clear
command:
invar_inv1:(n,p:nat)(invar (mkstate n p true))-> n0=(mult (two_power p) n). invar_inv2: (n,p:nat)(invar (mkstate n p false))-> (lt (mult (two_power p)n) n0) /\ (lt n0 (mult (two_power p) (S n))).etc ...
Realizer {nat x y; while (zerop x) { ... }; return y;}We hope this will be done in few weeks.
Demo.ml
, and open
it.
You will notice that the extraction algorithm generated a function of
type int -> (int, sumbool) sigS
and not of
int -> int*bool
.
In consequence, the result has the form existS(p,left)
for an exact logarithm, and existS(p,right)
for an
inexact (default) logarithm.
#load "Demo.ml";; #open "Demo";; #log2_impl;; - : int -> (int, sumbool) sigS =#log2_impl 1023;; - : (int, sumbool) sigS = existS (9, right) #log2_impl 1024;; - : (int, sumbool) sigS = existS (10, left) #log2_impl 1025;; - : (int, sumbool) sigS = existS (10, right)