| author | traytel | 
| Tue, 10 Dec 2019 01:06:39 +0100 | |
| changeset 71261 | 4765fec43414 | 
| parent 69584 | a91e32843310 | 
| child 71545 | b0b16088ccf2 | 
| permissions | -rw-r--r-- | 
| 
42288
 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 
wenzelm 
parents: 
42268 
diff
changeset
 | 
1  | 
(* Title: Pure/Syntax/syntax_ext.ML  | 
| 
911
 
55754d6d399c
new in mixfix annotations: "' " (quote space) separates delimiters without
 
wenzelm 
parents: 
865 
diff
changeset
 | 
2  | 
Author: Markus Wenzel and Carsten Clasohm, TU Muenchen  | 
| 240 | 3  | 
|
| 
42288
 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 
wenzelm 
parents: 
42268 
diff
changeset
 | 
4  | 
Syntax extension.  | 
| 240 | 5  | 
*)  | 
6  | 
||
| 
42288
 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 
wenzelm 
parents: 
42268 
diff
changeset
 | 
7  | 
signature SYNTAX_EXT =  | 
| 4050 | 8  | 
sig  | 
| 62753 | 9  | 
datatype mfix = Mfix of Symbol_Pos.T list * typ * string * int list * int * Position.T  | 
| 1510 | 10  | 
val typ_to_nonterm: typ -> string  | 
| 62789 | 11  | 
  type block_info = {markup: Markup.T, consistent: bool, unbreakable: bool, indent: int}
 | 
| 62783 | 12  | 
val block_indent: int -> block_info  | 
| 1510 | 13  | 
datatype xsymb =  | 
14  | 
Delim of string |  | 
|
15  | 
Argument of string * int |  | 
|
16  | 
Space of string |  | 
|
| 62783 | 17  | 
Bg of block_info |  | 
18  | 
Brk of int |  | 
|
19  | 
En  | 
|
| 1510 | 20  | 
datatype xprod = XProd of string * xsymb list * string * int  | 
21  | 
val chain_pri: int  | 
|
| 4701 | 22  | 
val delims_of: xprod list -> string list list  | 
| 1510 | 23  | 
datatype syn_ext =  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
35429 
diff
changeset
 | 
