| author | wenzelm | 
| Thu, 08 Jan 1998 18:24:45 +0100 | |
| changeset 4539 | 4227bd14dbe7 | 
| parent 4146 | 00136226f74b | 
| child 4701 | be8a8d60d962 | 
| permissions | -rw-r--r-- | 
| 240 | 1  | 
(* Title: Pure/Syntax/syn_ext.ML  | 
2  | 
ID: $Id$  | 
|
| 
911
 
55754d6d399c
new in mixfix annotations: "' " (quote space) separates delimiters without
 
wenzelm 
parents: 
865 
diff
changeset
 | 
3  | 
Author: Markus Wenzel and Carsten Clasohm, TU Muenchen  | 
| 240 | 4  | 
|
5  | 
Syntax extension (internal interface).  | 
|
6  | 
*)  | 
|
7  | 
||
8  | 
signature SYN_EXT0 =  | 
|
| 4050 | 9  | 
sig  | 
| 240 | 10  | 
val typeT: typ  | 
11  | 
val constrainC: string  | 
|
| 4050 | 12  | 
end;  | 
| 240 | 13  | 
|
14  | 
signature SYN_EXT =  | 
|
| 4050 | 15  | 
sig  | 
| 240 | 16  | 
include SYN_EXT0  | 
| 1510 | 17  | 
val logic: string  | 
18  | 
val args: string  | 
|
19  | 
val cargs: string  | 
|
20  | 
val any: string  | 
|
21  | 
val sprop: 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  | 
logtypes: string list,  | 
|
36  | 
xprods: xprod list,  | 
|
37  | 
consts: string list,  | 
|
| 2913 | 38  | 
prmodes: string list,  | 
| 1510 | 39  | 
parse_ast_translation: (string * (Ast.ast list -> Ast.ast)) list,  | 
40  | 
parse_rules: (Ast.ast * Ast.ast) list,  | 
|
41  | 
parse_translation: (string * (term list -> term)) list,  | 
|
| 4146 | 42  | 
print_translation: (string * (bool -> typ -> term list -> term)) list,  | 
| 1510 | 43  | 
print_rules: (Ast.ast * Ast.ast) list,  | 
| 2694 | 44  | 
print_ast_translation: (string * (Ast.ast list -> Ast.ast)) list,  | 
45  | 
token_translation: (string * string * (string -> string * int)) list}  | 
|
| 4054 | 46  | 
val mfix_args: string -> int  | 
| 1510 | 47  | 
val mk_syn_ext: bool -> string list -> mfix list ->  | 
48  | 
string list -> (string * (Ast.ast list -> Ast.ast)) list *  | 
|
49  | 
(string * (term list -> term)) list *  | 
|
| 4146 | 50  | 
(string * (bool -> typ -> term list -> term)) list * (string * (Ast.ast list -> Ast.ast)) list  | 
| 2694 | 51  | 
-> (string * string * (string -> string * int)) list  | 
| 1510 | 52  | 
-> (Ast.ast * Ast.ast) list * (Ast.ast * Ast.ast) list -> syn_ext  | 
53  | 
val syn_ext: string list -> mfix list -> string list ->  | 
|
54  | 
(string * (Ast.ast list -> Ast.ast)) list * (string * (term list -> term)) list *  | 
|
| 4146 | 55  | 
(string * (bool -> typ -> term list -> term)) list * (string * (Ast.ast list -> Ast.ast)) list  | 
| 2694 | 56  | 
-> (string * string * (string -> string * int)) list  | 
| 1510 | 57  | 
-> (Ast.ast * Ast.ast) list * (Ast.ast * Ast.ast) list -> syn_ext  | 
58  | 
val syn_ext_logtypes: string list -> syn_ext  | 
|
59  | 
val syn_ext_const_names: string list -> string list -> syn_ext  | 
|
60  | 
val syn_ext_rules: string list -> (Ast.ast * Ast.ast) list * (Ast.ast * Ast.ast) list -> syn_ext  | 
|
| 4146 | 61  | 
val fix_tr': (term list -> term) -> bool -> typ -> term list -> term  | 
| 1510 | 62  | 
val syn_ext_trfuns: string list ->  | 
63  | 
(string * (Ast.ast list -> Ast.ast)) list * (string * (term list -> term)) list *  | 
|
64  | 
(string * (term list -> term)) list * (string * (Ast.ast list -> Ast.ast)) list  | 
|
65  | 
-> syn_ext  | 
|
| 4146 | 66  | 
val syn_ext_trfunsT: string list ->  | 
67  | 
(string * (bool -> typ -> term list -> term)) list -> syn_ext  | 
|
| 2694 | 68  | 
val syn_ext_tokentrfuns: string list  | 
69  | 
-> (string * string * (string -> string * int)) list -> syn_ext  | 
|
| 1510 | 70  | 
val pure_ext: syn_ext  | 
| 4050 | 71  | 
end;  | 
| 240 | 72  | 
|
| 1510 | 73  | 
structure SynExt : SYN_EXT =  | 
| 240 | 74  | 
struct  | 
75  | 
||
76  | 
open Lexicon Ast;  | 
|
77  | 
||
| 2694 | 78  | 
|
| 240 | 79  | 
(** misc definitions **)  | 
80  | 
||
81  | 
(* syntactic categories *)  | 
|
82  | 
||
83  | 
val logic = "logic";  | 
|
84  | 
val logicT = Type (logic, []);  | 
|
85  | 
||
86  | 
val args = "args";  | 
|
| 1178 | 87  | 
val cargs = "cargs";  | 
| 240 | 88  | 
|
| 
330
 
