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