author | wenzelm |
Tue, 14 Mar 2006 22:06:42 +0100 | |
changeset 19271 | 967e6c2578f2 |
parent 17456 | bcf7544875b2 |
permissions | -rw-r--r-- |
1459 | 1 |
(* Title: 92/CCL/eval |
0 | 2 |
ID: $Id$ |
1459 | 3 |
Author: Martin Coen, Cambridge University Computer Laboratory |
0 | 4 |
Copyright 1992 University of Cambridge |
5 |
*) |
|
6 |
||
7 |
(*** Evaluation ***) |
|
8 |
||
9 |
val EVal_rls = ref [trueV,falseV,pairV,lamV,caseVtrue,caseVfalse,caseVpair,caseVlam]; |
|
10 |
val eval_tac = DEPTH_SOLVE_1 (resolve_tac (!EVal_rls) 1); |
|
11 |
fun ceval_tac rls = DEPTH_SOLVE_1 (resolve_tac (!EVal_rls@rls) 1); |
|
12 |
||
17456 | 13 |
val prems = goalw (the_context ()) [apply_def] |
3837 | 14 |
"[| f ---> lam x. b(x); b(a) ---> c |] ==> f ` a ---> c"; |
0 | 15 |
by (ceval_tac prems); |
757 | 16 |
qed "applyV"; |
0 | 17 |
|
18 |
EVal_rls := !EVal_rls @ [applyV]; |
|
19 |
||
17456 | 20 |
val major::prems = goalw (the_context ()) [let_def] |
0 | 21 |
"[| t ---> a; f(a) ---> c |] ==> let x be t in f(x) ---> c"; |
642
0db578095e6a
CCL/Gfp/coinduct2, coinduct3: modified proofs to suppress deep unification
lcp
parents:
289
diff
changeset
|
22 |
by (rtac (major RS canonical) 1); |
0 | 23 |
by (REPEAT (DEPTH_SOLVE_1 (resolve_tac ([major]@prems@(!EVal_rls)) 1 ORELSE |
642
0db578095e6a
CCL/Gfp/coinduct2, coinduct3: modified proofs to suppress deep unification
lcp
parents:
289
diff
changeset
|
24 |
etac substitute 1))); |
757 | 25 |
qed "letV"; |
0 | 26 |
|
17456 | 27 |
val prems = goalw (the_context ()) [fix_def] |
0 | 28 |
"f(fix(f)) ---> c ==> fix(f) ---> c"; |
642
0db578095e6a
CCL/Gfp/coinduct2, coinduct3: modified proofs to suppress deep unification
lcp
parents:
289
diff
changeset
|
29 |
by (rtac applyV 1); |
0db578095e6a
CCL/Gfp/coinduct2, coinduct3: modified proofs to suppress deep unification
lcp
parents:
289
diff
changeset
|
30 |
by (rtac lamV 1); |
0db578095e6a
CCL/Gfp/coinduct2, coinduct3: modified proofs to suppress deep unification
lcp
parents:
289
diff
changeset
|
31 |
by (resolve_tac prems 1); |
757 | 32 |
qed "fixV"; |
0 | 33 |
|
17456 | 34 |
val prems = goalw (the_context ()) [letrec_def] |
3837 | 35 |
"h(t,%y. letrec g x be h(x,g) in g(y)) ---> c ==> \ |
0 | 36 |
\ letrec g x be h(x,g) in g(t) ---> c"; |
37 |
by (REPEAT (resolve_tac (prems @ [fixV,applyV,lamV]) 1)); |
|
757 | 38 |
qed "letrecV"; |
0 | 39 |
|
40 |
EVal_rls := !EVal_rls @ [letV,letrecV,fixV]; |
|
41 |
||
17456 | 42 |
fun mk_V_rl s = prove_goalw (the_context ()) data_defs s (fn prems => [ceval_tac prems]); |
0 | 43 |
|
17456 | 44 |
val V_rls = map mk_V_rl |
0 | 45 |
["true ---> true", |
46 |
"false ---> false", |
|
47 |
"[| b--->true; t--->c |] ==> if b then t else u ---> c", |
|
48 |
"[| b--->false; u--->c |] ==> if b then t else u ---> c", |
|
49 |
"<a,b> ---> <a,b>", |
|
50 |
"[| t ---> <a,b>; h(a,b) ---> c |] ==> split(t,h) ---> c", |
|
51 |
"zero ---> zero", |
|
52 |
"succ(n) ---> succ(n)", |
|
53 |
"[| n ---> zero; t ---> c |] ==> ncase(n,t,u) ---> c", |
|
54 |
"[| n ---> succ(x); u(x) ---> c |] ==> ncase(n,t,u) ---> c", |
|
55 |
"[| n ---> zero; t ---> c |] ==> nrec(n,t,u) ---> c", |
|
56 |
"[| n--->succ(x); u(x,nrec(x,t,u))--->c |] ==> nrec(n,t,u)--->c", |
|
57 |
"[] ---> []", |
|
289 | 58 |
"h$t ---> h$t", |
0 | 59 |
"[| l ---> []; t ---> c |] ==> lcase(l,t,u) ---> c", |
289 | 60 |
"[| l ---> x$xs; u(x,xs) ---> c |] ==> lcase(l,t,u) ---> c", |
0 | 61 |
"[| l ---> []; t ---> c |] ==> lrec(l,t,u) ---> c", |
289 | 62 |
"[| l--->x$xs; u(x,xs,lrec(xs,t,u))--->c |] ==> lrec(l,t,u)--->c"]; |
0 | 63 |
|
64 |
EVal_rls := !EVal_rls @ V_rls; |
|
65 |
||
66 |
(* Factorial *) |
|
67 |
||
17456 | 68 |
val prems = goal (the_context ()) |
3837 | 69 |
"letrec f n be ncase(n,succ(zero),%x. nrec(n,zero,%y g. nrec(f(x),g,%z h. succ(h)))) \ |
0 | 70 |
\ in f(succ(succ(zero))) ---> ?a"; |
71 |
by (ceval_tac []); |
|
72 |
||
17456 | 73 |
val prems = goal (the_context ()) |
3837 | 74 |
"letrec f n be ncase(n,succ(zero),%x. nrec(n,zero,%y g. nrec(f(x),g,%z h. succ(h)))) \ |
0 | 75 |
\ in f(succ(succ(succ(zero)))) ---> ?a"; |
76 |
by (ceval_tac []); |
|
77 |
||
78 |
(* Less Than Or Equal *) |
|
79 |
||
17456 | 80 |
fun isle x y = prove_goal (the_context ()) |
3837 | 81 |
("letrec f p be split(p,%m n. ncase(m,true,%x. ncase(n,false,%y. f(<x,y>)))) \ |
0 | 82 |
\ in f(<"^x^","^y^">) ---> ?a") |
83 |
(fn prems => [ceval_tac []]); |
|
84 |
||
85 |
isle "succ(zero)" "succ(zero)"; |
|
86 |
isle "succ(zero)" "succ(succ(succ(succ(zero))))"; |
|
87 |
isle "succ(succ(succ(succ(succ(zero)))))" "succ(succ(succ(succ(zero))))"; |
|
88 |
||
89 |
||
90 |
(* Reverse *) |
|
91 |
||
17456 | 92 |
val prems = goal (the_context ()) |
3837 | 93 |
"letrec id l be lcase(l,[],%x xs. x$id(xs)) \ |
289 | 94 |
\ in id(zero$succ(zero)$[]) ---> ?a"; |
0 | 95 |
by (ceval_tac []); |
96 |
||
17456 | 97 |
val prems = goal (the_context ()) |
3837 | 98 |
"letrec rev l be lcase(l,[],%x xs. lrec(rev(xs),x$[],%y ys g. y$g)) \ |
99 |
\ in rev(zero$succ(zero)$(succ((lam x. x)`succ(zero)))$([])) ---> ?a"; |
|
0 | 100 |
by (ceval_tac []); |