| author | wenzelm | 
| Sun, 16 Jul 2023 19:13:08 +0200 | |
| changeset 78369 | ba71ea02d965 | 
| parent 69586 | 9171d1ce5a35 | 
| child 80748 | 43c4817375bf | 
| permissions | -rw-r--r-- | 
| 384 | 1  | 
(* Title: Pure/Syntax/mixfix.ML  | 
| 551 | 2  | 
Author: Tobias Nipkow and Markus Wenzel, TU Muenchen  | 
| 384 | 3  | 
|
| 
18719
 
dca3ae4f6dd6
moved pure syntax to Syntax/syntax.ML and pure_thy.ML;
 
wenzelm 
parents: 
18673 
diff
changeset
 | 
4  | 
Mixfix declarations, infixes, binders.  | 
| 384 | 5  | 
*)  | 
6  | 
||
| 
42287
 
d98eb048a2e4
discontinued special treatment of structure Mixfix;
 
wenzelm 
parents: 
42284 
diff
changeset
 | 
7  | 
signature BASIC_MIXFIX =  | 
| 
2199
 
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
 
wenzelm 
parents: 
1952 
diff
changeset
 | 
8  | 
sig  | 
| 384 | 9  | 
datatype mixfix =  | 
10  | 
NoSyn |  | 
|
| 62752 | 11  | 
Mixfix of Input.source * int list * int * Position.range |  | 
12  | 
Infix of Input.source * int * Position.range |  | 
|
13  | 
Infixl of Input.source * int * Position.range |  | 
|
14  | 
Infixr of Input.source * int * Position.range |  | 
|
15  | 
Binder of Input.source * int * int * Position.range |  | 
|
16  | 
Structure of Position.range  | 
|
| 
2199
 
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
 
wenzelm 
parents: 
1952 
diff
changeset
 | 
17  | 
end;  | 
| 384 | 18  | 
|
| 
42287
 
d98eb048a2e4
discontinued special treatment of structure Mixfix;
 
wenzelm 
parents: 
42284 
diff
changeset
 | 
19  | 
signature MIXFIX =  | 
| 
2199
 
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
 
wenzelm 
parents: 
1952 
diff
changeset
 | 
20  | 
sig  | 
| 
42287
 
d98eb048a2e4
discontinued special treatment of structure Mixfix;
 
wenzelm 
parents: 
42284 
diff
changeset
 | 
21  | 
include BASIC_MIXFIX  | 
| 62761 | 22  | 
val mixfix: string -> mixfix  | 
| 62752 | 23  | 
val is_empty: mixfix -> bool  | 
24  | 
val equal: mixfix * mixfix -> bool  | 
|
25  | 
val range_of: mixfix -> Position.range  | 
|
26  | 
val pos_of: mixfix -> Position.T  | 
|
| 62759 | 27  | 
val reset_pos: mixfix -> mixfix  | 
| 19271 | 28  | 
val pretty_mixfix: mixfix -> Pretty.T  | 
| 4823 | 29  | 
val mixfix_args: mixfix -> int  | 
| 
62959
 
19c2a58623ed
back to static Mixfix.default_constraint without any special tricks (reverting e6443edaebff);
 
wenzelm 
parents: 
62797 
diff
changeset
 | 
30  | 
val default_constraint: mixfix -> typ  | 
| 
35412
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
35390 
diff
changeset
 | 
31  | 
val make_type: int -> typ  | 
| 
42287
 
d98eb048a2e4
discontinued special treatment of structure Mixfix;
 
wenzelm 
parents: 
42284 
diff
changeset
 | 
32  | 
val binder_name: string -> string  | 
| 
42288
 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 
wenzelm 
parents: 
42287 
diff
changeset
 | 
33  | 
val syn_ext_types: (string * typ * mixfix) list -> Syntax_Ext.syn_ext  | 
| 59841 | 34  | 
val syn_ext_consts: string list -> (string * typ * mixfix) list -> Syntax_Ext.syn_ext  | 
| 
2199
 
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
 
