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