18
|
1 |
(* Title: Pure/Syntax/extension.ML
|
0
|
2 |
ID: $Id$
|
|
3 |
Author: Tobias Nipkow and Markus Wenzel, TU Muenchen
|
|
4 |
|
18
|
5 |
External grammar definition (internal interface).
|
|
6 |
|
|
7 |
TODO:
|
|
8 |
Ext of {roots, mfix, extra_consts} * {.._translation} * {.._rules}
|
|
9 |
remove synrules
|
|
10 |
extend_xgram: move '.. \\ roots1' to Syntax.extend_tables
|
0
|
11 |
*)
|
|
12 |
|
|
13 |
signature EXTENSION0 =
|
|
14 |
sig
|
|
15 |
val typeT: typ
|
|
16 |
val constrainC: string
|
|
17 |
end;
|
|
18 |
|
|
19 |
signature EXTENSION =
|
|
20 |
sig
|
|
21 |
include EXTENSION0
|
|
22 |
structure XGram: XGRAM
|
|
23 |
local open XGram XGram.Ast in
|
|
24 |
datatype mfix = Mfix of string * typ * string * int list * int
|
|
25 |
datatype ext =
|
|
26 |
Ext of {
|
|
27 |
roots: string list,
|
|
28 |
mfix: mfix list,
|
|
29 |
extra_consts: string list,
|
|
30 |
parse_ast_translation: (string * (ast list -> ast)) list,
|
|
31 |
parse_translation: (string * (term list -> term)) list,
|
|
32 |
print_translation: (string * (term list -> term)) list,
|
|
33 |
print_ast_translation: (string * (ast list -> ast)) list}
|
|
34 |
datatype synrules =
|
|
35 |
SynRules of {
|
|
36 |
parse_rules: (ast * ast) list,
|
|
37 |
print_rules: (ast * ast) list}
|
|
38 |
val logic: string
|
18
|
39 |
val args: string
|
0
|
40 |
val idT: typ
|
|
41 |
val varT: typ
|
|
42 |
val tfreeT: typ
|
|
43 |
val tvarT: typ
|
18
|
44 |
val typ_to_nonterm: typ -> string
|
0
|
45 |
val applC: string
|
|
46 |
val empty_synrules: synrules
|
18
|
47 |
val empty_xgram: xgram
|
|
48 |
val extend_xgram: xgram -> (ext * synrules) -> xgram
|
|
49 |
val mk_xgram: (ext * synrules) -> xgram
|
0
|
50 |
end
|
|
51 |
end;
|
|
52 |
|
18
|
53 |
functor ExtensionFun(structure XGram: XGRAM and Lexicon: LEXICON): EXTENSION =
|
0
|
54 |
struct
|
|
55 |
|
|
56 |
structure XGram = XGram;
|
18
|
57 |
open XGram XGram.Ast Lexicon;
|
0
|
58 |
|
|
59 |
|
|
60 |
(** datatype ext **)
|
|
61 |
|
18
|
62 |
(*Mfix (sy, ty, c, ps, p):
|
|
63 |
sy: rhs of production as symbolic string
|
0
|
64 |
ty: type description of production
|
18
|
65 |
c: head of parse tree
|
|
66 |
ps: priorities of arguments in sy
|
|
67 |
p: priority of production*)
|
0
|
68 |
|
|
69 |
datatype mfix = Mfix of string * typ * string * int list * int;
|
|
70 |
|
|
71 |
datatype ext =
|
|
72 |
Ext of {
|
|
73 |
roots: string list,
|
|
74 |
mfix: mfix list,
|
|
75 |
extra_consts: string list,
|
|
76 |
parse_ast_translation: (string * (ast list -> ast)) list,
|
|
77 |
parse_translation: (string * (term list -> term)) list,
|
|
78 |
print_translation: (string * (term list -> term)) list,
|
|
79 |
print_ast_translation: (string * (ast list -> ast)) list};
|
|
80 |
|
|
81 |
datatype synrules =
|
|
82 |
SynRules of {
|
|
83 |
parse_rules: (ast * ast) list,
|
|
84 |
print_rules: (ast * ast) list};
|
|
85 |
|
|
86 |
|
|
87 |
(* empty_synrules *)
|
|
88 |
|
|
89 |
val empty_synrules = SynRules {parse_rules = [], print_rules = []};
|
|
90 |
|
|
91 |
|
18
|
92 |
(* empty_xgram *)
|
0
|
93 |
|
18
|
94 |
val empty_xgram =
|
0
|
95 |
XGram {
|
|
96 |
roots = [], prods = [], consts = [],
|
|
97 |
parse_ast_translation = [],
|
|
98 |
parse_rules = [],
|
|
99 |
parse_translation = [],
|
|
100 |
print_translation = [],
|
|
101 |
print_rules = [],
|
|
102 |
print_ast_translation = []};
|
|
103 |
|
|
104 |
|
18
|
105 |
(* syntactic categories *)
|
0
|
106 |
|
|
107 |
val logic = "logic";
|
|
108 |
val logicT = Type (logic, []);
|
|
109 |
|
|
110 |
val logic1 = "logic1";
|
|
111 |
val logic1T = Type (logic1, []);
|
|
112 |
|
18
|
113 |
val args = "args";
|
|
114 |
val argsT = Type (args, []);
|
|
115 |
|
0
|
116 |
val funT = Type ("fun", []);
|
|
117 |
|
18
|
118 |
val typeT = Type ("type", []);
|
|
119 |
|
0
|
120 |
|
|
121 |
(* terminals *)
|
|
122 |
|
|
123 |
val idT = Type (id, []);
|
|
124 |
val varT = Type (var, []);
|
|
125 |
val tfreeT = Type (tfree, []);
|
|
126 |
val tvarT = Type (tvar, []);
|
|
127 |
|
|
128 |
|
18
|
129 |
(* constants *)
|
0
|
130 |
|
|
131 |
val applC = "_appl";
|
|
132 |
val constrainC = "_constrain";
|
|
133 |
|
|
134 |
|
18
|
135 |
(* typ_to_nonterm *)
|
|
136 |
|
|
137 |
fun typ_to_nonterm (Type (c, _)) = c
|
|
138 |
| typ_to_nonterm _ = logic;
|
|
139 |
|
|
140 |
fun typ_to_nonterm1 (Type (c, _)) = c
|
|
141 |
| typ_to_nonterm1 _ = logic1;
|
0
|
142 |
|
|
143 |
|
|
144 |
|
18
|
145 |
(** mfix_to_prod **)
|
0
|
146 |
|
18
|
147 |
fun mfix_to_prod (Mfix (sy, typ, const, pris, pri)) =
|
|
148 |
let
|
|
149 |
fun err msg =
|
|
150 |
(writeln ("Error in mixfix annotation " ^ quote sy ^ " for " ^ quote const);
|
|
151 |
error msg);
|
0
|
152 |
|
18
|
153 |
fun check_pri p =
|
|
154 |
if p >= 0 andalso p <= max_pri then ()
|
|
155 |
else err ("precedence out of range: " ^ string_of_int p);
|
0
|
156 |
|
18
|
157 |
fun blocks_ok [] 0 = true
|
|
158 |
| blocks_ok [] _ = false
|
|
159 |
| blocks_ok (Bg _ :: syms) n = blocks_ok syms (n + 1)
|
|
160 |
| blocks_ok (En :: _) 0 = false
|
|
161 |
| blocks_ok (En :: syms) n = blocks_ok syms (n - 1)
|
|
162 |
| blocks_ok (_ :: syms) n = blocks_ok syms n;
|
0
|
163 |
|
18
|
164 |
fun check_blocks syms =
|
|
165 |
if blocks_ok syms 0 then ()
|
|
166 |
else err "unbalanced block parentheses";
|
0
|
167 |
|
|
168 |
|
18
|
169 |
fun is_meta c = c mem ["(", ")", "/", "_"];
|
|
170 |
|
|
171 |
fun scan_delim_char ("'" :: c :: cs) =
|
|
172 |
if is_blank c then err "illegal spaces in delimiter" else (c, cs)
|
|
173 |
| scan_delim_char ["'"] = err "trailing escape character"
|
|
174 |
| scan_delim_char (chs as c :: cs) =
|
|
175 |
if is_blank c orelse is_meta c then raise LEXICAL_ERROR else (c, cs)
|
|
176 |
| scan_delim_char [] = raise LEXICAL_ERROR;
|
|
177 |
|
|
178 |
val scan_symb =
|
|
179 |
$$ "_" >> K (Nonterminal ("", 0)) ||
|
|
180 |
$$ "(" -- scan_int >> (Bg o #2) ||
|
|
181 |
$$ ")" >> K En ||
|
|
182 |
$$ "/" -- $$ "/" >> K (Brk ~1) ||
|
|
183 |
$$ "/" -- scan_any is_blank >> (Brk o length o #2) ||
|
|
184 |
scan_any1 is_blank >> (Space o implode) ||
|
|
185 |
repeat1 scan_delim_char >> (Terminal o implode);
|
|
186 |
|
|
187 |
|
|
188 |
val cons_fst = apfst o cons;
|
0
|
189 |
|
18
|
190 |
fun add_args [] ty [] = ([], typ_to_nonterm1 ty)
|
|
191 |
| add_args [] _ _ = err "too many precedences"
|
|
192 |
| add_args (Nonterminal _ :: syms) (Type ("fun", [ty, tys])) [] =
|
|
193 |
cons_fst (Nonterminal (typ_to_nonterm ty, 0)) (add_args syms tys [])
|
|
194 |
| add_args (Nonterminal _ :: syms) (Type ("fun", [ty, tys])) (p :: ps) =
|
|
195 |
cons_fst (Nonterminal (typ_to_nonterm ty, p)) (add_args syms tys ps)
|
|
196 |
| add_args (Nonterminal _ :: _) _ _ =
|
|
197 |
err "more arguments than in corresponding type"
|
|
198 |
| add_args (sym :: syms) ty ps = cons_fst sym (add_args syms ty ps);
|
|
199 |
|
|
200 |
|
|
201 |
fun is_arg (Nonterminal _) = true
|
|
202 |
| is_arg _ = false;
|
|
203 |
|
|
204 |
fun is_term (Terminal _) = true
|
|
205 |
| is_term (Nonterminal (s, _)) = is_terminal s
|
|
206 |
| is_term _ = false;
|
|
207 |
|
|
208 |
fun rem_pri (Nonterminal (s, _)) = Nonterminal (s, chain_pri)
|
|
209 |
| rem_pri sym = sym;
|
|
210 |
|
|
211 |
|
|
212 |
val (raw_symbs, _) = repeat scan_symb (explode sy);
|
|
213 |
val (symbs, lhs) = add_args raw_symbs typ pris;
|
|
214 |
val prod = Prod (lhs, symbs, const, pri);
|
|
215 |
in
|
|
216 |
seq check_pri pris;
|
|
217 |
check_pri pri;
|
|
218 |
check_blocks symbs;
|
|
219 |
if is_terminal lhs then err ("illegal lhs: " ^ lhs) else ();
|
|
220 |
|
|
221 |
if const <> "" then prod
|
|
222 |
else if length (filter is_arg symbs) <> 1 then
|
|
223 |
err "copy production must have exactly one argument"
|
|
224 |
else if exists is_term symbs then prod
|
|
225 |
else Prod (lhs, map rem_pri symbs, "", chain_pri)
|
|
226 |
end;
|
0
|
227 |
|
|
228 |
|
|
229 |
|
18
|
230 |
(** extend_xgram **)
|
0
|
231 |
|
18
|
232 |
fun extend_xgram (XGram xgram) (Ext ext, SynRules rules) =
|
|
233 |
let
|
0
|
234 |
fun descend (from, to) = Mfix ("_", to --> from, "", [0], 0);
|
|
235 |
|
|
236 |
fun parents T = Mfix ("(1'(_'))", T --> T, "", [0], max_pri);
|
|
237 |
|
|
238 |
fun mkappl T =
|
116
|
239 |
Mfix ("(1_/(1'(_')))", [funT, argsT] ---> T, applC, [max_pri, 0], max_pri);
|
0
|
240 |
|
|
241 |
fun mkid T = Mfix ("_", idT --> T, "", [], max_pri);
|
|
242 |
|
|
243 |
fun mkvar T = Mfix ("_", varT --> T, "", [], max_pri);
|
|
244 |
|
|
245 |
fun constrain T =
|
18
|
246 |
Mfix ("_::_", [T, typeT] ---> T, constrainC, [max_pri, 0], max_pri - 1);
|
0
|
247 |
|
|
248 |
|
|
249 |
val {roots = roots1, prods, consts,
|
|
250 |
parse_ast_translation = parse_ast_translation1,
|
|
251 |
parse_rules = parse_rules1,
|
|
252 |
parse_translation = parse_translation1,
|
|
253 |
print_translation = print_translation1,
|
|
254 |
print_rules = print_rules1,
|
|
255 |
print_ast_translation = print_ast_translation1} = xgram;
|
|
256 |
|
|
257 |
val {roots = roots2, mfix, extra_consts,
|
|
258 |
parse_ast_translation = parse_ast_translation2,
|
|
259 |
parse_translation = parse_translation2,
|
|
260 |
print_translation = print_translation2,
|
|
261 |
print_ast_translation = print_ast_translation2} = ext;
|
|
262 |
|
|
263 |
val {parse_rules = parse_rules2, print_rules = print_rules2} = rules;
|
|
264 |
|
|
265 |
val Troots = map (apr (Type, [])) (roots2 \\ roots1);
|
|
266 |
val Troots' = Troots \\ [typeT, propT, logicT];
|
|
267 |
val mfix' = mfix @ map parents (Troots \ logicT) @ map mkappl Troots' @
|
|
268 |
map mkid Troots' @ map mkvar Troots' @ map constrain Troots' @
|
|
269 |
map (apl (logicT, descend)) (Troots \\ [typeT, logicT]) @
|
|
270 |
map (apr (descend, logic1T)) Troots';
|
18
|
271 |
val mfix_consts =
|
|
272 |
distinct (filter is_xid (map (fn (Mfix (_, _, c, _, _)) => c) mfix'));
|
0
|
273 |
in
|
|
274 |
XGram {
|
|
275 |
roots = distinct (roots1 @ roots2),
|
18
|
276 |
prods = prods @ map mfix_to_prod mfix',
|
|
277 |
consts = extra_consts union (mfix_consts union consts),
|
0
|
278 |
parse_ast_translation = parse_ast_translation1 @ parse_ast_translation2,
|
|
279 |
parse_rules = parse_rules1 @ parse_rules2,
|
|
280 |
parse_translation = parse_translation1 @ parse_translation2,
|
|
281 |
print_translation = print_translation1 @ print_translation2,
|
|
282 |
print_rules = print_rules1 @ print_rules2,
|
|
283 |
print_ast_translation = print_ast_translation1 @ print_ast_translation2}
|
|
284 |
end;
|
|
285 |
|
|
286 |
|
18
|
287 |
(* mk_xgram *)
|
|
288 |
|
|
289 |
val mk_xgram = extend_xgram empty_xgram;
|
|
290 |
|
|
291 |
|
0
|
292 |
end;
|
|
293 |
|