wenzelm 
parents: 
1952 
diff
changeset
 | 
35  | 
end;  | 
| 
 
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
 
wenzelm 
parents: 
1952 
diff
changeset
 | 
36  | 
|
| 
15751
 
65e4790c7914
identify binder translations only once (admits remove);
 
wenzelm 
parents: 
15570 
diff
changeset
 | 
37  | 
structure Mixfix: MIXFIX =  | 
| 384 | 38  | 
struct  | 
39  | 
||
40  | 
(** mixfix declarations **)  | 
|
41  | 
||
42  | 
datatype mixfix =  | 
|
43  | 
NoSyn |  | 
|
| 62752 | 44  | 
Mixfix of Input.source * int list * int * Position.range |  | 
45  | 
Infix of Input.source * int * Position.range |  | 
|
46  | 
Infixl of Input.source * int * Position.range |  | 
|
47  | 
Infixr of Input.source * int * Position.range |  | 
|
48  | 
Binder of Input.source * int * int * Position.range |  | 
|
49  | 
Structure of Position.range;  | 
|
50  | 
||
| 62761 | 51  | 
fun mixfix s = Mixfix (Input.string s, [], 1000, Position.no_range);  | 
52  | 
||
| 62752 | 53  | 
fun is_empty NoSyn = true  | 
54  | 
| is_empty _ = false;  | 
|
55  | 
||
56  | 
fun equal (NoSyn, NoSyn) = true  | 
|
57  | 
| equal (Mixfix (sy, ps, p, _), Mixfix (sy', ps', p', _)) =  | 
|
58  | 
Input.equal_content (sy, sy') andalso ps = ps' andalso p = p'  | 
|
59  | 
| equal (Infix (sy, p, _), Infix (sy', p', _)) = Input.equal_content (sy, sy') andalso p = p'  | 
|
60  | 
| equal (Infixl (sy, p, _), Infixl (sy', p', _)) = Input.equal_content (sy, sy') andalso p = p'  | 
|
61  | 
| equal (Infixr (sy, p, _), Infixr (sy', p', _)) = Input.equal_content (sy, sy') andalso p = p'  | 
|
62  | 
| equal (Binder (sy, p, q, _), Binder (sy', p', q', _)) =  | 
|
63  | 
Input.equal_content (sy, sy') andalso p = p' andalso q = q'  | 
|
64  | 
| equal (Structure _, Structure _) = true  | 
|
65  | 
| equal _ = false;  | 
|
66  | 
||
67  | 
fun range_of NoSyn = Position.no_range  | 
|
68  | 
| range_of (Mixfix (_, _, _, range)) = range  | 
|
69  | 
| range_of (Infix (_, _, range)) = range  | 
|
70  | 
| range_of (Infixl (_, _, range)) = range  | 
|
71  | 
| range_of (Infixr (_, _, range)) = range  | 
|
72  | 
| range_of (Binder (_, _, _, range)) = range  | 
|
73  | 
| range_of (Structure range) = range;  | 
|
74  | 
||
| 62797 | 75  | 
val pos_of = Position.range_position o range_of;  | 
| 384 | 76  | 
|
| 62759 | 77  | 
fun reset_pos NoSyn = NoSyn  | 
78  | 
| reset_pos (Mixfix (sy, ps, p, _)) = Mixfix (Input.reset_pos sy, ps, p, Position.no_range)  | 
|
79  | 
| reset_pos (Infix (sy, p, _)) = Infix (Input.reset_pos sy, p, Position.no_range)  | 
|
80  | 
| reset_pos (Infixl (sy, p, _)) = Infixl (Input.reset_pos sy, p, Position.no_range)  | 
|
81  | 
| reset_pos (Infixr (sy, p, _)) = Infixr (Input.reset_pos sy, p, Position.no_range)  | 
|
82  | 
| reset_pos (Binder (sy, p, q, _)) = Binder (Input.reset_pos sy, p, q, Position.no_range)  | 
|
83  | 
| reset_pos (Structure _) = Structure Position.no_range;  | 
|
84  | 
||
| 384 | 85  | 
|
| 19271 | 86  | 
(* pretty_mixfix *)  | 
87  | 
||
88  | 
local  | 
|
| 11920 | 89  | 
|
| 69583 | 90  | 
val template = Pretty.cartouche o Pretty.str o #1 o Input.source_content;  | 
| 55763 | 91  | 
val keyword = Pretty.keyword2;  | 
| 19271 | 92  | 
val parens = Pretty.enclose "(" ")";
 | 
93  | 
val brackets = Pretty.enclose "[" "]";  | 
|
94  | 
val int = Pretty.str o string_of_int;  | 
|
95  | 
||
96  | 
in  | 
|
| 11920 | 97  | 
|
| 19271 | 98  | 
fun pretty_mixfix NoSyn = Pretty.str ""  | 
| 62752 | 99  | 
| pretty_mixfix (Mixfix (s, ps, p, _)) =  | 
| 
68273
 
53788963c4dc
pretty-print according to defaults of input syntax;
 
wenzelm 
parents: 
68271 
diff
changeset
 | 
100  | 
parens  | 
| 
 
53788963c4dc
pretty-print according to defaults of input syntax;
 
wenzelm 
parents: 
68271 
diff
changeset
 | 
101  | 
(Pretty.breaks  | 
| 69582 | 102  | 
(template s ::  | 
| 
68273
 
53788963c4dc
pretty-print according to defaults of input syntax;
 
wenzelm 
parents: 
68271 
diff
changeset
 | 
103  | 
(if null ps then [] else [brackets (Pretty.commas (map int ps))]) @  | 
| 
 
53788963c4dc
pretty-print according to defaults of input syntax;
 
wenzelm 
parents: 
68271 
diff
changeset
 | 
104  | 
(if p = 1000 then [] else [int p])))  | 
| 69582 | 105  | 
| pretty_mixfix (Infix (s, p, _)) = parens (Pretty.breaks [keyword "infix", template s, int p])  | 
106  | 
| pretty_mixfix (Infixl (s, p, _)) = parens (Pretty.breaks [keyword "infixl", template s, int p])  | 
|
107  | 
| pretty_mixfix (Infixr (s, p, _)) = parens (Pretty.breaks [keyword "infixr", template s, int p])  | 
|
| 62752 | 108  | 
| pretty_mixfix (Binder (s, p1, p2, _)) =  | 
| 
68273
 
53788963c4dc
pretty-print according to defaults of input syntax;
 
wenzelm 
parents: 
68271 
diff
changeset
 | 
109  | 
parens  | 
| 
 
53788963c4dc
pretty-print according to defaults of input syntax;
 
wenzelm 
parents: 
68271 
diff
changeset
 | 
110  | 
(Pretty.breaks  | 
| 69582 | 111  | 
([keyword "binder", template s] @  | 
112  | 
(if p1 = p2 then [] else [brackets [int p1]]) @ [int p2]))  | 
|
| 62752 | 113  | 
| pretty_mixfix (Structure _) = parens [keyword "structure"];  | 
| 19271 | 114  | 
|
115  | 
end;  | 
|
| 11920 | 116  | 
|
117  | 
||
| 12531 | 118  | 
(* syntax specifications *)  | 
| 384 | 119  | 
|
| 4053 | 120  | 
fun mixfix_args NoSyn = 0  | 
| 62752 | 121  | 
| mixfix_args (Mixfix (sy, _, _, _)) = Syntax_Ext.mixfix_args sy  | 
122  | 
| mixfix_args (Infix (sy, _, _)) = 2 + Syntax_Ext.mixfix_args sy  | 
|
123  | 
| mixfix_args (Infixl (sy, _, _)) = 2 + Syntax_Ext.mixfix_args sy  | 
|
124  | 
| mixfix_args (Infixr (sy, _, _)) = 2 + Syntax_Ext.mixfix_args sy  | 
|
| 18673 | 125  | 
| mixfix_args (Binder _) = 1  | 
| 62752 | 126  | 
| mixfix_args (Structure _) = 0;  | 
| 4053 | 127  | 
|
128  | 
||
| 
62959
 
19c2a58623ed
back to static Mixfix.default_constraint without any special tricks (reverting e6443edaebff);
 
wenzelm 
parents: 
62797 
diff
changeset
 | 
129  | 
(* default type constraint *)  | 
| 
 
19c2a58623ed
back to static Mixfix.default_constraint without any special tricks (reverting e6443edaebff);
 
wenzelm 
parents: 
62797 
diff
changeset
 | 
130  | 
|
| 
 
19c2a58623ed
back to static Mixfix.default_constraint without any special tricks (reverting e6443edaebff);
 
wenzelm 
parents: 
62797 
diff
changeset
 | 
131  | 
fun default_constraint (Binder _) = (dummyT --> dummyT) --> dummyT  | 
| 
 
19c2a58623ed
back to static Mixfix.default_constraint without any special tricks (reverting e6443edaebff);
 
wenzelm 
parents: 
62797 
diff
changeset
 | 
132  | 
| default_constraint mx = replicate (mixfix_args mx) dummyT ---> dummyT;  | 
| 
 
19c2a58623ed
back to static Mixfix.default_constraint without any special tricks (reverting e6443edaebff);
 
wenzelm 
parents: 
62797 
diff
changeset
 | 
133  | 
|
| 
 
19c2a58623ed
back to static Mixfix.default_constraint without any special tricks (reverting e6443edaebff);
 
wenzelm 
parents: 
62797 
diff
changeset
 | 
134  | 
|
| 
69586
 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 
wenzelm 
parents: 
69584 
diff
changeset
 | 
135  | 
(* mixfix template *)  | 
| 
 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 
wenzelm 
parents: 
69584 
diff
changeset
 | 
136  | 
|
| 
 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 
wenzelm 
parents: 
69584 
diff
changeset
 | 
137  | 
fun mixfix_template (Mixfix (sy, _, _, _)) = SOME sy  | 
| 
 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 
wenzelm 
parents: 
69584 
diff
changeset
 | 
138  | 
| mixfix_template (Infix (sy, _, _)) = SOME sy  | 
| 
 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 
wenzelm 
parents: 
69584 
diff
changeset
 | 
139  | 
| mixfix_template (Infixl (sy, _, _)) = SOME sy  | 
| 
 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 
wenzelm 
parents: 
69584 
diff
changeset
 | 
140  | 
| mixfix_template (Infixr (sy, _, _)) = SOME sy  | 
| 
 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 
wenzelm 
parents: 
69584 
diff
changeset
 | 
141  | 
| mixfix_template (Binder (sy, _, _, _)) = SOME sy  | 
| 
 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 
wenzelm 
parents: 
69584 
diff
changeset
 | 
142  | 
| mixfix_template _ = NONE;  | 
| 
 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 
wenzelm 
parents: 
69584 
diff
changeset
 | 
143  | 
|
| 
 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 
wenzelm 
parents: 
69584 
diff
changeset
 | 
144  | 
fun mixfix_template_reports mx =  | 
| 
 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 
wenzelm 
parents: 
69584 
diff
changeset
 | 
145  | 
if Options.default_bool "update_mixfix_cartouches" then  | 
| 
 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 
wenzelm 
parents: 
69584 
diff
changeset
 | 
146  | 
(case mixfix_template mx of  | 
| 
 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 
wenzelm 
parents: 
69584 
diff
changeset
 | 
147  | 
SOME sy => [((Input.pos_of sy, Markup.update), cartouche (#1 (Input.source_content sy)))]  | 
| 
 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 
wenzelm 
parents: 
69584 
diff
changeset
 | 
148  | 
| NONE => [])  | 
| 
 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 
wenzelm 
parents: 
69584 
diff
changeset
 | 
149  | 
else [];  | 
| 
 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 
wenzelm 
parents: 
69584 
diff
changeset
 | 
150  | 
|
| 
 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 
wenzelm 
parents: 
69584 
diff
changeset
 | 
151  | 
|
| 384 | 152  | 
(* syn_ext_types *)  | 
153  | 
||
| 
42293
 
6cca0343ea48
renamed sprop "prop#" to "prop'" -- proper identifier;
 
wenzelm 
parents: 
42288 
diff
changeset
 | 
154  | 
val typeT = Type ("type", []);
 | 
| 
 
6cca0343ea48
renamed sprop "prop#" to "prop'" -- proper identifier;
 
wenzelm 
parents: 
42288 
diff
changeset
 | 
155  | 
fun make_type n = replicate n typeT ---> typeT;  | 
| 
35412
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
35390 
diff
changeset
 | 
156  | 
|
| 14903 | 157  | 
fun syn_ext_types type_decls =  | 
| 384 | 158  | 
let  | 
| 62796 | 159  | 
fun mk_infix sy ty t p1 p2 p3 pos =  | 
| 62752 | 160  | 
Syntax_Ext.Mfix  | 
161  | 
(Symbol_Pos.explode0 "(_ " @ Input.source_explode sy @ Symbol_Pos.explode0 "/ _)",  | 
|
| 62796 | 162  | 
ty, t, [p1, p2], p3, pos);  | 
| 384 | 163  | 
|
| 62796 | 164  | 
fun mfix_of (mfix as (_, _, mx)) =  | 
165  | 
(case mfix of  | 
|
166  | 
(_, _, NoSyn) => NONE  | 
|
167  | 
| (t, ty, Mixfix (sy, ps, p, _)) =>  | 
|
168  | 
SOME (Syntax_Ext.Mfix (Input.source_explode sy, ty, t, ps, p, pos_of mx))  | 
|
169  | 
| (t, ty, Infix (sy, p, _)) => SOME (mk_infix sy ty t (p + 1) (p + 1) p (pos_of mx))  | 
|
170  | 
| (t, ty, Infixl (sy, p, _)) => SOME (mk_infix sy ty t p (p + 1) p (pos_of mx))  | 
|
171  | 
| (t, ty, Infixr (sy, p, _)) => SOME (mk_infix sy ty t (p + 1) p p (pos_of mx))  | 
|
172  | 
      | (t, _, _) => error ("Bad mixfix declaration for " ^ quote t ^ Position.here (pos_of mx)));
 | 
|
| 384 | 173  | 
|
| 62753 | 174  | 
fun check_args (_, ty, mx) (SOME (Syntax_Ext.Mfix (sy, _, _, _, _, _))) =  | 
| 
42288
 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 
wenzelm 
parents: 
42287 
diff
changeset
 | 
175  | 
if length (Term.binder_types ty) = Syntax_Ext.mfix_args sy then ()  | 
| 62752 | 176  | 
else  | 
177  | 
            error ("Bad number of type constructor arguments in mixfix annotation" ^
 | 
|
178  | 
Position.here (pos_of mx))  | 
|
| 
35351
 
7425aece4ee3
allow general mixfix syntax for type constructors;
 
wenzelm 
parents: 
35130 
diff
changeset
 | 
179  | 
| check_args _ NONE = ();  | 
| 
 
7425aece4ee3
allow general mixfix syntax for type constructors;
 
wenzelm 
parents: 
35130 
diff
changeset
 | 
180  | 
|
| 
69586
 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 
wenzelm 
parents: 
69584 
diff
changeset
 | 
181  | 
val _ = Position.reports_text (maps (mixfix_template_reports o #3) type_decls);  | 
| 
 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 
wenzelm 
parents: 
69584 
diff
changeset
 | 
182  | 
|
| 
35351
 
7425aece4ee3
allow general mixfix syntax for type constructors;
 
wenzelm 
parents: 
35130 
diff
changeset
 | 
183  | 
val mfix = map mfix_of type_decls;  | 
| 
 
7425aece4ee3
allow general mixfix syntax for type constructors;
 
wenzelm 
parents: 
35130 
diff
changeset
 | 
184  | 
val _ = map2 check_args type_decls mfix;  | 
| 
42298
 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 
wenzelm 
parents: 
42297 
diff
changeset
 | 
185  | 
val consts = map (fn (t, _, _) => (t, "")) type_decls;  | 
| 69584 | 186  | 
in Syntax_Ext.syn_ext [] (map_filter I mfix) consts ([], [], [], []) ([], []) end;  | 
| 384 | 187  | 
|
188  | 
||
189  | 
(* syn_ext_consts *)  | 
|
190  | 
||
| 
15751
 
65e4790c7914
identify binder translations only once (admits remove);
 
wenzelm 
parents: 
15570 
diff
changeset
 | 
191  | 
val binder_stamp = stamp ();  | 
| 
21534
 
68f805e9db0b
Binder: syntax const is determined by binder_name, not its syntax;
 
wenzelm 
parents: 
20892 
diff
changeset
 | 
192  | 
val binder_name = suffix "_binder";  | 
| 
15751
 
65e4790c7914
identify binder translations only once (admits remove);
 
wenzelm 
parents: 
15570 
diff
changeset
 | 
193  | 
|
| 67398 | 194  | 
fun mk_prefix sy =  | 
| 
69064
 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 
nipkow 
parents: 
68273 
diff
changeset
 | 
195  | 
let val sy' = Input.source_explode (Input.reset_pos sy);  | 
| 
 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 
nipkow 
parents: 
68273 
diff
changeset
 | 
196  | 
  in Symbol_Pos.explode0 "'(" @ sy' @ Symbol_Pos.explode0 "')" end
 | 
| 67398 | 197  | 
|
| 59841 | 198  | 
fun syn_ext_consts logical_types const_decls =  | 
| 384 | 199  | 
let  | 
| 62796 | 200  | 
fun mk_infix sy ty c p1 p2 p3 pos =  | 
| 67398 | 201  | 
[Syntax_Ext.Mfix (mk_prefix sy, ty, c, [], 1000, Position.none),  | 
| 62752 | 202  | 
Syntax_Ext.Mfix  | 
203  | 
(Symbol_Pos.explode0 "(_ " @ Input.source_explode sy @ Symbol_Pos.explode0 "/ _)",  | 
|
| 62796 | 204  | 
ty, c, [p1, p2], p3, pos)];  | 
| 384 | 205  | 
|
206  | 
    fun binder_typ _ (Type ("fun", [Type ("fun", [_, ty2]), ty3])) =
 | 
|
207  | 
          [Type ("idts", []), ty2] ---> ty3
 | 
|
| 14903 | 208  | 
      | binder_typ c _ = error ("Bad type of binder: " ^ quote c);
 | 
| 384 | 209  | 
|
| 62796 | 210  | 
fun mfix_of (mfix as (_, _, mx)) =  | 
211  | 
(case mfix of  | 
|
212  | 
(_, _, NoSyn) => []  | 
|
213  | 
| (c, ty, Mixfix (sy, ps, p, _)) =>  | 
|
214  | 
[Syntax_Ext.Mfix (Input.source_explode sy, ty, c, ps, p, pos_of mx)]  | 
|
215  | 
| (c, ty, Infix (sy, p, _)) => mk_infix sy ty c (p + 1) (p + 1) p (pos_of mx)  | 
|
216  | 
| (c, ty, Infixl (sy, p, _)) => mk_infix sy ty c p (p + 1) p (pos_of mx)  | 
|
217  | 
| (c, ty, Infixr (sy, p, _)) => mk_infix sy ty c (p + 1) p p (pos_of mx)  | 
|
218  | 
| (c, ty, Binder (sy, p, q, _)) =>  | 
|
| 62752 | 219  | 
[Syntax_Ext.Mfix  | 
220  | 
(Symbol_Pos.explode0 "(3" @ Input.source_explode sy @ Symbol_Pos.explode0 "_./ _)",  | 
|
| 62796 | 221  | 
binder_typ c ty, (binder_name c), [0, p], q, pos_of mx)]  | 
222  | 
| (c, _, mx) =>  | 
|
223  | 
          error ("Bad mixfix declaration for " ^ quote c ^ Position.here (pos_of mx)));
 | 
|
| 384 | 224  | 
|
| 
21534
 
68f805e9db0b
Binder: syntax const is determined by binder_name, not its syntax;
 
wenzelm 
parents: 
20892 
diff
changeset
 | 
225  | 
fun binder (c, _, Binder _) = SOME (binder_name c, c)  | 
| 15531 | 226  | 
| binder _ = NONE;  | 
| 384 | 227  | 
|
| 
69586
 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 
wenzelm 
parents: 
69584 
diff
changeset
 | 
228  | 
val _ = Position.reports_text (maps (mixfix_template_reports o #3) const_decls);  | 
| 
 
9171d1ce5a35
support for "isabelle update -u mixfix_cartouches";
 
wenzelm 
parents: 
69584 
diff
changeset
 | 
229  | 
|
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19467 
diff
changeset
 | 
230  | 
val mfix = maps mfix_of const_decls;  | 
| 
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19467 
diff
changeset
 | 
231  | 
val binders = map_filter binder const_decls;  | 
| 35129 | 232  | 
val binder_trs = binders  | 
| 52143 | 233  | 
|> map (Syntax_Ext.stamp_trfun binder_stamp o Syntax_Trans.mk_binder_tr);  | 
| 35129 | 234  | 
val binder_trs' = binders  | 
| 
42288
 
2074b31650e6
discontinued special treatment of structure Syntax_Ext (formerly Syn_Ext);
 
wenzelm 
parents: 
42287 
diff
changeset
 | 
235  | 
|> map (Syntax_Ext.stamp_trfun binder_stamp o  | 
| 52143 | 236  | 
apsnd Syntax_Trans.non_typed_tr' o Syntax_Trans.mk_binder_tr' o swap);  | 
| 
42298
 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 
wenzelm 
parents: 
42297 
diff
changeset
 | 
237  | 
|
| 
 
d622145603ee
more accurate markup for syntax consts, notably binders which point back to the original logical entity;
 
wenzelm 
parents: 
42297 
diff
changeset
 | 
238  | 
val consts = binders @ map (fn (c, _, _) => (c, "")) const_decls;  | 
| 69584 | 239  | 
in Syntax_Ext.syn_ext logical_types mfix consts ([], binder_trs, binder_trs', []) ([], []) end;  | 
| 384 | 240  | 
|
241  | 
end;  | 
|
| 
42287
 
d98eb048a2e4
discontinued special treatment of structure Mixfix;
 
wenzelm 
parents: 
42284 
diff
changeset
 | 
242  | 
|
| 
 
d98eb048a2e4
discontinued special treatment of structure Mixfix;
 
wenzelm 
parents: 
42284 
diff
changeset
 | 
243  | 
structure Basic_Mixfix: BASIC_MIXFIX = Mixfix;  | 
| 
 
d98eb048a2e4
discontinued special treatment of structure Mixfix;
 
wenzelm 
parents: 
42284 
diff
changeset
 | 
244  | 
open Basic_Mixfix;  |