18
|
1 |
(* Title: Pure/Syntax/parse_tree.ML
|
0
|
2 |
ID: $Id$
|
|
3 |
Author: Tobias Nipkow and Markus Wenzel, TU Muenchen
|
18
|
4 |
|
|
5 |
TODO:
|
|
6 |
move parsetree to parser.ML
|
|
7 |
move pt_to_ast before ast_to_term (sextension.ML (?))
|
|
8 |
delete this file
|
0
|
9 |
*)
|
|
10 |
|
|
11 |
signature PARSE_TREE =
|
|
12 |
sig
|
18
|
13 |
structure Ast: AST
|
0
|
14 |
structure Lexicon: LEXICON
|
18
|
15 |
local open Ast Lexicon in
|
|
16 |
datatype parsetree =
|
|
17 |
Node of string * parsetree list |
|
|
18 |
Tip of token
|
|
19 |
val pt_to_ast: (string -> (ast list -> ast) option) -> parsetree -> ast
|
0
|
20 |
end
|
|
21 |
end;
|
|
22 |
|
18
|
23 |
functor ParseTreeFun(structure Ast: AST and Lexicon: LEXICON): PARSE_TREE =
|
0
|
24 |
struct
|
|
25 |
|
18
|
26 |
structure Ast = Ast;
|
0
|
27 |
structure Lexicon = Lexicon;
|
18
|
28 |
open Ast Lexicon;
|
0
|
29 |
|
|
30 |
|
18
|
31 |
(* datatype parsetree *)
|
0
|
32 |
|
18
|
33 |
datatype parsetree =
|
|
34 |
Node of string * parsetree list |
|
|
35 |
Tip of token;
|
0
|
36 |
|
|
37 |
|
|
38 |
(* pt_to_ast *)
|
|
39 |
|
|
40 |
fun pt_to_ast trf pt =
|
|
41 |
let
|
|
42 |
fun trans a args =
|
|
43 |
(case trf a of
|
|
44 |
None => mk_appl (Constant a) args
|
18
|
45 |
| Some f => f args handle exn
|
|
46 |
=> (writeln ("Error in parse ast translation for " ^ quote a); raise exn));
|
0
|
47 |
|
18
|
48 |
fun ast_of (Node (a, pts)) = trans a (map ast_of pts)
|
|
49 |
| ast_of (Tip tok) = Variable (str_of_token tok);
|
0
|
50 |
in
|
18
|
51 |
ast_of pt
|
0
|
52 |
end;
|
|
53 |
|
|
54 |
|
|
55 |
end;
|
|
56 |
|