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