author | paulson |
Fri, 11 Dec 1998 10:36:39 +0100 | |
changeset 6022 | 259e4f2114e1 |
parent 5690 | 4b056ee5435c |
child 6761 | aa71a04f4b93 |
permissions | -rw-r--r-- |
548 | 1 |
(* Title: Pure/Syntax/syn_trans.ML |
2 |
ID: $Id$ |
|
3 |
Author: Tobias Nipkow and Markus Wenzel, TU Muenchen |
|
4 |
||
5 |
Syntax translation functions. |
|
6 |
*) |
|
7 |
||
8 |
signature SYN_TRANS0 = |
|
2698 | 9 |
sig |
548 | 10 |
val eta_contract: bool ref |
11 |
val mk_binder_tr: string * string -> string * (term list -> term) |
|
12 |
val mk_binder_tr': string * string -> string * (term list -> term) |
|
13 |
val dependent_tr': string * string -> term list -> term |
|
5084 | 14 |
val quote_antiquote_tr: string -> string -> string -> string * (term list -> term) |
15 |
val quote_antiquote_tr': string -> string -> string -> string * (term list -> term) |
|
2698 | 16 |
val mark_bound: string -> term |
17 |
val mark_boundT: string * typ -> term |
|
18 |
val variant_abs': string * typ * term -> string * term |
|
19 |
end; |
|
548 | 20 |
|
21 |
signature SYN_TRANS1 = |
|
2698 | 22 |
sig |
548 | 23 |
include SYN_TRANS0 |
1511 | 24 |
val constrainAbsC: string |
25 |
val pure_trfuns: |
|
26 |
(string * (Ast.ast list -> Ast.ast)) list * |
|
548 | 27 |
(string * (term list -> term)) list * |
28 |
(string * (term list -> term)) list * |
|
1511 | 29 |
(string * (Ast.ast list -> Ast.ast)) list |
4148 | 30 |
val pure_trfunsT: (string * (bool -> typ -> term list -> term)) list |
2698 | 31 |
end; |
548 | 32 |
|
33 |
signature SYN_TRANS = |
|
2698 | 34 |
sig |
548 | 35 |
include SYN_TRANS1 |
1511 | 36 |
val abs_tr': term -> term |
4148 | 37 |
val prop_tr': term -> term |
1511 | 38 |
val appl_ast_tr': Ast.ast * Ast.ast list -> Ast.ast |
39 |
val applC_ast_tr': Ast.ast * Ast.ast list -> Ast.ast |
|
40 |
val pt_to_ast: (string -> (Ast.ast list -> Ast.ast) option) -> Parser.parsetree -> Ast.ast |
|
41 |
val ast_to_term: (string -> (term list -> term) option) -> Ast.ast -> term |
|
2698 | 42 |
end; |
548 | 43 |
|
2698 | 44 |
structure SynTrans: SYN_TRANS = |
548 | 45 |
struct |
2698 | 46 |
|
47 |
||
548 | 48 |
(** parse (ast) translations **) |
49 |
||
50 |
(* application *) |
|
51 |
||
5690 | 52 |
fun appl_ast_tr [f, args] = Ast.Appl (f :: Ast.unfold_ast "_args" args) |
53 |
| appl_ast_tr asts = raise Ast.AST ("appl_ast_tr", asts); |
|
922
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
639
diff
changeset
|
54 |
|
5690 | 55 |
fun applC_ast_tr [f, args] = Ast.Appl (f :: Ast.unfold_ast "_cargs" args) |
56 |
| applC_ast_tr asts = raise Ast.AST ("applC_ast_tr", asts); |
|
548 | 57 |
|
58 |
||
59 |
(* abstraction *) |
|
60 |
||
5690 | 61 |
fun idtyp_ast_tr (*"_idtyp"*) [x, ty] = Ast.Appl [Ast.Constant SynExt.constrainC, x, ty] |
62 |
| idtyp_ast_tr (*"_idtyp"*) asts = raise Ast.AST ("idtyp_ast_tr", asts); |
|
548 | 63 |
|
3691
f0396ac63e12
tuned lambda_ast_tr, idtyp_ast_tr' to accomodate fix of idt/idts
wenzelm
parents:
2698
diff
changeset
|
64 |
fun lambda_ast_tr (*"_lambda"*) [pats, body] = |
5690 | 65 |
Ast.fold_ast_p "_abs" (Ast.unfold_ast "_pttrns" pats, body) |
66 |
| lambda_ast_tr (*"_lambda"*) asts = raise Ast.AST ("lambda_ast_tr", asts); |
|
548 | 67 |
|
68 |
val constrainAbsC = "_constrainAbs"; |
|
69 |
||
5690 | 70 |
fun abs_tr (*"_abs"*) [Free (x, T), body] = Term.absfree (x, T, body) |
548 | 71 |
| abs_tr (*"_abs"*) (ts as [Const (c, _) $ Free (x, T) $ tT, body]) = |
5690 | 72 |
if c = SynExt.constrainC |
73 |
then Lexicon.const constrainAbsC $ Term.absfree (x, T, body) $ tT |
|
3777 | 74 |
else raise TERM ("abs_tr", ts) |
75 |
| abs_tr (*"_abs"*) ts = raise TERM ("abs_tr", ts); |
|
548 | 76 |
|
77 |
||
78 |
(* nondependent abstraction *) |
|
79 |
||
5690 | 80 |
fun k_tr (*"_K"*) [t] = Abs ("uu", dummyT, Term.incr_boundvars 1 t) |
3777 | 81 |
| k_tr (*"_K"*) ts = raise TERM ("k_tr", ts); |
548 | 82 |
|
83 |
||
84 |
(* binder *) |
|
85 |
||
86 |
fun mk_binder_tr (sy, name) = |
|
87 |
let |
|
5690 | 88 |
fun tr (Free (x, T), t) = Lexicon.const name $ Term.absfree (x, T, t) |
548 | 89 |
| tr (Const ("_idts", _) $ idt $ idts, t) = tr (idt, tr (idts, t)) |
90 |
| tr (t1 as Const (c, _) $ Free (x, T) $ tT, t) = |
|
5690 | 91 |
if c = SynExt.constrainC then |
92 |
Lexicon.const name $ (Lexicon.const constrainAbsC $ Term.absfree (x, T, t) $ tT) |
|
3777 | 93 |
else raise TERM ("binder_tr", [t1, t]) |
94 |
| tr (t1, t2) = raise TERM ("binder_tr", [t1, t2]); |
|
548 | 95 |
|
96 |
fun binder_tr (*sy*) [idts, body] = tr (idts, body) |
|
3777 | 97 |
| binder_tr (*sy*) ts = raise TERM ("binder_tr", ts); |
548 | 98 |
in |
99 |
(sy, binder_tr) |
|
100 |
end; |
|
101 |
||
102 |
||
103 |
(* meta propositions *) |
|
104 |
||
5690 | 105 |
fun aprop_tr (*"_aprop"*) [t] = Lexicon.const SynExt.constrainC $ t $ Lexicon.const "prop" |
3777 | 106 |
| aprop_tr (*"_aprop"*) ts = raise TERM ("aprop_tr", ts); |
548 | 107 |
|
108 |
fun ofclass_tr (*"_ofclass"*) [ty, cls] = |
|
5690 | 109 |
cls $ (Lexicon.const SynExt.constrainC $ Lexicon.const "TYPE" $ |
110 |
(Lexicon.const "itself" $ ty)) |
|
3777 | 111 |
| ofclass_tr (*"_ofclass"*) ts = raise TERM ("ofclass_tr", ts); |
548 | 112 |
|
113 |
||
114 |
(* meta implication *) |
|
115 |
||
116 |
fun bigimpl_ast_tr (*"_bigimpl"*) [asms, concl] = |
|
5690 | 117 |
Ast.fold_ast_p "==>" (Ast.unfold_ast "_asms" asms, concl) |
118 |
| bigimpl_ast_tr (*"_bigimpl"*) asts = raise Ast.AST ("bigimpl_ast_tr", asts); |
|
548 | 119 |
|
120 |
||
4148 | 121 |
(* type reflection *) |
122 |
||
123 |
fun type_tr (*"_TYPE"*) [ty] = |
|
5690 | 124 |
Lexicon.const SynExt.constrainC $ Lexicon.const "TYPE" $ (Lexicon.const "itself" $ ty) |
4148 | 125 |
| type_tr (*"_TYPE"*) ts = raise TERM ("type_tr", ts); |
126 |
||
548 | 127 |
|
5288 | 128 |
(* binds *) |
129 |
||
5690 | 130 |
fun bind_ast_tr (*"_BIND"*) [Ast.Variable x] = |
131 |
Ast.Variable (Lexicon.string_of_vname (Lexicon.binding x, 0)) |
|
132 |
| bind_ast_tr (*"_BIND"*) asts = raise Ast.AST ("bind_ast_tr", asts); |
|
5288 | 133 |
|
134 |
||
5084 | 135 |
(* quote / antiquote *) |
136 |
||
137 |
fun quote_antiquote_tr quoteN antiquoteN name = |
|
138 |
let |
|
139 |
fun antiquote_tr i ((a as Const (c, _)) $ t) = |
|
140 |
if c = antiquoteN then t $ Bound i |
|
141 |
else a $ antiquote_tr i t |
|
142 |
| antiquote_tr i (t $ u) = antiquote_tr i t $ antiquote_tr i u |
|
143 |
| antiquote_tr i (Abs (x, T, t)) = Abs (x, T, antiquote_tr (i + 1) t) |
|
144 |
| antiquote_tr _ a = a; |
|
145 |
||
5690 | 146 |
fun quote_tr [t] = Lexicon.const name $ |
147 |
Abs ("uu", dummyT, antiquote_tr 0 (Term.incr_boundvars 1 t)) |
|
5084 | 148 |
| quote_tr ts = raise TERM ("quote_tr", ts); |
149 |
in (quoteN, quote_tr) end; |
|
150 |
||
151 |
||
152 |
||
548 | 153 |
(** print (ast) translations **) |
154 |
||
155 |
(* application *) |
|
156 |
||
5690 | 157 |
fun appl_ast_tr' (f, []) = raise Ast.AST ("appl_ast_tr'", [f]) |
158 |
| appl_ast_tr' (f, args) = Ast.Appl [Ast.Constant "_appl", f, Ast.fold_ast "_args" args]; |
|
548 | 159 |
|
5690 | 160 |
fun applC_ast_tr' (f, []) = raise Ast.AST ("applC_ast_tr'", [f]) |
161 |
| applC_ast_tr' (f, args) = Ast.Appl [Ast.Constant "_applC", f, Ast.fold_ast "_cargs" args]; |
|
922
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
639
diff
changeset
|
162 |
|
548 | 163 |
|
164 |
(* abstraction *) |
|
165 |
||
5690 | 166 |
fun mark_boundT x_T = Lexicon.const "_bound" $ Free x_T; |
2698 | 167 |
fun mark_bound x = mark_boundT (x, dummyT); |
168 |
||
548 | 169 |
fun strip_abss vars_of body_of tm = |
170 |
let |
|
171 |
val vars = vars_of tm; |
|
172 |
val body = body_of tm; |
|
173 |
val rev_new_vars = rename_wrt_term body vars; |
|
174 |
in |
|
2698 | 175 |
(map mark_boundT (rev rev_new_vars), |
176 |
subst_bounds (map (mark_bound o #1) rev_new_vars, body)) |
|
548 | 177 |
end; |
178 |
||
3928
787d2659ce4a
no longer tries bogus eta-contract involving aprops;
wenzelm
parents:
3777
diff
changeset
|
179 |
|
548 | 180 |
(*do (partial) eta-contraction before printing*) |
181 |
||
1326 | 182 |
val eta_contract = ref true; |
548 | 183 |
|
184 |
fun eta_contr tm = |
|
185 |
let |
|
3928
787d2659ce4a
no longer tries bogus eta-contract involving aprops;
wenzelm
parents:
3777
diff
changeset
|
186 |
fun is_aprop (Const ("_aprop", _)) = true |
787d2659ce4a
no longer tries bogus eta-contract involving aprops;
wenzelm
parents:
3777
diff
changeset
|
187 |
| is_aprop _ = false; |
787d2659ce4a
no longer tries bogus eta-contract involving aprops;
wenzelm
parents:
3777
diff
changeset
|
188 |
|
548 | 189 |
fun eta_abs (Abs (a, T, t)) = |
190 |
(case eta_abs t of |
|
191 |
t' as f $ u => |
|
192 |
(case eta_abs u of |
|
193 |
Bound 0 => |
|
5084 | 194 |
if Term.loose_bvar1 (f, 0) orelse is_aprop f then Abs (a, T, t') |
3928
787d2659ce4a
no longer tries bogus eta-contract involving aprops;
wenzelm
parents:
3777
diff
changeset
|
195 |
else incr_boundvars ~1 f |
548 | 196 |
| _ => Abs (a, T, t')) |
197 |
| t' => Abs (a, T, t')) |
|
198 |
| eta_abs t = t; |
|
199 |
in |
|
200 |
if ! eta_contract then eta_abs tm else tm |
|
201 |
end; |
|
202 |
||
203 |
||
204 |
fun abs_tr' tm = |
|
5690 | 205 |
foldr (fn (x, t) => Lexicon.const "_abs" $ x $ t) |
548 | 206 |
(strip_abss strip_abs_vars strip_abs_body (eta_contr tm)); |
207 |
||
208 |
||
209 |
fun abs_ast_tr' (*"_abs"*) asts = |
|
5690 | 210 |
(case Ast.unfold_ast_p "_abs" (Ast.Appl (Ast.Constant "_abs" :: asts)) of |
211 |
([], _) => raise Ast.AST ("abs_ast_tr'", asts) |
|
212 |
| (xs, body) => Ast.Appl [Ast.Constant "_lambda", Ast.fold_ast "_pttrns" xs, body]); |
|
548 | 213 |
|
214 |
||
215 |
(* binder *) |
|
216 |
||
217 |
fun mk_binder_tr' (name, sy) = |
|
218 |
let |
|
219 |
fun mk_idts [] = raise Match (*abort translation*) |
|
220 |
| mk_idts [idt] = idt |
|
5690 | 221 |
| mk_idts (idt :: idts) = Lexicon.const "_idts" $ idt $ mk_idts idts; |
548 | 222 |
|
223 |
fun tr' t = |
|
224 |
let |
|
225 |
val (xs, bd) = strip_abss (strip_qnt_vars name) (strip_qnt_body name) t; |
|
5690 | 226 |
in Lexicon.const sy $ mk_idts xs $ bd end; |
548 | 227 |
|
5690 | 228 |
fun binder_tr' (*name*) (t :: ts) = Term.list_comb (tr' (Lexicon.const name $ t), ts) |
548 | 229 |
| binder_tr' (*name*) [] = raise Match; |
230 |
in |
|
231 |
(name, binder_tr') |
|
232 |
end; |
|
233 |
||
234 |
||
3691
f0396ac63e12
tuned lambda_ast_tr, idtyp_ast_tr' to accomodate fix of idt/idts
wenzelm
parents:
2698
diff
changeset
|
235 |
(* idtyp constraints *) |
548 | 236 |
|
5690 | 237 |
fun idtyp_ast_tr' a [Ast.Appl [Ast.Constant c, x, ty], xs] = |
238 |
if c = SynExt.constrainC then |
|
239 |
Ast.Appl [ Ast.Constant a, Ast.Appl [Ast.Constant "_idtyp", x, ty], xs] |
|
548 | 240 |
else raise Match |
3691
f0396ac63e12
tuned lambda_ast_tr, idtyp_ast_tr' to accomodate fix of idt/idts
wenzelm
parents:
2698
diff
changeset
|
241 |
| idtyp_ast_tr' _ _ = raise Match; |
548 | 242 |
|
243 |
||
244 |
(* meta propositions *) |
|
245 |
||
4148 | 246 |
fun prop_tr' tm = |
548 | 247 |
let |
5690 | 248 |
fun aprop t = Lexicon.const "_aprop" $ t; |
548 | 249 |
|
2698 | 250 |
fun is_prop Ts t = |
251 |
fastype_of1 (Ts, t) = propT handle TERM _ => false; |
|
548 | 252 |
|
253 |
fun tr' _ (t as Const _) = t |
|
2698 | 254 |
| tr' _ (t as Free (x, T)) = |
5690 | 255 |
if T = propT then aprop (Lexicon.free x) else t |
2698 | 256 |
| tr' _ (t as Var (xi, T)) = |
5690 | 257 |
if T = propT then aprop (Lexicon.var xi) else t |
2698 | 258 |
| tr' Ts (t as Bound _) = |
259 |
if is_prop Ts t then aprop t else t |
|
260 |
| tr' Ts (Abs (x, T, t)) = Abs (x, T, tr' (T :: Ts) t) |
|
261 |
| tr' Ts (t as t1 $ (t2 as Const ("TYPE", Type ("itself", [T])))) = |
|
4148 | 262 |
if is_prop Ts t then Const ("_mk_ofclass", T) $ tr' Ts t1 |
2698 | 263 |
else tr' Ts t1 $ tr' Ts t2 |
264 |
| tr' Ts (t as t1 $ t2) = |
|
5690 | 265 |
(if is_Const (Term.head_of t) orelse not (is_prop Ts t) |
2698 | 266 |
then I else aprop) (tr' Ts t1 $ tr' Ts t2); |
548 | 267 |
in |
268 |
tr' [] tm |
|
269 |
end; |
|
270 |
||
4148 | 271 |
fun mk_ofclass_tr' show_sorts (*"_mk_ofclass"*) T [t] = |
5690 | 272 |
Lexicon.const "_ofclass" $ TypeExt.term_of_typ show_sorts T $ t |
4148 | 273 |
| mk_ofclass_tr' _ (*"_mk_ofclass"*) T ts = raise TYPE ("mk_ofclass_tr'", [T], ts); |
2698 | 274 |
|
275 |
||
548 | 276 |
(* meta implication *) |
277 |
||
278 |
fun impl_ast_tr' (*"==>"*) asts = |
|
5690 | 279 |
(case Ast.unfold_ast_p "==>" (Ast.Appl (Ast.Constant "==>" :: asts)) of |
548 | 280 |
(asms as _ :: _ :: _, concl) |
5690 | 281 |
=> Ast.Appl [Ast.Constant "_bigimpl", Ast.fold_ast "_asms" asms, concl] |
548 | 282 |
| _ => raise Match); |
283 |
||
284 |
||
4148 | 285 |
(* type reflection *) |
286 |
||
287 |
fun type_tr' show_sorts (*"TYPE"*) (Type ("itself", [T])) ts = |
|
5690 | 288 |
Term.list_comb (Lexicon.const "_TYPE" $ TypeExt.term_of_typ show_sorts T, ts) |
4148 | 289 |
| type_tr' _ _ _ = raise Match; |
290 |
||
291 |
||
548 | 292 |
(* dependent / nondependent quantifiers *) |
293 |
||
2698 | 294 |
fun variant_abs' (x, T, B) = |
295 |
let val x' = variant (add_term_names (B, [])) x in |
|
296 |
(x', subst_bound (mark_boundT (x', T), B)) |
|
297 |
end; |
|
298 |
||
548 | 299 |
fun dependent_tr' (q, r) (A :: Abs (x, T, B) :: ts) = |
5084 | 300 |
if Term.loose_bvar1 (B, 0) then |
2698 | 301 |
let val (x', B') = variant_abs' (x, dummyT, B); |
5690 | 302 |
in Term.list_comb (Lexicon.const q $ mark_boundT (x', T) $ A $ B', ts) end |
303 |
else Term.list_comb (Lexicon.const r $ A $ B, ts) |
|
548 | 304 |
| dependent_tr' _ _ = raise Match; |
305 |
||
306 |
||
5084 | 307 |
(* quote / antiquote *) |
308 |
||
309 |
fun no_loose_bvar i t = |
|
310 |
if Term.loose_bvar1 (t, i) then raise Match else t; |
|
311 |
||
312 |
fun quote_antiquote_tr' quoteN antiquoteN name = |
|
313 |
let |
|
314 |
fun antiquote_tr' i (t $ u) = |
|
5690 | 315 |
if u = Bound i then Lexicon.const antiquoteN $ no_loose_bvar i t |
5084 | 316 |
else antiquote_tr' i t $ antiquote_tr' i u |
317 |
| antiquote_tr' i (Abs (x, T, t)) = Abs (x, T, antiquote_tr' (i + 1) t) |
|
318 |
| antiquote_tr' i a = no_loose_bvar i a; |
|
319 |
||
320 |
fun quote_tr' (Abs (x, T, t) :: ts) = |
|
5690 | 321 |
Term.list_comb ( Lexicon.const quoteN $ Term.incr_boundvars ~1 (antiquote_tr' 0 t), ts) |
5084 | 322 |
| quote_tr' _ = raise Match; |
323 |
in (name, quote_tr') end; |
|
324 |
||
325 |
||
548 | 326 |
|
327 |
(** pure_trfuns **) |
|
328 |
||
329 |
val pure_trfuns = |
|
922
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
639
diff
changeset
|
330 |
([("_appl", appl_ast_tr), ("_applC", applC_ast_tr), |
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
639
diff
changeset
|
331 |
("_lambda", lambda_ast_tr), ("_idtyp", idtyp_ast_tr), |
5288 | 332 |
("_bigimpl", bigimpl_ast_tr), ("_BIND", bind_ast_tr)], |
922
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents:
639
diff
changeset
|
333 |
[("_abs", abs_tr), ("_aprop", aprop_tr), ("_ofclass", ofclass_tr), |
4148 | 334 |
("_TYPE", type_tr), ("_K", k_tr)], |
3700 | 335 |
[]: (string * (term list -> term)) list, |
3691
f0396ac63e12
tuned lambda_ast_tr, idtyp_ast_tr' to accomodate fix of idt/idts
wenzelm
parents:
2698
diff
changeset
|
336 |
[("_abs", abs_ast_tr'), ("_idts", idtyp_ast_tr' "_idts"), |
f0396ac63e12
tuned lambda_ast_tr, idtyp_ast_tr' to accomodate fix of idt/idts
wenzelm
parents:
2698
diff
changeset
|
337 |
("_pttrns", idtyp_ast_tr' "_pttrns"), ("==>", impl_ast_tr')]); |
548 | 338 |
|
2698 | 339 |
val pure_trfunsT = |
4148 | 340 |
[("_mk_ofclass", mk_ofclass_tr'), ("TYPE", type_tr')]; |
2698 | 341 |
|
548 | 342 |
|
343 |
||
344 |
(** pt_to_ast **) |
|
345 |
||
346 |
fun pt_to_ast trf pt = |
|
347 |
let |
|
348 |
fun trans a args = |
|
349 |
(case trf a of |
|
5690 | 350 |
None => Ast.mk_appl (Ast.Constant a) args |
548 | 351 |
| Some f => f args handle exn |
987 | 352 |
=> (writeln ("Error in parse ast translation for " ^ quote a); |
353 |
raise exn)); |
|
548 | 354 |
|
987 | 355 |
(*translate pt bottom-up*) |
5690 | 356 |
fun ast_of (Parser.Node (a, pts)) = trans a (map ast_of pts) |
357 |
| ast_of (Parser.Tip tok) = Ast.Variable (Lexicon.str_of_token tok); |
|
548 | 358 |
in |
1178 | 359 |
ast_of pt |
548 | 360 |
end; |
361 |
||
362 |
||
363 |
||
364 |
(** ast_to_term **) |
|
365 |
||
366 |
fun ast_to_term trf ast = |
|
367 |
let |
|
368 |
fun trans a args = |
|
369 |
(case trf a of |
|
5690 | 370 |
None => Term.list_comb (Lexicon.const a, args) |
548 | 371 |
| Some f => f args handle exn |
987 | 372 |
=> (writeln ("Error in parse translation for " ^ quote a); |
373 |
raise exn)); |
|
548 | 374 |
|
5690 | 375 |
fun term_of (Ast.Constant a) = trans a [] |
376 |
| term_of (Ast.Variable x) = Lexicon.read_var x |
|
377 |
| term_of (Ast.Appl (Ast.Constant a :: (asts as _ :: _))) = |
|
548 | 378 |
trans a (map term_of asts) |
5690 | 379 |
| term_of (Ast.Appl (ast :: (asts as _ :: _))) = |
380 |
Term.list_comb (term_of ast, map term_of asts) |
|
381 |
| term_of (ast as Ast.Appl _) = raise Ast.AST ("ast_to_term: malformed ast", [ast]); |
|
548 | 382 |
in |
383 |
term_of ast |
|
384 |
end; |
|
385 |
||
386 |
end; |