2fda15dd1e0f
changed the way a grammar is generated to allow the new parser to work;
 
clasohm 
parents: 
240 
diff
changeset
 | 
89  | 
val typeT = Type ("type", []);
 | 
| 240 | 90  | 
|
| 
764
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
91  | 
val sprop = "#prop";  | 
| 
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
92  | 
val spropT = Type (sprop, []);  | 
| 240 | 93  | 
|
| 
764
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
94  | 
val any = "any";  | 
| 624 | 95  | 
val anyT = Type (any, []);  | 
96  | 
||
| 780 | 97  | 
|
| 240 | 98  | 
(* constants *)  | 
99  | 
||
100  | 
val constrainC = "_constrain";  | 
|
101  | 
||
102  | 
||
103  | 
||
104  | 
(** datatype xprod **)  | 
|
105  | 
||
106  | 
(*Delim s: delimiter s  | 
|
107  | 
Argument (s, p): nonterminal s requiring priority >= p, or valued token  | 
|
108  | 
Space s: some white space for printing  | 
|
109  | 
Bg, Brk, En: blocks and breaks for pretty printing*)  | 
|
110  | 
||
111  | 
datatype xsymb =  | 
|
112  | 
Delim of string |  | 
|
113  | 
Argument of string * int |  | 
|
114  | 
Space of string |  | 
|
115  | 
Bg of int | Brk of int | En;  | 
|
116  | 
||
117  | 
||
118  | 
(*XProd (lhs, syms, c, p):  | 
|
119  | 
lhs: name of nonterminal on the lhs of the production  | 
|
120  | 
syms: list of symbols on the rhs of the production  | 
|
121  | 
c: head of parse tree  | 
|
122  | 
p: priority of this production*)  | 
|
123  | 
||
124  | 
datatype xprod = XProd of string * xsymb list * string * int;  | 
|
125  | 
||
126  | 
val max_pri = 1000; (*maximum legal priority*)  | 
|
127  | 
val chain_pri = ~1; (*dummy for chain productions*)  | 
|
128  | 
||
129  | 
||
130  | 
(* delims_of *)  | 
|
131  | 
||
132  | 
fun delims_of xprods =  | 
|
133  | 
let  | 
|
134  | 
fun del_of (Delim s) = Some s  | 
|
135  | 
| del_of _ = None;  | 
|
136  | 
||
137  | 
fun dels_of (XProd (_, xsymbs, _, _)) =  | 
|
138  | 
mapfilter del_of xsymbs;  | 
|
139  | 
in  | 
|
140  | 
distinct (flat (map dels_of xprods))  | 
|
141  | 
end;  | 
|
142  | 
||
143  | 
||
144  | 
||
145  | 
(** datatype mfix **)  | 
|
146  | 
||
147  | 
(*Mfix (sy, ty, c, ps, p):  | 
|
148  | 
sy: rhs of production as symbolic string  | 
|
149  | 
ty: type description of production  | 
|
150  | 
c: head of parse tree  | 
|
151  | 
ps: priorities of arguments in sy  | 
|
152  | 
p: priority of production*)  | 
|
153  | 
||
154  | 
datatype mfix = Mfix of string * typ * string * int list * int;  | 
|
155  | 
||
156  | 
||
157  | 
(* typ_to_nonterm *)  | 
|
158  | 
||
| 865 | 159  | 
fun typ_to_nt _ (Type (c, _)) = c  | 
160  | 
| typ_to_nt default _ = default;  | 
|
161  | 
||
| 
764
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
162  | 
(*get nonterminal for rhs*)  | 
| 865 | 163  | 
val typ_to_nonterm = typ_to_nt any;  | 
| 240 | 164  | 
|
| 
764
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
165  | 
(*get nonterminal for lhs*)  | 
| 865 | 166  | 
val typ_to_nonterm1 = typ_to_nt logic;  | 
| 240 | 167  | 
|
168  | 
||
| 4050 | 169  | 
(* scan_mixfix, mixfix_args *)  | 
170  | 
||
171  | 
local  | 
|
172  | 
  fun is_meta c = c mem ["(", ")", "/", "_"];
 | 