24  | 
    Syn_Ext of {
 | 
| 1510 | 25  | 
xprods: xprod list,  | 
| 
42298
 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 
wenzelm 
parents: 
42297 
diff
changeset
 | 
26  | 
consts: (string * string) list,  | 
| 21772 | 27  | 
parse_ast_translation: (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list,  | 
| 1510 | 28  | 
parse_rules: (Ast.ast * Ast.ast) list,  | 
| 21772 | 29  | 
parse_translation: (string * ((Proof.context -> term list -> term) * stamp)) list,  | 
| 
42247
 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 
wenzelm 
parents: 
42245 
diff
changeset
 | 
30  | 
print_translation: (string * ((Proof.context -> typ -> term list -> term) * stamp)) list,  | 
| 1510 | 31  | 
print_rules: (Ast.ast * Ast.ast) list,  | 
| 42268 | 32  | 
print_ast_translation: (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list}  | 
| 62752 | 33  | 
val mfix_args: Symbol_Pos.T list -> int  | 
34  | 
val mixfix_args: Input.source -> int  | 
|
| 
42288
 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 
wenzelm 
parents: 
42268 
diff
changeset
 | 
35  | 
val escape: string -> string  | 
| 69584 | 36  | 
val syn_ext: string list -> mfix list ->  | 
| 
42298
 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 
wenzelm 
parents: 
42297 
diff
changeset
 | 
37  | 
(string * string) list -> (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list *  | 
| 21772 | 38  | 
(string * ((Proof.context -> term list -> term) * stamp)) list *  | 
| 
42247
 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 
wenzelm 
parents: 
42245 
diff
changeset
 | 
39  | 
(string * ((Proof.context -> typ -> term list -> term) * stamp)) list *  | 
| 42268 | 40  | 
(string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list ->  | 
41  | 
(Ast.ast * Ast.ast) list * (Ast.ast * Ast.ast) list -> syn_ext  | 
|
| 14903 | 42  | 
val syn_ext_rules: (Ast.ast * Ast.ast) list * (Ast.ast * Ast.ast) list -> syn_ext  | 
43  | 
val syn_ext_trfuns:  | 
|
| 21772 | 44  | 
(string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list *  | 
45  | 
(string * ((Proof.context -> term list -> term) * stamp)) list *  | 
|
| 
42247
 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 
wenzelm 
parents: 
42245 
diff
changeset
 | 
46  | 
(string * ((Proof.context -> typ -> term list -> term) * stamp)) list *  | 
| 21772 | 47  | 
(string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list -> syn_ext  | 
| 
42288
 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 
wenzelm 
parents: 
42268 
diff
changeset
 | 
48  | 
  val stamp_trfun: stamp -> string * 'a -> string * ('a * stamp)
 | 
| 
 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 
wenzelm 
parents: 
42268 
diff
changeset
 | 
49  | 
  val mk_trfun: string * 'a -> string * ('a * stamp)
 | 
| 
 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 
wenzelm 
parents: 
42268 
diff
changeset
 | 
50  | 
  val eq_trfun: ('a * stamp) * ('a * stamp) -> bool
 | 
| 4050 | 51  | 
end;  | 
| 240 | 52  | 
|
| 
42288
 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 
wenzelm 
parents: 
42268 
diff
changeset
 | 
53  | 
structure Syntax_Ext: SYNTAX_EXT =  | 
| 240 | 54  | 
struct  | 
55  | 
||
56  | 
(** datatype xprod **)  | 
|
57  | 
||
58  | 
(*Delim s: delimiter s  | 
|
59  | 
Argument (s, p): nonterminal s requiring priority >= p, or valued token  | 
|
60  | 
Space s: some white space for printing  | 
|
61  | 
Bg, Brk, En: blocks and breaks for pretty printing*)  | 
|
62  | 
||
| 62789 | 63  | 
type block_info = {markup: Markup.T, consistent: bool, unbreakable: bool, indent: int};
 | 
64  | 
||
65  | 
fun block_indent indent : block_info =  | 
|
66  | 
  {markup = Markup.empty, consistent = false, unbreakable = false, indent = indent};
 | 
|
| 62783 | 67  | 
|
| 240 | 68  | 
datatype xsymb =  | 
69  | 
Delim of string |  | 
|
70  | 
Argument of string * int |  | 
|
71  | 
Space of string |  | 
|
| 62783 | 72  | 
Bg of block_info |  | 
73  | 
Brk of int |  | 
|
74  | 
En;  | 
|
| 240 | 75  | 
|
| 
12513
 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 
wenzelm 
parents: 
11546 
diff
changeset
 | 
76  | 
fun is_delim (Delim _) = true  | 
| 
 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 
wenzelm 
parents: 
11546 
diff
changeset
 | 
77  | 
| is_delim _ = false;  | 
| 
 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 
wenzelm 
parents: 
11546 
diff
changeset
 | 
78  | 
|
| 
 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 
wenzelm 
parents: 
11546 
diff
changeset
 | 
79  | 
fun is_terminal (Delim _) = true  | 
| 
 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 
wenzelm 
parents: 
11546 
diff
changeset
 | 
80  | 
| is_terminal (Argument (s, _)) = Lexicon.is_terminal s  | 
| 
 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 
wenzelm 
parents: 
11546 
diff
changeset
 | 
81  | 
| is_terminal _ = false;  | 
| 
 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 
wenzelm 
parents: 
11546 
diff
changeset
 | 
82  | 
|
| 
 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 
wenzelm 
parents: 
11546 
diff
changeset
 | 
83  | 
fun is_argument (Argument _) = true  | 
| 
 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 
wenzelm 
parents: 
11546 
diff
changeset
 | 
84  | 
| is_argument _ = false;  | 
| 
 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 
wenzelm 
parents: 
11546 
diff
changeset
 | 
85  | 
|
| 
 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 
wenzelm 
parents: 
11546 
diff
changeset
 | 
86  | 
fun is_index (Argument ("index", _)) = true
 | 
| 
 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 
wenzelm 
parents: 
11546 
diff
changeset
 | 
87  | 
| is_index _ = false;  | 
| 
 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 
wenzelm 
parents: 
11546 
diff
changeset
 | 
88  | 
|
| 
 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 
wenzelm 
parents: 
11546 
diff
changeset
 | 
89  | 
val index = Argument ("index", 1000);
 | 
| 
 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 
wenzelm 
parents: 
11546 
diff
changeset
 | 
90  | 
|
| 240 | 91  | 
|
92  | 
(*XProd (lhs, syms, c, p):  | 
|
93  | 
lhs: name of nonterminal on the lhs of the production  | 
|
94  | 
syms: list of symbols on the rhs of the production  | 
|
95  | 
c: head of parse tree  | 
|
96  | 
p: priority of this production*)  | 
|
97  | 
||
98  | 
datatype xprod = XProd of string * xsymb list * string * int;  | 
|
99  | 
||
100  | 
val chain_pri = ~1; (*dummy for chain productions*)  | 
|
101  | 
||
102  | 
fun delims_of xprods =  | 
|
| 19004 | 103  | 
fold (fn XProd (_, xsymbs, _, _) =>  | 
104  | 
fold (fn Delim s => insert (op =) s | _ => I) xsymbs) xprods []  | 
|
105  | 
|> map Symbol.explode;  | 
|
| 240 | 106  | 
|
107  | 
||
108  | 
||
109  | 
(** datatype mfix **)  | 
|
110  | 
||
| 62753 | 111  | 
(*Mfix (sy, ty, c, ps, p, pos):  | 
| 62752 | 112  | 
sy: rhs of production as symbolic text  | 
| 240 | 113  | 
ty: type description of production  | 
114  | 
c: head of parse tree  | 
|
115  | 
ps: priorities of arguments in sy  | 
|
| 62753 | 116  | 
p: priority of production  | 
117  | 
pos: source position*)  | 
|
| 240 | 118  | 
|
| 62753 | 119  | 
datatype mfix = Mfix of Symbol_Pos.T list * typ * string * int list * int * Position.T;  | 
| 
12513
 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 
wenzelm 
parents: 
11546 
diff
changeset
 | 
120  | 
|
| 240 | 121  | 
|
122  | 
(* typ_to_nonterm *)  | 
|
123  | 
||
| 865 | 124  | 
fun typ_to_nt _ (Type (c, _)) = c  | 
125  | 
| typ_to_nt default _ = default;  | 
|
126  | 
||
| 
764
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
127  | 
(*get nonterminal for rhs*)  | 
| 
42293
 
6cca0343ea48
renamed sprop "prop#" to "prop'" -- proper identifier;
 
wenzelm 
parents: 
42288 
diff
changeset
 | 
128  | 
val typ_to_nonterm = typ_to_nt "any";  | 
| 240 | 129  | 
|
| 
764
 
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
 
clasohm 
parents: 
624 
diff
changeset
 | 
130  | 
(*get nonterminal for lhs*)  | 
| 
42293
 
6cca0343ea48
renamed sprop "prop#" to "prop'" -- proper identifier;
 
wenzelm 
parents: 
42288 
diff
changeset
 | 
131  | 
val typ_to_nonterm1 = typ_to_nt "logic";  | 
| 240 | 132  | 
|
133  | 
||
| 62786 | 134  | 
(* properties *)  | 
135  | 
||
136  | 
local  | 
|
137  | 
||
138  | 
open Basic_Symbol_Pos;  | 
|
139  | 
||
140  | 
val err_prefix = "Error in mixfix block properties: ";  | 
|
| 
62801
 
f9d102ef13f1
clarified errors -- disallow cartouche fragments as delimiter;
 
wenzelm 
parents: 
62795 
diff
changeset
 | 
141  | 
val !!! = Symbol_Pos.!!! (fn () => err_prefix ^ "atom expected (identifier, numeral, cartouche)");  | 
| 62786 | 142  | 
|
143  | 
val scan_atom =  | 
|
144  | 
Symbol_Pos.scan_ident ||  | 
|
145  | 
($$$ "-" @@@ (Symbol_Pos.scan_float || Symbol_Pos.scan_nat)) ||  | 
|
146  | 
Symbol_Pos.scan_float || Symbol_Pos.scan_nat ||  | 
|
147  | 
Symbol_Pos.scan_cartouche_content err_prefix;  | 
|
148  | 
||
149  | 
val scan_blanks = Scan.many (Symbol.is_blank o Symbol_Pos.symbol);  | 
|
150  | 
val scan_item =  | 
|
151  | 
scan_blanks |-- scan_atom --| scan_blanks  | 
|
152  | 
>> (fn ss => (Symbol_Pos.content ss, #1 (Symbol_Pos.range ss)));  | 
|
153  | 
||
154  | 
val scan_prop =  | 
|
155  | 
scan_item -- Scan.optional ($$ "=" |-- !!! scan_item >> #1) "true"  | 
|
156  | 
>> (fn ((x, pos), y) => (x, (y, pos)));  | 
|
157  | 
||
158  | 
fun get_property default parse name props =  | 
|
159  | 
(case AList.lookup (op =) props name of  | 
|
160  | 
NONE => default  | 
|
161  | 
| SOME (s, pos) =>  | 
|
162  | 
(parse s handle Fail msg => error (msg ^ " for property " ^ quote name ^ Position.here pos)));  | 
|
163  | 
||
164  | 
in  | 
|
165  | 
||
166  | 
fun read_properties ss =  | 
|
167  | 
let  | 
|
| 
62801
 
f9d102ef13f1
clarified errors -- disallow cartouche fragments as delimiter;
 
wenzelm 
parents: 
62795 
diff
changeset
 | 
168  | 
val props =  | 
| 
 
f9d102ef13f1
clarified errors -- disallow cartouche fragments as delimiter;
 
wenzelm 
parents: 
62795 
diff
changeset
 | 
169  | 
(case Scan.error (Scan.finite Symbol_Pos.stopper (Scan.repeat scan_prop)) ss of  | 
| 
 
f9d102ef13f1
clarified errors -- disallow cartouche fragments as delimiter;
 
wenzelm 
parents: 
62795 
diff
changeset
 | 
170  | 
(props, []) => props  | 
| 
 
f9d102ef13f1
clarified errors -- disallow cartouche fragments as delimiter;
 
wenzelm 
parents: 
62795 
diff
changeset
 | 
171  | 
| (_, (_, pos) :: _) => error (err_prefix ^ "bad input" ^ Position.here pos));  | 
| 62786 | 172  | 
val _ =  | 
| 62802 | 173  | 
(case AList.group (op =) props |> filter (fn (_, [_]) => false | _ => true) of  | 
| 62786 | 174  | 
[] => ()  | 
| 62802 | 175  | 
      | dups => error ("Duplicate properties: " ^ commas_quote (map #1 dups) ^
 | 
176  | 
Position.here_list (map #2 (maps #2 dups))));  | 
|
| 62786 | 177  | 
in props end;  | 
178  | 
||
179  | 
val get_string = get_property "" I;  | 
|
| 63806 | 180  | 
val get_bool = get_property false Value.parse_bool;  | 
181  | 
val get_nat = get_property 0 Value.parse_nat;  | 
|
| 62786 | 182  | 
|
183  | 
end;  | 
|
184  | 
||
185  | 
||
| 19004 | 186  | 
(* read mixfix annotations *)  | 
| 4050 | 187  | 
|
188  | 
local  | 
|
| 19004 | 189  | 
|
| 62752 | 190  | 
open Basic_Symbol_Pos;  | 
191  | 
||
| 
62801
 
f9d102ef13f1
clarified errors -- disallow cartouche fragments as delimiter;
 
wenzelm 
parents: 
62795 
diff
changeset
 | 
192  | 
val err_prefix = "Error in mixfix annotation: ";  | 
| 
 
f9d102ef13f1
clarified errors -- disallow cartouche fragments as delimiter;
 
wenzelm 
parents: 
62795 
diff
changeset
 | 
193  | 
|
| 62752 | 194  | 
fun scan_one pred = Scan.one (pred o Symbol_Pos.symbol);  | 
195  | 
fun scan_many pred = Scan.many (pred o Symbol_Pos.symbol);  | 
|
196  | 
fun scan_many1 pred = Scan.many1 (pred o Symbol_Pos.symbol);  | 
|
197  | 
||
| 62806 | 198  | 
fun reports_of_block pos =  | 
199  | 
[(pos, Markup.expression "mixfix block begin"), (pos, Markup.keyword3)];  | 
|
200  | 
||
| 62808 | 201  | 
fun reports_of (xsym, pos) =  | 
| 62806 | 202  | 
(case xsym of  | 
203  | 
Delim _ => [(pos, Markup.expression "mixfix delimiter"), (pos, Markup.literal)]  | 
|
204  | 
| Argument _ => [(pos, Markup.expression "mixfix argument")]  | 
|
205  | 
| Space _ => [(pos, Markup.expression "mixfix space")]  | 
|
206  | 
| Bg _ => reports_of_block pos  | 
|
207  | 
| Brk _ => [(pos, Markup.expression "mixfix break"), (pos, Markup.keyword3)]  | 
|
208  | 
| En => [(pos, Markup.expression "mixfix block end"), (pos, Markup.keyword3)]);  | 
|
209  | 
||
| 62808 | 210  | 
fun reports_text_of (xsym, pos) =  | 
211  | 
(case xsym of  | 
|
212  | 
Delim s =>  | 
|
213  | 
if Position.is_reported pos andalso exists Symbol.is_utf8 (Symbol.explode s) then  | 
|
| 
64677
 
8dc24130e8fe
more uniform treatment of "bad" like other messages (with serial number);
 
wenzelm 
parents: 
63933 
diff
changeset
 | 
214  | 
[((pos, Markup.bad ()),  | 
| 62808 | 215  | 
"Mixfix delimiter contains raw Unicode -- this is non-portable and unreliable")]  | 
216  | 
else []  | 
|
217  | 
| _ => []);  | 
|
218  | 
||
| 62786 | 219  | 
fun read_block_properties ss =  | 
220  | 
let  | 
|
221  | 
val props = read_properties ss;  | 
|
| 62783 | 222  | 
|
| 62786 | 223  | 
val markup_name = get_string Markup.markupN props;  | 
224  | 
val markup_props = fold (AList.delete (op =)) Markup.block_properties props;  | 
|
225  | 
val markup = (markup_name, map (apsnd #1) markup_props);  | 
|
226  | 
val _ =  | 
|
227  | 
if markup_name = "" andalso not (null markup_props) then  | 
|
| 62802 | 228  | 
        error ("Markup name required for block properties: " ^
 | 
229  | 
commas_quote (map #1 markup_props) ^ Position.here_list (map (#2 o #2) markup_props))  | 
|
| 62786 | 230  | 
else ();  | 
| 62783 | 231  | 
|
| 62786 | 232  | 
val consistent = get_bool Markup.consistentN props;  | 
| 62789 | 233  | 
val unbreakable = get_bool Markup.unbreakableN props;  | 
| 62786 | 234  | 
val indent = get_nat Markup.indentN props;  | 
| 62789 | 235  | 
  in Bg {markup = markup, consistent = consistent, unbreakable = unbreakable, indent = indent} end
 | 
| 62806 | 236  | 
handle ERROR msg =>  | 
237  | 
let  | 
|
238  | 
val reported_texts =  | 
|
239  | 
reports_of_block (#1 (Symbol_Pos.range ss))  | 
|
240  | 
|> map (fn (p, m) => Markup.markup_report (Position.reported_text p m ""))  | 
|
241  | 
in error (msg ^ implode reported_texts) end;  | 
|
| 62783 | 242  | 
|
243  | 
val read_block_indent =  | 
|
| 62789 | 244  | 
Bg o block_indent o #1 o Library.read_int o map Symbol_Pos.symbol;  | 
| 62783 | 245  | 
|
| 
62801
 
f9d102ef13f1
clarified errors -- disallow cartouche fragments as delimiter;
 
wenzelm 
parents: 
62795 
diff
changeset
 | 
246  | 
val is_meta = member (op =) ["(", ")", "/", "_", "\<index>", Symbol.open_, Symbol.close];
 | 
| 19004 | 247  | 
|
| 63933 | 248  | 
val scan_delim =  | 
249  | 
scan_one Symbol.is_control ::: Symbol_Pos.scan_cartouche "Mixfix error: " ||  | 
|
250  | 
$$ "'" |-- scan_one ((not o Symbol.is_blank) andf Symbol.not_eof) >> single ||  | 
|
251  | 
scan_one ((not o is_meta) andf (not o Symbol.is_blank) andf Symbol.not_eof) >> single;  | 
|
| 19004 | 252  | 
|
253  | 
val scan_sym =  | 
|
254  | 
  $$ "_" >> K (Argument ("", 0)) ||
 | 
|
| 
62529
 
8b7bdfc09f3b
clarified treatment of fragments of Isabelle symbols during bootstrap;
 
wenzelm 
parents: 
59841 
diff
changeset
 | 
255  | 
$$ "\<index>" >> K index ||  | 
| 62783 | 256  | 
  $$ "(" |--
 | 
| 
62801
 
f9d102ef13f1
clarified errors -- disallow cartouche fragments as delimiter;
 
wenzelm 
parents: 
62795 
diff
changeset
 | 
257  | 
(Symbol_Pos.scan_cartouche_content err_prefix >> read_block_properties ||  | 
| 62783 | 258  | 
scan_many Symbol.is_digit >> read_block_indent) ||  | 
| 19004 | 259  | 
$$ ")" >> K En ||  | 
260  | 
$$ "/" -- $$ "/" >> K (Brk ~1) ||  | 
|
| 62805 | 261  | 
$$ "/" |-- scan_many Symbol.is_space >> (Brk o length) ||  | 
262  | 
scan_many1 Symbol.is_space >> (Space o Symbol_Pos.content) ||  | 
|
| 63933 | 263  | 
Scan.repeat1 scan_delim >> (Delim o Symbol_Pos.content o flat);  | 
| 14819 | 264  | 
|
| 19004 | 265  | 
val scan_symb =  | 
| 
62795
 
063d2f23cdf6
removed redundant Position.set_range -- already done in Position.range;
 
wenzelm 
parents: 
62789 
diff
changeset
 | 
266  | 
Scan.trace scan_sym >> (fn (syms, trace) => SOME (syms, #1 (Symbol_Pos.range trace))) ||  | 
| 62805 | 267  | 
$$ "'" -- scan_one Symbol.is_space >> K NONE;  | 
| 19004 | 268  | 
|
| 19305 | 269  | 
val scan_symbs = Scan.repeat scan_symb --| Scan.ahead (~$$ "'");  | 
| 4050 | 270  | 
|
| 19004 | 271  | 
in  | 
| 
12513
 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 
wenzelm 
parents: 
11546 
diff
changeset
 | 
272  | 
|
| 62764 | 273  | 
fun read_mfix ss =  | 
274  | 
let  | 
|
| 
62801
 
f9d102ef13f1
clarified errors -- disallow cartouche fragments as delimiter;
 
wenzelm 
parents: 
62795 
diff
changeset
 | 
275  | 
val xsymbs =  | 
| 
 
f9d102ef13f1
clarified errors -- disallow cartouche fragments as delimiter;
 
wenzelm 
parents: 
62795 
diff
changeset
 | 
276  | 
(case Scan.error (Scan.finite Symbol_Pos.stopper scan_symbs) ss of  | 
| 
 
f9d102ef13f1
clarified errors -- disallow cartouche fragments as delimiter;
 
wenzelm 
parents: 
62795 
diff
changeset
 | 
277  | 
(res, []) => map_filter I res  | 
| 
 
f9d102ef13f1
clarified errors -- disallow cartouche fragments as delimiter;
 
wenzelm 
parents: 
62795 
diff
changeset
 | 
278  | 
| (_, (_, pos) :: _) => error (err_prefix ^ "bad input" ^ Position.here pos));  | 
| 62764 | 279  | 
val _ = Position.reports (maps reports_of xsymbs);  | 
| 62808 | 280  | 
val _ = Position.reports_text (maps reports_text_of xsymbs);  | 
| 62764 | 281  | 
in xsymbs end;  | 
| 19004 | 282  | 
|
| 62764 | 283  | 
val mfix_args = length o filter (is_argument o #1) o read_mfix o map (apsnd (K Position.none));  | 
| 62752 | 284  | 
val mixfix_args = mfix_args o Input.source_explode;  | 
| 19004 | 285  | 
|
| 35390 | 286  | 
val escape = implode o map (fn s => if is_meta s then "'" ^ s else s) o Symbol.explode;  | 
| 19004 | 287  | 
|
| 4050 | 288  | 
end;  | 
289  | 
||
290  | 
||
| 240 | 291  | 
(* mfix_to_xprod *)  | 
292  | 
||
| 62753 | 293  | 
fun mfix_to_xprod logical_types (Mfix (sy, typ, const, pris, pri, pos)) =  | 
| 240 | 294  | 
let  | 
| 62772 | 295  | 
val _ = Position.report pos Markup.language_mixfix;  | 
| 62764 | 296  | 
val symbs0 = read_mfix sy;  | 
297  | 
||
| 62762 | 298  | 
fun err_in_mixfix msg = error (msg ^ " in mixfix annotation" ^ Position.here pos);  | 
| 240 | 299  | 
|
| 62762 | 300  | 
fun check_blocks [] pending bad = pending @ bad  | 
301  | 
| check_blocks ((Bg _, pos) :: rest) pending bad = check_blocks rest (pos :: pending) bad  | 
|
302  | 
| check_blocks ((En, pos) :: rest) [] bad = check_blocks rest [] (pos :: bad)  | 
|
303  | 
| check_blocks ((En, _) :: rest) (_ :: pending) bad = check_blocks rest pending bad  | 
|
304  | 
| check_blocks (_ :: rest) pending bad = check_blocks rest pending bad;  | 
|
| 240 | 305  | 
|
306  | 
fun add_args [] ty [] = ([], typ_to_nonterm1 ty)  | 
|
| 62762 | 307  | 
| add_args [] _ _ = err_in_mixfix "Too many precedences"  | 
308  | 
      | add_args ((sym as (Argument ("index", _), _)) :: syms) ty ps =
 | 
|
309  | 
add_args syms ty ps |>> cons sym  | 
|
310  | 
      | add_args ((Argument _, pos) :: syms) (Type ("fun", [ty, tys])) [] =
 | 
|
311  | 
add_args syms tys [] |>> cons (Argument (typ_to_nonterm ty, 0), pos)  | 
|
312  | 
      | add_args ((Argument _, pos) :: syms) (Type ("fun", [ty, tys])) (p :: ps) =
 | 
|
313  | 
add_args syms tys ps |>> cons (Argument (typ_to_nonterm ty, p), pos)  | 
|
314  | 
| add_args ((Argument _, _) :: _) _ _ =  | 
|
315  | 
err_in_mixfix "More arguments than in corresponding type"  | 
|
316  | 
| add_args (sym :: syms) ty ps = add_args syms ty ps |>> cons sym;  | 
|
317  | 
||
318  | 
fun logical_args (a as (Argument (s, p))) =  | 
|
319  | 
          if s <> "prop" andalso member (op =) logical_types s then Argument ("logic", p) else a
 | 
|
320  | 
| logical_args a = a;  | 
|
| 240 | 321  | 
|
322  | 
fun rem_pri (Argument (s, _)) = Argument (s, chain_pri)  | 
|
323  | 
| rem_pri sym = sym;  | 
|
324  | 
||
| 62764 | 325  | 
val indexes = filter (is_index o #1) symbs0;  | 
| 62762 | 326  | 
val _ =  | 
327  | 
if length indexes <= 1 then ()  | 
|
328  | 
      else error ("More than one index argument" ^ Position.here_list (map #2 indexes));
 | 
|
| 2364 | 329  | 
|
| 62764 | 330  | 
val args = map_filter (fn (arg as Argument _, _) => SOME arg | _ => NONE) symbs0;  | 
| 
44470
 
6c6c31ef6bb2
more accurate treatment of index syntax constants, for proper entity references in concrete notation (e.g. infix "\<oplus>\<index>");
 
wenzelm 
parents: 
43329 
diff
changeset
 | 
331  | 
val (const', typ', syntax_consts, parse_rules) =  | 
| 
 
6c6c31ef6bb2
more accurate treatment of index syntax constants, for proper entity references in concrete notation (e.g. infix "\<oplus>\<index>");
 
wenzelm 
parents: 
43329 
diff
changeset
 | 
332  | 
if not (exists is_index args) then (const, typ, NONE, NONE)  | 
| 
12513
 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 
wenzelm 
parents: 
11546 
diff
changeset
 | 
333  | 
else  | 
| 
 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 
wenzelm 
parents: 
11546 
diff
changeset
 | 
334  | 
let  | 
| 
35429
 
afa8cf9e63d8
authentic syntax for classes and type constructors;
 
wenzelm 
parents: 
35390 
diff
changeset
 | 
335  | 
val indexed_const =  | 
| 
 
afa8cf9e63d8
authentic syntax for classes and type constructors;
 
wenzelm 
parents: 
35390 
diff
changeset
 | 
336  | 
if const <> "" then const ^ "_indexed"  | 
| 62762 | 337  | 
else err_in_mixfix "Missing constant name for indexed syntax";  | 
| 14697 | 338  | 
val rangeT = Term.range_type typ handle Match =>  | 
| 62762 | 339  | 
err_in_mixfix "Missing structure argument for indexed syntax";  | 
| 
12513
 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 
wenzelm 
parents: 
11546 
diff
changeset
 | 
340  | 
|
| 
43329
 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 
wenzelm 
parents: 
43323 
diff
changeset
 | 
341  | 
val xs = map Ast.Variable (Name.invent Name.context "xa" (length args - 1));  | 
| 19012 | 342  | 
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
 | 
343  | 
val i = Ast.Variable "i";  | 
| 62762 | 344  | 
val lhs =  | 
345  | 
Ast.mk_appl (Ast.Constant indexed_const)  | 
|
346  | 
(xs1 @ [Ast.mk_appl (Ast.Constant "_index") [i]] @ xs2);  | 
|
| 14697 | 347  | 
val rhs = Ast.mk_appl (Ast.Constant const) (i :: xs);  | 
| 
44470
 
6c6c31ef6bb2
more accurate treatment of index syntax constants, for proper entity references in concrete notation (e.g. infix "\<oplus>\<index>");
 
wenzelm 
parents: 
43329 
diff
changeset
 | 
348  | 
in (indexed_const, rangeT, SOME (indexed_const, const), SOME (lhs, rhs)) end;  | 
| 
12513
 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 
wenzelm 
parents: 
11546 
diff
changeset
 | 
349  | 
|
| 62764 | 350  | 
val (symbs1, lhs) = add_args symbs0 typ' pris;  | 
| 
12513
 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 
wenzelm 
parents: 
11546 
diff
changeset
 | 
351  | 
|
| 2364 | 352  | 
val copy_prod =  | 
| 20675 | 353  | 
(lhs = "prop" orelse lhs = "logic")  | 
| 2364 | 354  | 
andalso const <> ""  | 
| 62764 | 355  | 
andalso not (null symbs1)  | 
356  | 
andalso not (exists (is_delim o #1) symbs1);  | 
|
| 2364 | 357  | 
val lhs' =  | 
| 62764 | 358  | 
      if copy_prod orelse lhs = "prop" andalso map #1 symbs1 = [Argument ("prop'", 0)] then lhs
 | 
| 
42293
 
6cca0343ea48
renamed sprop "prop#" to "prop'" -- proper identifier;
 
wenzelm 
parents: 
42288 
diff
changeset
 | 
359  | 
else if lhs = "prop" then "prop'"  | 
| 62762 | 360  | 
else if member (op =) logical_types lhs then "logic"  | 
| 2364 | 361  | 
else lhs;  | 
| 62764 | 362  | 
val symbs2 = map (apfst logical_args) symbs1;  | 
| 240 | 363  | 
|
| 62762 | 364  | 
val _ =  | 
365  | 
(pri :: pris) |> List.app (fn p =>  | 
|
366  | 
if p >= 0 andalso p <= 1000 then ()  | 
|
367  | 
        else err_in_mixfix ("Precedence " ^ string_of_int p ^ " out of range"));
 | 
|
368  | 
val _ =  | 
|
| 62764 | 369  | 
(case check_blocks symbs2 [] [] of  | 
| 62762 | 370  | 
[] => ()  | 
371  | 
      | bad => error ("Unbalanced block parentheses" ^ Position.here_list bad));
 | 
|
372  | 
||
| 62764 | 373  | 
val xprod = XProd (lhs', map #1 symbs2, const', pri);  | 
| 
12513
 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 
wenzelm 
parents: 
11546 
diff
changeset
 | 
374  | 
val xprod' =  | 
| 62762 | 375  | 
if Lexicon.is_terminal lhs' then  | 
376  | 
        err_in_mixfix ("Illegal use of terminal " ^ quote lhs' ^ " as nonterminal")
 | 
|
| 
12513
 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 
wenzelm 
parents: 
11546 
diff
changeset
 | 
377  | 
else if const <> "" then xprod  | 
| 62764 | 378  | 
else if length (filter (is_argument o #1) symbs2) <> 1 then  | 
| 62762 | 379  | 
err_in_mixfix "Copy production must have exactly one argument"  | 
| 62764 | 380  | 
else if exists (is_terminal o #1) symbs2 then xprod  | 
381  | 
else XProd (lhs', map (rem_pri o #1) symbs2, "", chain_pri);  | 
|
| 
12513
 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 
wenzelm 
parents: 
11546 
diff
changeset
 | 
382  | 
|
| 
44470
 
6c6c31ef6bb2
more accurate treatment of index syntax constants, for proper entity references in concrete notation (e.g. infix "\<oplus>\<index>");
 
wenzelm 
parents: 
43329 
diff
changeset
 | 
383  | 
in (xprod', syntax_consts, parse_rules) end;  | 
| 
12513
 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 
wenzelm 
parents: 
11546 
diff
changeset
 | 
384  | 
|
| 240 | 385  | 
|
386  | 
||
387  | 
(** datatype syn_ext **)  | 
|
388  | 
||
389  | 
datatype syn_ext =  | 
|
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
35429 
diff
changeset
 | 
390  | 
  Syn_Ext of {
 | 
| 240 | 391  | 
xprods: xprod list,  | 
| 
42298
 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 
wenzelm 
parents: 
42297 
diff
changeset
 | 
392  | 
consts: (string * string) list,  | 
| 21772 | 393  | 
parse_ast_translation: (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list,  | 
| 1510 | 394  | 
parse_rules: (Ast.ast * Ast.ast) list,  | 
| 21772 | 395  | 
parse_translation: (string * ((Proof.context -> term list -> term) * stamp)) list,  | 
| 
42247
 
12fe41a92cd5
typed_print_translation: discontinued show_sorts argument;
 
wenzelm 
parents: 
42245 
diff
changeset
 | 
396  | 
print_translation: (string * ((Proof.context -> typ -> term list -> term) * stamp)) list,  | 
| 1510 | 397  | 
print_rules: (Ast.ast * Ast.ast) list,  | 
| 42268 | 398  | 
print_ast_translation: (string * ((Proof.context -> Ast.ast list -> Ast.ast) * stamp)) list};  | 
| 240 | 399  | 
|
400  | 
||
401  | 
(* syn_ext *)  | 
|
402  | 
||
| 69584 | 403  | 
fun syn_ext logical_types mfixes consts trfuns (parse_rules, print_rules) =  | 
| 240 | 404  | 
let  | 
405  | 
val (parse_ast_translation, parse_translation, print_translation,  | 
|
406  | 
print_ast_translation) = trfuns;  | 
|
407  | 
||
| 59841 | 408  | 
val xprod_results = map (mfix_to_xprod logical_types) mfixes;  | 
| 
44470
 
6c6c31ef6bb2
more accurate treatment of index syntax constants, for proper entity references in concrete notation (e.g. infix "\<oplus>\<index>");
 
wenzelm 
parents: 
43329 
diff
changeset
 | 
409  | 
val xprods = map #1 xprod_results;  | 
| 
 
6c6c31ef6bb2
more accurate treatment of index syntax constants, for proper entity references in concrete notation (e.g. infix "\<oplus>\<index>");
 
wenzelm 
parents: 
43329 
diff
changeset
 | 
410  | 
val consts' = map_filter #2 xprod_results;  | 
| 
 
6c6c31ef6bb2
more accurate treatment of index syntax constants, for proper entity references in concrete notation (e.g. infix "\<oplus>\<index>");
 
wenzelm 
parents: 
43329 
diff
changeset
 | 
411  | 
val parse_rules' = rev (map_filter #3 xprod_results);  | 
| 
42298
 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 
wenzelm 
parents: 
42297 
diff
changeset
 | 
412  | 
val mfix_consts = map (fn Mfix x => (#3 x, "")) mfixes @ map (fn XProd x => (#3 x, "")) xprods;  | 
| 240 | 413  | 
in  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
35429 
diff
changeset
 | 
414  | 
    Syn_Ext {
 | 
| 624 | 415  | 
xprods = xprods,  | 
| 
44470
 
6c6c31ef6bb2
more accurate treatment of index syntax constants, for proper entity references in concrete notation (e.g. infix "\<oplus>\<index>");
 
wenzelm 
parents: 
43329 
diff
changeset
 | 
416  | 
consts = mfix_consts @ consts' @ consts,  | 
| 240 | 417  | 
parse_ast_translation = parse_ast_translation,  | 
| 
12513
 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 
wenzelm 
parents: 
11546 
diff
changeset
 | 
418  | 
parse_rules = parse_rules' @ parse_rules,  | 
| 240 | 419  | 
parse_translation = parse_translation,  | 
420  | 
print_translation = print_translation,  | 
|
| 
12513
 
0ffb824dc95c
support for ``indexed syntax'' (using "\<index>" argument instead of "_");
 
wenzelm 
parents: 
11546 
diff
changeset
 | 
421  | 
print_rules = map swap parse_rules' @ print_rules,  | 
| 42268 | 422  | 
print_ast_translation = print_ast_translation}  | 
| 240 | 423  | 
end;  | 
424  | 
||
| 2382 | 425  | 
|
| 69584 | 426  | 
fun syn_ext_rules rules = syn_ext [] [] [] ([], [], [], []) rules;  | 
427  | 
fun syn_ext_trfuns trfuns = syn_ext [] [] [] trfuns ([], []);  | 
|
| 16610 | 428  | 
|
| 15754 | 429  | 
fun stamp_trfun s (c, f) = (c, (f, s));  | 
430  | 
fun mk_trfun tr = stamp_trfun (stamp ()) tr;  | 
|
| 
29318
 
6337d1cb2ba0
added numeral, which supercedes num, xnum, float;
 
wenzelm 
parents: 
28904 
diff
changeset
 | 
431  | 
fun eq_trfun ((_, s1: stamp), (_, s2)) = s1 = s2;  | 
| 15754 | 432  | 
|
| 240 | 433  | 
end;  |