20856
|
1 |
(* Title: Pure/nbe_eval.ML
|
|
2 |
ID: $Id$
|
19177
|
3 |
Author: Klaus Aehlig, LMU Muenchen; Tobias Nipkow, TU Muenchen
|
|
4 |
|
|
5 |
Evaluator infrastructure for "normalization by evaluation".
|
19116
|
6 |
*)
|
|
7 |
|
|
8 |
(* Optimizations:
|
|
9 |
int instead of string in Constr
|
|
10 |
treat int via ML ints
|
|
11 |
*)
|
|
12 |
|
|
13 |
signature NBE_EVAL =
|
|
14 |
sig
|
19177
|
15 |
(*named terms used for output only*)
|
|
16 |
datatype nterm =
|
|
17 |
C of string (*named constants*)
|
|
18 |
| V of string (*free variable*)
|
|
19 |
| B of int (*named(!!) bound variables*)
|
|
20 |
| A of nterm * nterm (*application*)
|
|
21 |
| AbsN of int * nterm (*binding of named variables*);
|
|
22 |
datatype Univ =
|
|
23 |
Constr of string*(Univ list) (*named constructors*)
|
|
24 |
| Var of string*(Univ list) (*free variables*)
|
|
25 |
| BVar of int*(Univ list) (*bound named variables*)
|
|
26 |
| Fun of (Univ list -> Univ) * (Univ list) * int * (unit -> nterm)
|
|
27 |
(*functions*);
|
19116
|
28 |
|
20856
|
29 |
val eval: theory -> Univ Symtab.table -> term -> nterm
|
19795
|
30 |
val apply: Univ -> Univ -> Univ
|
20846
|
31 |
val prep_term: theory -> term -> term
|
19177
|
32 |
|
19795
|
33 |
val to_term: Univ -> nterm
|
19147
|
34 |
|
19795
|
35 |
val mk_Fun: string * (Univ list -> Univ) * int -> string * Univ
|
|
36 |
val new_name: unit -> int
|
19116
|
37 |
|
19795
|
38 |
val string_of_nterm: nterm -> string
|
19116
|
39 |
end;
|
|
40 |
|
19177
|
41 |
structure NBE_Eval: NBE_EVAL =
|
19116
|
42 |
struct
|
|
43 |
|
|
44 |
datatype nterm =
|
19177
|
45 |
C of string
|
|
46 |
| V of string
|
|
47 |
| B of int
|
|
48 |
| A of nterm * nterm
|
|
49 |
| AbsN of int * nterm;
|
19116
|
50 |
|
19795
|
51 |
fun string_of_nterm(C s) = "(C \"" ^ s ^ "\")"
|
|
52 |
| string_of_nterm(V s) = "(V \"" ^ s ^ "\")"
|
|
53 |
| string_of_nterm(B n) = "(B " ^ string_of_int n ^ ")"
|
|
54 |
| string_of_nterm(A(s,t)) =
|
|
55 |
"(A " ^ string_of_nterm s ^ string_of_nterm t ^ ")"
|
|
56 |
| string_of_nterm(AbsN(n,t)) =
|
|
57 |
"(Abs " ^ string_of_int n ^ " " ^ string_of_nterm t ^ ")";
|
|
58 |
|
19116
|
59 |
fun apps t args = foldr (fn (y,x) => A(x,y)) t args;
|
|
60 |
|
|
61 |
(* ------------------------------ The semantical universe --------------------- *)
|
|
62 |
|
|
63 |
(*
|
|
64 |
Functions are given by their semantical function value. To avoid
|
|
65 |
trouble with the ML-type system, these functions have the most
|
|
66 |
generic type, that is "Univ list -> Univ". The calling convention is
|
|
67 |
that the arguments come as a list, the last argument first. In
|
|
68 |
other words, a function call that usually would look like
|
|
69 |
|
|
70 |
f x_1 x_2 ... x_n or f(x_1,x_2, ..., x_n)
|
|
71 |
|
|
72 |
would be in our convention called as
|
|
73 |
|
|
74 |
f [x_n,..,x_2,x_1]
|
|
75 |
|
|
76 |
Moreover, to handle functions that are still waiting for some
|
|
77 |
arguments we have additionally a list of arguments collected to far
|
|
78 |
and the number of arguments we're still waiting for.
|
|
79 |
|
|
80 |
Finally, it might happen, that a function does not get all the
|
|
81 |
arguments it needs. In this case the function must provide means to
|
|
82 |
present itself as a string. As this might be a heavy-wight
|
|
83 |
operation, we delay it.
|
|
84 |
*)
|
|
85 |
|
|
86 |
datatype Univ =
|
|
87 |
Constr of string*(Univ list) (* Named Constructors *)
|
|
88 |
| Var of string*(Univ list) (* Free variables *)
|
|
89 |
| BVar of int*(Univ list) (* Bound named variables *)
|
|
90 |
| Fun of (Univ list -> Univ) * (Univ list) * int * (unit -> nterm);
|
|
91 |
|
|
92 |
fun to_term (Constr(name, args)) = apps (C name) (map to_term args)
|
20596
|
93 |
| to_term (Var (name, args)) = apps (V name) (map to_term args)
|
|
94 |
| to_term (BVar (name, args)) = apps (B name) (map to_term args)
|
|
95 |
| to_term (Fun(_,args,_,name_thunk)) = apps (name_thunk ()) (map to_term args);
|
19116
|
96 |
|
|
97 |
(*
|
|
98 |
A typical operation, why functions might be good for, is
|
|
99 |
application. Moreover, functions are the only values that can
|
|
100 |
reasonably we applied in a semantical way.
|
|
101 |
*)
|
|
102 |
|
19177
|
103 |
fun apply (Fun(f,xs,1,name)) x = f (x::xs)
|
|
104 |
| apply (Fun(f,xs,n,name)) x = Fun(f,x::xs,n-1,name)
|
|
105 |
| apply (Constr(name,args)) x = Constr(name,x::args)
|
|
106 |
| apply (Var(name,args)) x = Var(name,x::args)
|
|
107 |
| apply (BVar(name,args)) x = BVar(name,x::args);
|
19116
|
108 |
|
19147
|
109 |
fun mk_Fun(name,v,0) = (name,v [])
|
|
110 |
| mk_Fun(name,v,n) = (name,Fun(v,[],n, fn () => C name));
|
|
111 |
|
19116
|
112 |
|
|
113 |
(* ------------------ evaluation with greetings to Tarski ------------------ *)
|
|
114 |
|
20846
|
115 |
fun prep_term thy (Const c) = Const (CodegenNames.const thy (CodegenConsts.norm_of_typ thy c), dummyT)
|
|
116 |
| prep_term thy (Free v_ty) = Free v_ty
|
|
117 |
| prep_term thy (s $ t) = prep_term thy s $ prep_term thy t
|
|
118 |
| prep_term thy (Abs (raw_v, ty, raw_t)) =
|
|
119 |
let
|
22038
|
120 |
val (v, t) = Syntax.variant_abs (raw_v, ty, raw_t);
|
20846
|
121 |
in Abs (v, ty, prep_term thy t) end;
|
|
122 |
|
|
123 |
|
19116
|
124 |
(* generation of fresh names *)
|
|
125 |
|
|
126 |
val counter = ref 0;
|
20856
|
127 |
fun new_name () = (counter := !counter +1; !counter);
|
19116
|
128 |
|
|
129 |
|
|
130 |
(* greetings to Tarski *)
|
|
131 |
|
20856
|
132 |
fun eval thy tab t =
|
19202
|
133 |
let
|
20856
|
134 |
fun evl vars (Const (s, _)) =
|
|
135 |
the_default (Constr (s, [])) (Symtab.lookup tab s)
|
|
136 |
| evl vars (Free (v, _)) =
|
|
137 |
the_default (Var (v, [])) (AList.lookup (op =) vars v)
|
|
138 |
| evl vars (s $ t) =
|
|
139 |
apply (evl vars s) (evl vars t)
|
20846
|
140 |
| evl vars (Abs (v, _, t)) =
|
20856
|
141 |
Fun (fn [x] => evl (AList.update (op =) (v, x) vars) t, [], 1,
|
20846
|
142 |
fn () => let val var = new_name() in
|
|
143 |
AbsN (var, to_term (evl (AList.update (op =) (v, BVar (var, [])) vars) t)) end)
|
20856
|
144 |
in (counter := 0; to_term (evl [] (prep_term thy t))) end;
|
19116
|
145 |
|
|
146 |
end;
|