develop
to validate the chain
generation algorithm specified in
spec
.
power_algo: (n:nat)(lt O n)->(addchain_spec n).To do that, we first prove a lemma:
chain_algo: (c:Call)(conform c)->(gencode c).and use the inductive definition of
addchain_spec
(by a
Split
).
chain_algo
We have in mind a proof of the form:
Realizer <Code>rec REC :: :: {Call_lt} [c:Call] <Code>Case c of (* Call_M p *) [p:nat] (COND3 Code p M3:: ::{(gencode (Call_M three))} [q:nat](M2q q) [q:nat] (C2M p q (REC (Call_C p q)))) (* Call_C p n0 *) [p,n0:nat] <Code>let (q,r:nat)=(eucl_dev n0 p) in <Code>if (zerop r) then (MMC n0 q (REC (Call_M n0)) (REC (Call_M q))) else (KMC n0 q r (REC (Call_K n0 r)) (REC (Call_M q))) (* Call_K p n0 *) [p,n0:nat] <Code>let (q,r:nat)= (eucl_dev n0 p) in <Code>if (zerop r) then (MMK n0 q (REC (Call_M n0)) (REC (Call_M q))) else (KMK n0 q r (REC (Call_K n0 r)) (REC (Call_M q))) end. Program_all. ... Qed.
Call
,
which ensures the termination of the code generation algorithm.
We give a definition which allows to build this algorithm from the lemmas
of develop
.
Inductive Call_lt:Call->Call->Prop := M_M_lt:(p,q:nat)(lt p q)-> (Call_lt (Call_M p) (Call_M q)) | C_M_lt:(p,q:nat)(lt p q)-> (Call_lt (Call_C q p) (Call_M q)) | M_C_lt:(p,q,r:nat)(lt p r)-> (Call_lt (Call_M p) (Call_C r q)) | M_K_lt:(p,q,r:nat)(le p r)-> (Call_lt (Call_M p) (Call_K r q)) | K_C_lt:(p,q,r,s:nat)(lt q s)-> (Call_lt (Call_K q p) (Call_C s r)) | K_K_lt:(p,q,r,s:nat)(lt q s)-> (Call_lt (Call_K q p) (Call_K s r)).
Call_lt
is well founded Call_lt
,
we use the lemma well_founded_ltof
of the library
ARITH
: we define an integer measure Call_measure
and use the lemma wf_coarser
of the module
Wf_compl
.
Definition Call_measure:= [c:Call] <nat>Case c of (* Call_M n *) [n:nat](S (mult three n)) (* Call_C n p *) [n,p:nat](mult three n) (* Call_K n p *) [n,p:nat](S (S (mult three n))) end.
This technique (using an integer measure to prove well-foundedness) is
classical, perhaps it will be interesting to have a more direct proof,
based directly on the definition of Call_lt
.
(Call_M p)
develop
tell us how to build code for computing x^p:
M3
),
M2q
),
C2M
)
In order to build a switch between these 3 cases, we begin to prove:
COND3:(A:nat->Set)(p:nat) (p=three ->(A p))-> ((q:nat)(p=(two_power q))->(A p))-> ((q:nat)(lt q p )->(le two q)->(A p))-> (lt O p)-> (A p).such that a
Real
term of the form:
(Cond3 A p case3 case_2q case_gamma)has the intuitive meaning:
if p = 3 then case3 else if p is the q_th power of 2, then (case_2q q) else (case_gamma (gamma p))With all these notations, the part of the
chain_algo
realizer dealing with (Call_M p)
is:
(* Call_M p *) [p:nat] (COND3 Code p M3:: ::{(gencode (Call_M three))} [q:nat](M2q q) [q:nat] (C2M p q (REC (Call_C p q))))It should be noticed that the piece of code:
if p is the q-th power of 2 then ... else ...is realized with the help of a logarithm function
log2_r
, specified in
log2_spec.v
.
In the module
log2_implementation
, we propose a realization of this
specification.
(Call_C p n0)
(conform (Call_C p n0)
, we know
that 1 < n0 < p (this is done with an Elim
);
It is therefore legal to do an euclidean division of p by
n0, and check the rest r (with zerop
of
the module $COQTH/ARITH/Compare_dec).
Let q denote the quotient of p by n0.
MMC
of
the module develop
.
KMC
of
the module develop
.
We now have the piece of Real
term for
(Call_C p n0)
:
(* Call_C p n0 *) [p,n0:nat] <Code>let (q,r:nat)=(eucl_dev n0 p) in <Code>if (zerop r) then (MMC n0 q (REC (Call_M n0)) (REC (Call_M q))) else (KMC n0 q r (REC (Call_K n0 r))(REC (Call_M q)))
(Call_K p n0)
(* Call_K p n0 *) [p,n0:nat] <Code>let (q,r:nat)= (eucl_dev n0 p) in <Code>if (zerop r) then (MMK n0 q (REC (Call_M n0)) (REC (Call_M q))) else (KMK n0 q r (REC (Call_K n0 r)) (REC (Call_M q)))
chain_algo
is:
Lemma chain_algo: (c:Call)(conform c)->(gencode c). Proof. Realizer see above ... ...The informative goals are solved using
Program
or
Program_all
.
There remain a lot of logical goals, mainly for proving that each
recursive call is conform and less (w.r.t. Call_lt
)
than the current call.
One may think that the proof of these logical goals is long and tedious, but remind that there are 7 relationships between recursive calls to study.
Lemma eucl_dev : (b:nat)(gt b O)->(a:nat)(diveucl a b).where
diveucl
is defined as:
Inductive diveucl [a,b:nat] : Set := divex : (q,r:nat)(a=(plus (mult q b) r))-> (gt b r)-> (diveucl a b).(see the files $COQTH/PROGRAMS/Euclid*.v).
<Code>let(q,r)=(eucl_dev n0 p) in ...
Program
and
Do n Program
instead of Program_all
.
We did that to control the introduction of some logical properties.
For instance, consider the goal:
gamma : strategy log2_r : (n:nat)(lt O n)->(log2_spec n) a : Call c : Call REC : (y:Call)(Call_lt y c)->(conform y)->(gencode y) ============================ (n,n0:nat) ((y:Call)(Call_lt y (Call_C n n0))-> (conform y)-> (gencode y)) -> (conform (Call_C n n0))-> (gencode (Call_C n n0))A
Program_all
would give us 11 subgoals, 6 of them containing
the hypothesis:
e : p=(plus (mult q n0) r) e0 : r=OIf we want to use (by a
Cut
) an equality
e0:p=(mult q n0)
, we have to repeat a sequence
Cut ... Intro
in the 6 first subgoals; we prefered to control
by hand the Program
tactic, and introduce this equality before
the subgoals explosion.
Perhaps a good solution would be to put in the Real
term a comment
of the form:
(* Call_C p n0 *) [p,n0:nat] <Code>let (q,r:nat)=(eucl_dev n0 p) in <Code>if (zerop r) then (* p = (mult q n0) *) (MMC n0 q (REC (Call_M n0)) (REC (Call_M q))) else (KMC n0 q r (REC (Call_K n0 r)) (REC (Call_M q)))and let the Program tactic do a
Cut
on this comment.
(*)
Demo.ml
, and open
it.
The main theorems in this module have their functional counterparts:
power_algo (dicho log2_impl) log2_impl 16;; - : Code = seq (SQR, seq (SQR, seq (SQR, seq (SQR, End)))) #chain_algo (dicho log2_impl) log2_impl (Call_M 87);; - : Code = seq (PUSH, seq (PUSH, seq (SQR, seq (MUL, seq (PUSH, seq (SWAP, seq (SQR, seq (MUL, seq (PUSH, seq(SWAP, seq (MUL, seq (SQR, seq (SQR, seq (SQR, seq (MUL, End))))))))))))))) #chain_algo (dicho log2_impl) log2_impl (Call_K(7,3));; - : Code = seq (PUSH, seq (PUSH, seq (SQR, seq (MUL, seq (PUSH, seq (SWAP, seq (SQR, seq (MUL, End)))))))) (* generates a call for computing x^n, using the dichotomic strategy *) let demoCOND3 n = COND3 (dicho log2_impl) log2_impl n (Call_M 3) (function p -> (Call_M (two_power p))) (function p ->(Call_C(n,p)));; #demoCOND3 3;; - : Call = Call_M 3 #demoCOND3 4096;; - : Call = Call_M 4096 #demoCOND3 87;; - : Call = Call_C (87, 10) #demoCOND3 4095;; - : Call = Call_C (4095, 63)