|
173  | 
||
174  | 
  fun scan_delim_char ("'" :: c :: cs) =
 | 
|
175  | 
if is_blank c then raise LEXICAL_ERROR else (c, cs)  | 
|
176  | 
| scan_delim_char ["'"] = error "Trailing escape character"  | 
|
177  | 
| scan_delim_char (chs as c :: cs) =  | 
|
178  | 
if is_blank c orelse is_meta c then raise LEXICAL_ERROR else (c, cs)  | 
|
179  | 
| scan_delim_char [] = raise LEXICAL_ERROR;  | 
|
180  | 
||
181  | 
val scan_sym =  | 
|
182  | 
    $$ "_" >> K (Argument ("", 0)) ||
 | 
|
183  | 
    $$ "(" -- Term.scan_int >> (Bg o #2) ||
 | 
|
184  | 
$$ ")" >> K En ||  | 
|
185  | 
$$ "/" -- $$ "/" >> K (Brk ~1) ||  | 
|
186  | 
$$ "/" -- scan_any is_blank >> (Brk o length o #2) ||  | 
|
187  | 
scan_any1 is_blank >> (Space o implode) ||  | 
|
188  | 
repeat1 scan_delim_char >> (Delim o implode);  | 
|
189  | 
||
190  | 
val scan_symb =  | 
|
191  | 
scan_sym >> Some ||  | 
|
192  | 
$$ "'" -- scan_one is_blank >> K None;  | 
|
193  | 
||
194  | 
val scan_symbs = mapfilter I o #1 o repeat scan_symb;  | 
|
195  | 
in  | 
|
196  | 
val scan_mixfix = scan_symbs o SymbolFont.read_charnames o explode;  | 
|
197  | 
end;  | 
|
198  | 
||
| 4054 | 199  | 
fun mfix_args sy =  | 
200  | 
foldl (fn (i, Argument _) => i + 1 | (i, _) => i) (0, scan_mixfix sy);  | 
|
| 4050 | 201  | 
|
202  | 
||
| 240 | 203  | 
(* mfix_to_xprod *)  | 
204  | 
||
| 
764
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
205  | 
fun mfix_to_xprod convert logtypes (Mfix (sy, typ, const, pris, pri)) =  | 
| 240 | 206  | 
let  | 
207  | 
fun err msg =  | 
|
| 4050 | 208  | 
(if msg = "" then () else error_msg msg;  | 
209  | 
        error ("in mixfix annotation " ^ quote sy ^ " for " ^ quote const));
 | 
|
| 240 | 210  | 
|
211  | 
fun check_pri p =  | 
|
212  | 
if p >= 0 andalso p <= max_pri then ()  | 
|
| 4050 | 213  | 
      else err ("Precedence out of range: " ^ string_of_int p);
 | 
| 240 | 214  | 
|
215  | 
fun blocks_ok [] 0 = true  | 
|
216  | 
| blocks_ok [] _ = false  | 
|
217  | 
| blocks_ok (Bg _ :: syms) n = blocks_ok syms (n + 1)  | 
|
218  | 
| blocks_ok (En :: _) 0 = false  | 
|
219  | 
| blocks_ok (En :: syms) n = blocks_ok syms (n - 1)  | 
|
220  | 
| blocks_ok (_ :: syms) n = blocks_ok syms n;  | 
|
221  | 
||
222  | 
fun check_blocks syms =  | 
|
223  | 
if blocks_ok syms 0 then ()  | 
|
| 4050 | 224  | 
else err "Unbalanced block parentheses";  | 
| 240 | 225  | 
|
226  | 
||
227  | 
val cons_fst = apfst o cons;  | 
|
228  | 
||
229  | 
fun add_args [] ty [] = ([], typ_to_nonterm1 ty)  | 
|
| 4050 | 230  | 
| add_args [] _ _ = err "Too many precedences"  | 
| 240 | 231  | 
      | add_args (Argument _ :: syms) (Type ("fun", [ty, tys])) [] =
 | 
232  | 
cons_fst (Argument (typ_to_nonterm ty, 0)) (add_args syms tys [])  | 
|
233  | 
      | add_args (Argument _ :: syms) (Type ("fun", [ty, tys])) (p :: ps) =
 | 
|
234  | 
cons_fst (Argument (typ_to_nonterm ty, p)) (add_args syms tys ps)  | 
|
235  | 
| add_args (Argument _ :: _) _ _ =  | 
|
| 4050 | 236  | 
err "More arguments than in corresponding type"  | 
| 240 | 237  | 
| add_args (sym :: syms) ty ps = cons_fst sym (add_args syms ty ps);  | 
238  | 
||
239  | 
||
240  | 
fun is_arg (Argument _) = true  | 
|
241  | 
| is_arg _ = false;  | 
|
242  | 
||
243  | 
fun is_term (Delim _) = true  | 
|
244  | 
| is_term (Argument (s, _)) = is_terminal s  | 
|
245  | 
| is_term _ = false;  | 
|
246  | 
||
247  | 
fun rem_pri (Argument (s, _)) = Argument (s, chain_pri)  | 
|
248  | 
| rem_pri sym = sym;  | 
|
249  | 
||
| 
764
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
250  | 
fun is_delim (Delim _) = true  | 
| 
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
251  | 
| is_delim _ = false;  | 
| 
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
252  | 
|
| 4054 | 253  | 
fun logify_types copy_prod (a as (Argument (s, p))) =  | 
254  | 
if s mem logtypes then Argument (logic, p) else a  | 
|
255  | 
| logify_types _ a = a;  | 
|
| 240 | 256  | 
|
| 2364 | 257  | 
|
| 4050 | 258  | 
val raw_symbs = scan_mixfix sy handle ERROR => err "";  | 
| 240 | 259  | 
val (symbs, lhs) = add_args raw_symbs typ pris;  | 
| 2364 | 260  | 
val copy_prod =  | 
261  | 
lhs mem ["prop", "logic"]  | 
|
262  | 
andalso const <> ""  | 
|
263  | 
andalso not (null symbs)  | 
|
264  | 
andalso not (exists is_delim symbs);  | 
|
265  | 
val lhs' =  | 
|
266  | 
if convert andalso not copy_prod then  | 
|
267  | 
(if lhs mem logtypes then logic  | 
|
268  | 
else if lhs = "prop" then sprop else lhs)  | 
|
269  | 
else lhs;  | 
|
| 4054 | 270  | 
val symbs' = map (logify_types copy_prod) symbs;  | 
| 
764
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
271  | 
val xprod = XProd (lhs', symbs', const, pri);  | 
| 240 | 272  | 
in  | 
273  | 
seq check_pri pris;  | 
|
274  | 
check_pri pri;  | 
|
| 
764
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
275  | 
check_blocks symbs';  | 
| 240 | 276  | 
|
| 4050 | 277  | 
    if is_terminal lhs' then err ("Illegal lhs: " ^ lhs')
 | 
| 240 | 278  | 
else if const <> "" then xprod  | 
| 
764
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
279  | 
else if length (filter is_arg symbs') <> 1 then  | 
| 4050 | 280  | 
err "Copy production must have exactly one argument"  | 
| 
764
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
281  | 
else if exists is_term symbs' then xprod  | 
| 
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
282  | 
else XProd (lhs', map rem_pri symbs', "", chain_pri)  | 
| 240 | 283  | 
end;  | 
284  | 
||
285  | 
||
286  | 
(** datatype syn_ext **)  | 
|
287  | 
||
288  | 
datatype syn_ext =  | 
|
289  | 
  SynExt of {
 | 
|
| 
764
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
290  | 
logtypes: string list,  | 
| 240 | 291  | 
xprods: xprod list,  | 
292  | 
consts: string list,  | 
|
| 2913 | 293  | 
prmodes: string list,  | 
| 1510 | 294  | 
parse_ast_translation: (string * (Ast.ast list -> Ast.ast)) list,  | 
295  | 
parse_rules: (Ast.ast * Ast.ast) list,  | 
|
| 240 | 296  | 
parse_translation: (string * (term list -> term)) list,  | 
| 4146 | 297  | 
print_translation: (string * (bool -> typ -> term list -> term)) list,  | 
| 1510 | 298  | 
print_rules: (Ast.ast * Ast.ast) list,  | 
| 2694 | 299  | 
print_ast_translation: (string * (Ast.ast list -> Ast.ast)) list,  | 
| 2913 | 300  | 
token_translation: (string * string * (string -> string * int)) list}  | 
| 240 | 301  | 
|
302  | 
||
303  | 
(* syn_ext *)  | 
|
304  | 
||
| 2694 | 305  | 
fun mk_syn_ext convert logtypes mfixes consts trfuns tokentrfuns rules =  | 
| 240 | 306  | 
let  | 
307  | 
val (parse_ast_translation, parse_translation, print_translation,  | 
|
308  | 
print_ast_translation) = trfuns;  | 
|
309  | 
val (parse_rules, print_rules) = rules;  | 
|
| 
764
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
310  | 
val logtypes' = logtypes \ "prop";  | 
| 240 | 311  | 
|
| 624 | 312  | 
val mfix_consts = distinct (map (fn (Mfix (_, _, c, _, _)) => c) mfixes);  | 
| 
764
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
313  | 
val xprods = map (mfix_to_xprod convert logtypes') mfixes;  | 
| 240 | 314  | 
in  | 
315  | 
    SynExt {
 | 
|
| 
764
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
316  | 
logtypes = logtypes',  | 
| 624 | 317  | 
xprods = xprods,  | 
| 368 | 318  | 
consts = filter is_xid (consts union mfix_consts),  | 
| 2913 | 319  | 
prmodes = distinct (map (fn (m, _, _) => m) tokentrfuns),  | 
| 240 | 320  | 
parse_ast_translation = parse_ast_translation,  | 
321  | 
parse_rules = parse_rules,  | 
|
322  | 
parse_translation = parse_translation,  | 
|
323  | 
print_translation = print_translation,  | 
|
324  | 
print_rules = print_rules,  | 
|
| 2694 | 325  | 
print_ast_translation = print_ast_translation,  | 
326  | 
token_translation = tokentrfuns}  | 
|
| 240 | 327  | 
end;  | 
328  | 
||
| 2382 | 329  | 
|
| 
764
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
330  | 
val syn_ext = mk_syn_ext true;  | 
| 240 | 331  | 
|
| 
764
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
332  | 
fun syn_ext_logtypes logtypes =  | 
| 2694 | 333  | 
syn_ext logtypes [] [] ([], [], [], []) [] ([], []);  | 
| 
764
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
334  | 
|
| 
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
335  | 
fun syn_ext_const_names logtypes cs =  | 
| 2694 | 336  | 
syn_ext logtypes [] cs ([], [], [], []) [] ([], []);  | 
| 
555
 
a7f397a14b16
removed idT, varT, tidT, tvarT (now in lexicon.ML);
 
wenzelm 
parents: 
441 
diff
changeset
 | 
337  | 
|
| 
764
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
338  | 
fun syn_ext_rules logtypes rules =  | 
| 2694 | 339  | 
syn_ext logtypes [] [] ([], [], [], []) [] rules;  | 
| 
764
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
340  | 
|
| 4146 | 341  | 
fun fix_tr' f _ _ ts = f ts;  | 
| 2382 | 342  | 
|
343  | 
fun syn_ext_trfuns logtypes (atrs, trs, tr's, atr's) =  | 
|
| 2694 | 344  | 
syn_ext logtypes [] [] (atrs, trs, map (apsnd fix_tr') tr's, atr's) [] ([], []);  | 
| 2382 | 345  | 
|
346  | 
fun syn_ext_trfunsT logtypes tr's =  | 
|
| 2694 | 347  | 
syn_ext logtypes [] [] ([], [], tr's, []) [] ([], []);  | 
348  | 
||
349  | 
fun syn_ext_tokentrfuns logtypes tokentrfuns =  | 
|
350  | 
syn_ext logtypes [] [] ([], [], [], []) tokentrfuns ([], []);  | 
|
| 2382 | 351  | 
|
| 240 | 352  | 
|
| 
764
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
353  | 
(* pure_ext *)  | 
| 240 | 354  | 
|
| 
764
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
355  | 
val pure_ext = mk_syn_ext false []  | 
| 
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
356  | 
  [Mfix ("_", spropT --> propT, "", [0], 0),
 | 
| 
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
357  | 
   Mfix ("_", logicT --> anyT, "", [0], 0),
 | 
| 
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
358  | 
   Mfix ("_", spropT --> anyT, "", [0], 0),
 | 
| 
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
359  | 
   Mfix ("'(_')", logicT --> logicT, "", [0], max_pri),
 | 
| 
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
360  | 
   Mfix ("'(_')", spropT --> spropT, "", [0], max_pri),
 | 
| 
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
361  | 
   Mfix ("_::_",  [logicT, typeT] ---> logicT, "_constrain", [4, 0], 3),
 | 
| 
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
362  | 
   Mfix ("_::_",  [spropT, typeT] ---> spropT, "_constrain", [4, 0], 3)]
 | 
| 
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
363  | 
[]  | 
| 
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
364  | 
([], [], [], [])  | 
| 2694 | 365  | 
[]  | 
| 
764
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
366  | 
([], []);  | 
| 240 | 367  | 
|
368  | 
end;  |