author | wenzelm |
Fri, 19 Aug 1994 15:38:50 +0200 | |
changeset 558 | c4092ae47210 |
parent 555 | a7f397a14b16 |
child 624 | 33b9b5da3e6f |
permissions | -rw-r--r-- |
240 | 1 |
(* Title: Pure/Syntax/syn_ext.ML |
2 |
ID: $Id$ |
|
3 |
Author: Markus Wenzel, TU Muenchen |
|
4 |
||
5 |
Syntax extension (internal interface). |
|
6 |
*) |
|
7 |
||
8 |
signature SYN_EXT0 = |
|
9 |
sig |
|
10 |
val typeT: typ |
|
11 |
val constrainC: string |
|
12 |
end; |
|
13 |
||
14 |
signature SYN_EXT = |
|
15 |
sig |
|
16 |
include SYN_EXT0 |
|
17 |
structure Ast: AST |
|
18 |
local open Ast in |
|
19 |
val logic: string |
|
20 |
val args: string |
|
21 |
val applC: string |
|
22 |
val typ_to_nonterm: typ -> string |
|
23 |
datatype xsymb = |
|
24 |
Delim of string | |
|
25 |
Argument of string * int | |
|
26 |
Space of string | |
|
27 |
Bg of int | Brk of int | En |
|
28 |
datatype xprod = XProd of string * xsymb list * string * int |
|
29 |
val max_pri: int |
|
30 |
val chain_pri: int |
|
31 |
val delims_of: xprod list -> string list |
|
32 |
datatype mfix = Mfix of string * typ * string * int list * int |
|
33 |
datatype syn_ext = |
|
34 |
SynExt of { |
|
35 |
roots: string list, |
|
36 |
xprods: xprod list, |
|
37 |
consts: string list, |
|
38 |
parse_ast_translation: (string * (ast list -> ast)) list, |
|
39 |
parse_rules: (ast * ast) list, |
|
40 |
parse_translation: (string * (term list -> term)) list, |
|
41 |
print_translation: (string * (term list -> term)) list, |
|
42 |
print_rules: (ast * ast) list, |
|
43 |
print_ast_translation: (string * (ast list -> ast)) list} |
|
330
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
44 |
val syn_ext: string list -> string list -> mfix list -> string list -> |
240 | 45 |
(string * (ast list -> ast)) list * (string * (term list -> term)) list * |
46 |
(string * (term list -> term)) list * (string * (ast list -> ast)) list |
|
47 |
-> (ast * ast) list * (ast * ast) list -> syn_ext |
|
555
a7f397a14b16
removed idT, varT, tidT, tvarT (now in lexicon.ML);
wenzelm
parents:
441
diff
changeset
|
48 |
val syn_ext_roots: string list -> string list -> syn_ext |
a7f397a14b16
removed idT, varT, tidT, tvarT (now in lexicon.ML);
wenzelm
parents:
441
diff
changeset
|
49 |
val syn_ext_const_names: string list -> string list -> syn_ext |
330
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
50 |
val syn_ext_rules: string list -> (ast * ast) list * (ast * ast) list -> syn_ext |
555
a7f397a14b16
removed idT, varT, tidT, tvarT (now in lexicon.ML);
wenzelm
parents:
441
diff
changeset
|
51 |
val syn_ext_trfuns: string list -> |
a7f397a14b16
removed idT, varT, tidT, tvarT (now in lexicon.ML);
wenzelm
parents:
441
diff
changeset
|
52 |
(string * (ast list -> ast)) list * (string * (term list -> term)) list * |
a7f397a14b16
removed idT, varT, tidT, tvarT (now in lexicon.ML);
wenzelm
parents:
441
diff
changeset
|
53 |
(string * (term list -> term)) list * (string * (ast list -> ast)) list |
a7f397a14b16
removed idT, varT, tidT, tvarT (now in lexicon.ML);
wenzelm
parents:
441
diff
changeset
|
54 |
-> syn_ext |
240 | 55 |
end |
56 |
end; |
|
57 |
||
58 |
functor SynExtFun(structure Lexicon: LEXICON and Ast: AST): SYN_EXT = |
|
59 |
struct |
|
60 |
||
61 |
structure Ast = Ast; |
|
62 |
open Lexicon Ast; |
|
63 |
||
64 |
||
65 |
(** misc definitions **) |
|
66 |
||
67 |
(* syntactic categories *) |
|
68 |
||
69 |
val logic = "logic"; |
|
70 |
val logicT = Type (logic, []); |
|
71 |
||
72 |
val logic1 = "logic1"; |
|
73 |
val logic1T = Type (logic1, []); |
|
74 |
||
75 |
val args = "args"; |
|
76 |
val argsT = Type (args, []); |
|
77 |
||
330
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
78 |
val typeT = Type ("type", []); |
240 | 79 |
|
330
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
80 |
val funT = Type ("fun", []); |
240 | 81 |
|
82 |
||
83 |
(* constants *) |
|
84 |
||
85 |
val applC = "_appl"; |
|
86 |
val constrainC = "_constrain"; |
|
87 |
||
88 |
||
89 |
||
90 |
(** datatype xprod **) |
|
91 |
||
92 |
(*Delim s: delimiter s |
|
93 |
Argument (s, p): nonterminal s requiring priority >= p, or valued token |
|
94 |
Space s: some white space for printing |
|
95 |
Bg, Brk, En: blocks and breaks for pretty printing*) |
|
96 |
||
97 |
datatype xsymb = |
|
98 |
Delim of string | |
|
99 |
Argument of string * int | |
|
100 |
Space of string | |
|
101 |
Bg of int | Brk of int | En; |
|
102 |
||
103 |
||
104 |
(*XProd (lhs, syms, c, p): |
|
105 |
lhs: name of nonterminal on the lhs of the production |
|
106 |
syms: list of symbols on the rhs of the production |
|
107 |
c: head of parse tree |
|
108 |
p: priority of this production*) |
|
109 |
||
110 |
datatype xprod = XProd of string * xsymb list * string * int; |
|
111 |
||
112 |
val max_pri = 1000; (*maximum legal priority*) |
|
113 |
val chain_pri = ~1; (*dummy for chain productions*) |
|
114 |
||
115 |
||
116 |
(* delims_of *) |
|
117 |
||
118 |
fun delims_of xprods = |
|
119 |
let |
|
120 |
fun del_of (Delim s) = Some s |
|
121 |
| del_of _ = None; |
|
122 |
||
123 |
fun dels_of (XProd (_, xsymbs, _, _)) = |
|
124 |
mapfilter del_of xsymbs; |
|
125 |
in |
|
126 |
distinct (flat (map dels_of xprods)) |
|
127 |
end; |
|
128 |
||
129 |
||
130 |
||
131 |
(** datatype mfix **) |
|
132 |
||
133 |
(*Mfix (sy, ty, c, ps, p): |
|
134 |
sy: rhs of production as symbolic string |
|
135 |
ty: type description of production |
|
136 |
c: head of parse tree |
|
137 |
ps: priorities of arguments in sy |
|
138 |
p: priority of production*) |
|
139 |
||
140 |
datatype mfix = Mfix of string * typ * string * int list * int; |
|
141 |
||
142 |
||
143 |
(* typ_to_nonterm *) |
|
144 |
||
145 |
fun typ_to_nonterm (Type (c, _)) = c |
|
146 |
| typ_to_nonterm _ = logic; |
|
147 |
||
148 |
fun typ_to_nonterm1 (Type (c, _)) = c |
|
149 |
| typ_to_nonterm1 _ = logic1; |
|
150 |
||
151 |
||
152 |
(* mfix_to_xprod *) |
|
153 |
||
154 |
fun mfix_to_xprod (Mfix (sy, typ, const, pris, pri)) = |
|
155 |
let |
|
156 |
fun err msg = |
|
157 |
(writeln ("Error in mixfix annotation " ^ quote sy ^ " for " ^ quote const); |
|
158 |
error msg); |
|
159 |
||
160 |
fun check_pri p = |
|
161 |
if p >= 0 andalso p <= max_pri then () |
|
162 |
else err ("precedence out of range: " ^ string_of_int p); |
|
163 |
||
164 |
fun blocks_ok [] 0 = true |
|
165 |
| blocks_ok [] _ = false |
|
166 |
| blocks_ok (Bg _ :: syms) n = blocks_ok syms (n + 1) |
|
167 |
| blocks_ok (En :: _) 0 = false |
|
168 |
| blocks_ok (En :: syms) n = blocks_ok syms (n - 1) |
|
169 |
| blocks_ok (_ :: syms) n = blocks_ok syms n; |
|
170 |
||
171 |
fun check_blocks syms = |
|
172 |
if blocks_ok syms 0 then () |
|
173 |
else err "unbalanced block parentheses"; |
|
174 |
||
175 |
||
176 |
fun is_meta c = c mem ["(", ")", "/", "_"]; |
|
177 |
||
178 |
fun scan_delim_char ("'" :: c :: cs) = |
|
179 |
if is_blank c then err "illegal spaces in delimiter" else (c, cs) |
|
180 |
| scan_delim_char ["'"] = err "trailing escape character" |
|
181 |
| scan_delim_char (chs as c :: cs) = |
|
182 |
if is_blank c orelse is_meta c then raise LEXICAL_ERROR else (c, cs) |
|
183 |
| scan_delim_char [] = raise LEXICAL_ERROR; |
|
184 |
||
185 |
val scan_symb = |
|
186 |
$$ "_" >> K (Argument ("", 0)) || |
|
187 |
$$ "(" -- scan_int >> (Bg o #2) || |
|
188 |
$$ ")" >> K En || |
|
189 |
$$ "/" -- $$ "/" >> K (Brk ~1) || |
|
190 |
$$ "/" -- scan_any is_blank >> (Brk o length o #2) || |
|
191 |
scan_any1 is_blank >> (Space o implode) || |
|
192 |
repeat1 scan_delim_char >> (Delim o implode); |
|
193 |
||
194 |
||
195 |
val cons_fst = apfst o cons; |
|
196 |
||
197 |
fun add_args [] ty [] = ([], typ_to_nonterm1 ty) |
|
198 |
| add_args [] _ _ = err "too many precedences" |
|
199 |
| add_args (Argument _ :: syms) (Type ("fun", [ty, tys])) [] = |
|
200 |
cons_fst (Argument (typ_to_nonterm ty, 0)) (add_args syms tys []) |
|
201 |
| add_args (Argument _ :: syms) (Type ("fun", [ty, tys])) (p :: ps) = |
|
202 |
cons_fst (Argument (typ_to_nonterm ty, p)) (add_args syms tys ps) |
|
203 |
| add_args (Argument _ :: _) _ _ = |
|
204 |
err "more arguments than in corresponding type" |
|
205 |
| add_args (sym :: syms) ty ps = cons_fst sym (add_args syms ty ps); |
|
206 |
||
207 |
||
208 |
fun is_arg (Argument _) = true |
|
209 |
| is_arg _ = false; |
|
210 |
||
211 |
fun is_term (Delim _) = true |
|
212 |
| is_term (Argument (s, _)) = is_terminal s |
|
213 |
| is_term _ = false; |
|
214 |
||
215 |
fun rem_pri (Argument (s, _)) = Argument (s, chain_pri) |
|
216 |
| rem_pri sym = sym; |
|
217 |
||
218 |
||
219 |
val (raw_symbs, _) = repeat scan_symb (explode sy); |
|
220 |
val (symbs, lhs) = add_args raw_symbs typ pris; |
|
221 |
val xprod = XProd (lhs, symbs, const, pri); |
|
222 |
in |
|
223 |
seq check_pri pris; |
|
224 |
check_pri pri; |
|
225 |
check_blocks symbs; |
|
226 |
||
227 |
if is_terminal lhs then err ("illegal lhs: " ^ lhs) |
|
228 |
else if const <> "" then xprod |
|
229 |
else if length (filter is_arg symbs) <> 1 then |
|
230 |
err "copy production must have exactly one argument" |
|
231 |
else if exists is_term symbs then xprod |
|
232 |
else XProd (lhs, map rem_pri symbs, "", chain_pri) |
|
233 |
end; |
|
234 |
||
235 |
||
236 |
(** datatype syn_ext **) |
|
237 |
||
238 |
datatype syn_ext = |
|
239 |
SynExt of { |
|
240 |
roots: string list, |
|
241 |
xprods: xprod list, |
|
242 |
consts: string list, |
|
243 |
parse_ast_translation: (string * (ast list -> ast)) list, |
|
244 |
parse_rules: (ast * ast) list, |
|
245 |
parse_translation: (string * (term list -> term)) list, |
|
246 |
print_translation: (string * (term list -> term)) list, |
|
247 |
print_rules: (ast * ast) list, |
|
248 |
print_ast_translation: (string * (ast list -> ast)) list}; |
|
249 |
||
250 |
||
251 |
(* syn_ext *) |
|
252 |
||
330
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
253 |
fun syn_ext all_roots new_roots mfixes consts trfuns rules = |
240 | 254 |
let |
255 |
val (parse_ast_translation, parse_translation, print_translation, |
|
256 |
print_ast_translation) = trfuns; |
|
257 |
val (parse_rules, print_rules) = rules; |
|
258 |
||
330
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
259 |
val Troots = map (apr (Type, [])) new_roots; |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
260 |
val Troots' = Troots \\ [typeT, propT]; |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
261 |
|
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
262 |
fun change_name T ext = |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
263 |
let val Type (name, ts) = T |
373
68400ea32f7b
fixed a bug in syntax_error, added "Building new grammar" message;
clasohm
parents:
369
diff
changeset
|
264 |
in Type ("@" ^ name ^ ext, ts) end; |
330
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
265 |
|
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
266 |
(* Append "_H" to lhs if production is not a copy or chain production *) |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
267 |
fun hide_xprod roots (XProd (lhs, symbs, const, pri)) = |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
268 |
let fun is_delim (Delim _) = true |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
269 |
| is_delim _ = false |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
270 |
in if const <> "" andalso lhs mem roots andalso exists is_delim symbs then |
373
68400ea32f7b
fixed a bug in syntax_error, added "Building new grammar" message;
clasohm
parents:
369
diff
changeset
|
271 |
XProd ("@" ^ lhs ^ "_H", symbs, const, pri) |
330
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
272 |
else XProd (lhs, symbs, const, pri) |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
273 |
end; |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
274 |
|
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
275 |
(* Make descend production and append "_H" to rhs nonterminal *) |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
276 |
fun descend_right (from, to) = |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
277 |
Mfix ("_", change_name to "_H" --> from, "", [0], 0); |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
278 |
|
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
279 |
(* Make descend production and append "_H" to lhs *) |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
280 |
fun descend_left (from, to) = |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
281 |
Mfix ("_", to --> change_name from "_H", "", [0], 0); |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
282 |
|
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
283 |
(* Make descend production and append "_A" to lhs *) |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
284 |
fun descend1 (from, to) = |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
285 |
Mfix ("_", to --> change_name from "_A", "", [0], 0); |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
286 |
|
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
287 |
(* Make parentheses production for 'hidden' and 'automatic' nonterminal *) |
555
a7f397a14b16
removed idT, varT, tidT, tvarT (now in lexicon.ML);
wenzelm
parents:
441
diff
changeset
|
288 |
fun parents T = |
330
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
289 |
if T = typeT then |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
290 |
[Mfix ("'(_')", T --> T, "", [0], max_pri)] |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
291 |
else |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
292 |
[Mfix ("'(_')", change_name T "_H" --> change_name T "_H", "", [0], max_pri), |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
293 |
Mfix ("'(_')", change_name T "_A" --> change_name T "_A", "", [0], max_pri)]; |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
294 |
|
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
295 |
fun mkappl T = |
555
a7f397a14b16
removed idT, varT, tidT, tvarT (now in lexicon.ML);
wenzelm
parents:
441
diff
changeset
|
296 |
Mfix ("(1_/(1'(_')))", [funT, argsT] ---> change_name T "_A", applC, |
330
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
297 |
[max_pri, 0], max_pri); |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
298 |
|
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
299 |
fun mkid T = |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
300 |
Mfix ("_", idT --> change_name T "_A", "", [], max_pri); |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
301 |
|
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
302 |
fun mkvar T = |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
303 |
Mfix ("_", varT --> change_name T "_A", "", [], max_pri); |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
304 |
|
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
305 |
fun constrain T = |
555
a7f397a14b16
removed idT, varT, tidT, tvarT (now in lexicon.ML);
wenzelm
parents:
441
diff
changeset
|
306 |
Mfix ("_::_", [T, typeT] ---> change_name T "_A", constrainC, |
441 | 307 |
[4, 0], 3) |
330
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
308 |
|
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
309 |
fun unhide T = |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
310 |
if T <> logicT then |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
311 |
[Mfix ("_", change_name T "_H" --> T, "", [0], 0), |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
312 |
Mfix ("_", change_name T "_A" --> T, "", [0], 0)] |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
313 |
else |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
314 |
[Mfix ("_", change_name T "_A" --> T, "", [0], 0)]; |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
315 |
|
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
316 |
val mfixes' = flat (map parents Troots) @ map mkappl Troots' @ |
240 | 317 |
map mkid Troots' @ map mkvar Troots' @ map constrain Troots' @ |
330
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
318 |
map (apl (logicT, descend_right)) (Troots \\ [logicT, typeT]) @ |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
319 |
map (apr (descend1, logic1T)) (Troots') @ |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
320 |
flat (map unhide (Troots \\ [typeT])); |
240 | 321 |
val mfix_consts = |
368 | 322 |
distinct (map (fn (Mfix (_, _, c, _, _)) => c) (mfixes @ mfixes')); |
330
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
323 |
val xprods = map mfix_to_xprod mfixes; |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
324 |
val xprods' = map mfix_to_xprod mfixes'; |
240 | 325 |
in |
326 |
SynExt { |
|
330
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
327 |
roots = new_roots, |
555
a7f397a14b16
removed idT, varT, tidT, tvarT (now in lexicon.ML);
wenzelm
parents:
441
diff
changeset
|
328 |
xprods = (map (hide_xprod (all_roots \\ ["logic", "type"])) xprods) |
330
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
329 |
@ xprods', (* hide only productions that weren't created |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
330 |
automatically *) |
368 | 331 |
consts = filter is_xid (consts union mfix_consts), |
240 | 332 |
parse_ast_translation = parse_ast_translation, |
333 |
parse_rules = parse_rules, |
|
334 |
parse_translation = parse_translation, |
|
335 |
print_translation = print_translation, |
|
336 |
print_rules = print_rules, |
|
337 |
print_ast_translation = print_ast_translation} |
|
338 |
end; |
|
339 |
||
340 |
||
555
a7f397a14b16
removed idT, varT, tidT, tvarT (now in lexicon.ML);
wenzelm
parents:
441
diff
changeset
|
341 |
fun syn_ext_roots roots new_roots = |
a7f397a14b16
removed idT, varT, tidT, tvarT (now in lexicon.ML);
wenzelm
parents:
441
diff
changeset
|
342 |
syn_ext roots new_roots [] [] ([], [], [], []) ([], []); |
a7f397a14b16
removed idT, varT, tidT, tvarT (now in lexicon.ML);
wenzelm
parents:
441
diff
changeset
|
343 |
|
a7f397a14b16
removed idT, varT, tidT, tvarT (now in lexicon.ML);
wenzelm
parents:
441
diff
changeset
|
344 |
fun syn_ext_const_names roots cs = |
a7f397a14b16
removed idT, varT, tidT, tvarT (now in lexicon.ML);
wenzelm
parents:
441
diff
changeset
|
345 |
syn_ext roots [] [] cs ([], [], [], []) ([], []); |
240 | 346 |
|
330
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
347 |
fun syn_ext_rules roots rules = |
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
clasohm
parents:
240
diff
changeset
|
348 |
syn_ext roots [] [] [] ([], [], [], []) rules; |
240 | 349 |
|
555
a7f397a14b16
removed idT, varT, tidT, tvarT (now in lexicon.ML);
wenzelm
parents:
441
diff
changeset
|
350 |
fun syn_ext_trfuns roots trfuns = |
a7f397a14b16
removed idT, varT, tidT, tvarT (now in lexicon.ML);
wenzelm
parents:
441
diff
changeset
|
351 |
syn_ext roots [] [] [] trfuns ([], []); |
240 | 352 |
|
353 |
||
354 |
end; |
|
355 |