| author | wenzelm | 
| Thu, 07 Apr 2011 18:24:59 +0200 | |
| changeset 42268 | 01401287c3f7 | 
| parent 42130 | e10e2cce85c8 | 
| child 42284 | 326f57825e1a | 
| 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  | 
||
7  | 
signature MIXFIX0 =  | 
|
| 
2199
 
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
 
wenzelm 
parents: 
1952 
diff
changeset
 | 
8  | 
sig  | 
| 384 | 9  | 
datatype mixfix =  | 
10  | 
NoSyn |  | 
|
11  | 
Mixfix of string * int list * int |  | 
|
12  | 
Delimfix of string |  | 
|
| 35130 | 13  | 
Infix of string * int |  | 
14  | 
Infixl of string * int |  | 
|
15  | 
Infixr of string * int |  | 
|
| 18673 | 16  | 
Binder of string * int * int |  | 
17  | 
Structure  | 
|
| 
21534
 
68f805e9db0b
Binder: syntax const is determined by binder_name, not its syntax;
 
wenzelm 
parents: 
20892 
diff
changeset
 | 
18  | 
val binder_name: string -> string  | 
| 
2199
 
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
 
wenzelm 
parents: 
1952 
diff
changeset
 | 
19  | 
end;  | 
| 384 | 20  | 
|
21  | 
signature MIXFIX1 =  | 
|
| 
2199
 
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
 
wenzelm 
parents: 
1952 
diff
changeset
 | 
22  | 
sig  | 
| 384 | 23  | 
include MIXFIX0  | 
| 4823 | 24  | 
val no_syn: 'a * 'b -> 'a * 'b * mixfix  | 
| 19271 | 25  | 
val pretty_mixfix: mixfix -> Pretty.T  | 
| 4823 | 26  | 
val mixfix_args: mixfix -> int  | 
| 22702 | 27  | 
val mixfixT: mixfix -> typ  | 
| 
35412
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
35390 
diff
changeset
 | 
28  | 
val make_type: int -> typ  | 
| 
2199
 
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
 
wenzelm 
parents: 
1952 
diff
changeset
 | 
29  | 
end;  | 
| 384 | 30  | 
|
31  | 
signature MIXFIX =  | 
|
| 
2199
 
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
 
wenzelm 
parents: 
1952 
diff
changeset
 | 
32  | 
sig  | 
| 384 | 33  | 
include MIXFIX1  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
34  | 
val syn_ext_types: (string * typ * mixfix) list -> Syn_Ext.syn_ext  | 
| 
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
35  | 
val syn_ext_consts: (string -> bool) -> (string * typ * mixfix) list -> Syn_Ext.syn_ext  | 
| 
2199
 
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
 
wenzelm 
parents: 
1952 
diff
changeset
 | 
36  | 
end;  | 
| 
 
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
 
wenzelm 
parents: 
1952 
diff
changeset
 | 
37  | 
|
| 
15751
 
65e4790c7914
identify binder translations only once (admits remove);
 
wenzelm 
parents: 
15570 
diff
changeset
 | 
38  | 
structure Mixfix: MIXFIX =  | 
| 384 | 39  | 
struct  | 
40  | 
||
| 
2199
 
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
 
wenzelm 
parents: 
1952 
diff
changeset
 | 
41  | 
|
| 384 | 42  | 
(** mixfix declarations **)  | 
43  | 
||
44  | 
datatype mixfix =  | 
|
45  | 
NoSyn |  | 
|
46  | 
Mixfix of string * int list * int |  | 
|
47  | 
Delimfix of string |  | 
|
| 35130 | 48  | 
Infix of string * int |  | 
49  | 
Infixl of string * int |  | 
|
50  | 
Infixr of string * int |  | 
|
| 18673 | 51  | 
Binder of string * int * int |  | 
52  | 
Structure;  | 
|
| 384 | 53  | 
|
| 4823 | 54  | 
fun no_syn (x, y) = (x, y, NoSyn);  | 
55  | 
||
| 384 | 56  | 
|
| 19271 | 57  | 
(* pretty_mixfix *)  | 
58  | 
||
59  | 
local  | 
|
| 11920 | 60  | 
|
| 19271 | 61  | 
val quoted = Pretty.quote o Pretty.str;  | 
62  | 
val keyword = Pretty.keyword;  | 
|
63  | 
val parens = Pretty.enclose "(" ")";
 | 
|
64  | 
val brackets = Pretty.enclose "[" "]";  | 
|
65  | 
val int = Pretty.str o string_of_int;  | 
|
66  | 
||
67  | 
in  | 
|
| 11920 | 68  | 
|
| 19271 | 69  | 
fun pretty_mixfix NoSyn = Pretty.str ""  | 
70  | 
| pretty_mixfix (Mixfix (s, ps, p)) =  | 
|
71  | 
parens (Pretty.breaks [quoted s, brackets (Pretty.commas (map int ps)), int p])  | 
|
72  | 
| pretty_mixfix (Delimfix s) = parens [quoted s]  | 
|
| 35130 | 73  | 
| pretty_mixfix (Infix (s, p)) = parens (Pretty.breaks [keyword "infix", quoted s, int p])  | 
74  | 
| pretty_mixfix (Infixl (s, p)) = parens (Pretty.breaks [keyword "infixl", quoted s, int p])  | 
|
75  | 
| pretty_mixfix (Infixr (s, p)) = parens (Pretty.breaks [keyword "infixl", quoted s, int p])  | 
|
| 19271 | 76  | 
| pretty_mixfix (Binder (s, p1, p2)) =  | 
77  | 
parens (Pretty.breaks [keyword "binder", quoted s, brackets [int p1], int p2])  | 
|
78  | 
| pretty_mixfix Structure = parens [keyword "structure"];  | 
|
79  | 
||
80  | 
end;  | 
|
| 11920 | 81  | 
|
82  | 
||
| 12531 | 83  | 
(* syntax specifications *)  | 
| 384 | 84  | 
|
| 4053 | 85  | 
fun mixfix_args NoSyn = 0  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
86  | 
| mixfix_args (Mixfix (sy, _, _)) = Syn_Ext.mfix_args sy  | 
| 
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
87  | 
| mixfix_args (Delimfix sy) = Syn_Ext.mfix_args sy  | 
| 
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
88  | 
| mixfix_args (Infix (sy, _)) = 2 + Syn_Ext.mfix_args sy  | 
| 
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
89  | 
| mixfix_args (Infixl (sy, _)) = 2 + Syn_Ext.mfix_args sy  | 
| 
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
90  | 
| mixfix_args (Infixr (sy, _)) = 2 + Syn_Ext.mfix_args sy  | 
| 18673 | 91  | 
| mixfix_args (Binder _) = 1  | 
92  | 
| mixfix_args Structure = 0;  | 
|
| 4053 | 93  | 
|
| 22709 | 94  | 
fun mixfixT (Binder _) = (dummyT --> dummyT) --> dummyT  | 
95  | 
| mixfixT mx = replicate (mixfix_args mx) dummyT ---> dummyT;  | 
|
| 22702 | 96  | 
|
| 4053 | 97  | 
|
| 384 | 98  | 
(* syn_ext_types *)  | 
99  | 
||
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
100  | 
fun make_type n = replicate n Syn_Ext.typeT ---> Syn_Ext.typeT;  | 
| 
35412
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
35390 
diff
changeset
 | 
101  | 
|
| 14903 | 102  | 
fun syn_ext_types type_decls =  | 
| 384 | 103  | 
let  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
104  | 
    fun mk_infix sy ty t p1 p2 p3 = Syn_Ext.Mfix ("(_ " ^ sy ^ "/ _)", ty, t, [p1, p2], p3);
 | 
| 384 | 105  | 
|
| 35129 | 106  | 
fun mfix_of (_, _, NoSyn) = NONE  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
107  | 
| mfix_of (t, ty, Mixfix (sy, ps, p)) = SOME (Syn_Ext.Mfix (sy, ty, t, ps, p))  | 
| 
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
108  | 
| mfix_of (t, ty, Delimfix sy) = SOME (Syn_Ext.Mfix (sy, ty, t, [], Syn_Ext.max_pri))  | 
| 
35412
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
35390 
diff
changeset
 | 
109  | 
| mfix_of (t, ty, Infix (sy, p)) = SOME (mk_infix sy ty t (p + 1) (p + 1) p)  | 
| 
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
35390 
diff
changeset
 | 
110  | 
| mfix_of (t, ty, Infixl (sy, p)) = SOME (mk_infix sy ty t p (p + 1) p)  | 
| 
 
b8dead547d9e
more uniform treatment of syntax for types vs. consts;
 
wenzelm 
parents: 
35390 
diff
changeset
 | 
111  | 
| mfix_of (t, ty, Infixr (sy, p)) = SOME (mk_infix sy ty t (p + 1) p p)  | 
| 
42130
 
e10e2cce85c8
removed unclear comments stemming from ed24ba6f69aa;
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
112  | 
      | mfix_of (t, _, _) = error ("Bad mixfix declaration for " ^ quote t);
 | 
| 384 | 113  | 
|
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
114  | 
fun check_args (_, ty, _) (SOME (mfix as Syn_Ext.Mfix (sy, _, _, _, _))) =  | 
| 
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
115  | 
if length (Term.binder_types ty) = Syn_Ext.mfix_args sy then ()  | 
| 
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
116  | 
else Syn_Ext.err_in_mfix "Bad number of type constructor arguments" mfix  | 
| 
35351
 
7425aece4ee3
allow general mixfix syntax for type constructors;
 
wenzelm 
parents: 
35130 
diff
changeset
 | 
117  | 
| check_args _ NONE = ();  | 
| 
 
7425aece4ee3
allow general mixfix syntax for type constructors;
 
wenzelm 
parents: 
35130 
diff
changeset
 | 
118  | 
|
| 
 
7425aece4ee3
allow general mixfix syntax for type constructors;
 
wenzelm 
parents: 
35130 
diff
changeset
 | 
119  | 
val mfix = map mfix_of type_decls;  | 
| 
 
7425aece4ee3
allow general mixfix syntax for type constructors;
 
wenzelm 
parents: 
35130 
diff
changeset
 | 
120  | 
val _ = map2 check_args type_decls mfix;  | 
| 35129 | 121  | 
val xconsts = map #1 type_decls;  | 
| 42268 | 122  | 
in Syn_Ext.syn_ext (map_filter I mfix) xconsts ([], [], [], []) ([], []) end;  | 
| 384 | 123  | 
|
124  | 
||
125  | 
(* syn_ext_consts *)  | 
|
126  | 
||
| 
15751
 
65e4790c7914
identify binder translations only once (admits remove);
 
wenzelm 
parents: 
15570 
diff
changeset
 | 
127  | 
val binder_stamp = stamp ();  | 
| 
21534
 
68f805e9db0b
Binder: syntax const is determined by binder_name, not its syntax;
 
wenzelm 
parents: 
20892 
diff
changeset
 | 
128  | 
val binder_name = suffix "_binder";  | 
| 
15751
 
65e4790c7914
identify binder translations only once (admits remove);
 
wenzelm 
parents: 
15570 
diff
changeset
 | 
129  | 
|
| 14903 | 130  | 
fun syn_ext_consts is_logtype const_decls =  | 
| 384 | 131  | 
let  | 
132  | 
fun mk_infix sy ty c p1 p2 p3 =  | 
|
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
133  | 
      [Syn_Ext.Mfix ("op " ^ sy, ty, c, [], Syn_Ext.max_pri),
 | 
| 
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
134  | 
       Syn_Ext.Mfix ("(_ " ^ sy ^ "/ _)", ty, c, [p1, p2], p3)];
 | 
| 384 | 135  | 
|
136  | 
    fun binder_typ _ (Type ("fun", [Type ("fun", [_, ty2]), ty3])) =
 | 
|
137  | 
          [Type ("idts", []), ty2] ---> ty3
 | 
|
| 14903 | 138  | 
      | binder_typ c _ = error ("Bad type of binder: " ^ quote c);
 | 
| 384 | 139  | 
|
| 35129 | 140  | 
fun mfix_of (_, _, NoSyn) = []  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
141  | 
| mfix_of (c, ty, Mixfix (sy, ps, p)) = [Syn_Ext.Mfix (sy, ty, c, ps, p)]  | 
| 
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
142  | 
| mfix_of (c, ty, Delimfix sy) = [Syn_Ext.Mfix (sy, ty, c, [], Syn_Ext.max_pri)]  | 
| 35130 | 143  | 
| mfix_of (c, ty, Infix (sy, p)) = mk_infix sy ty c (p + 1) (p + 1) p  | 
144  | 
| mfix_of (c, ty, Infixl (sy, p)) = mk_infix sy ty c p (p + 1) p  | 
|
145  | 
| mfix_of (c, ty, Infixr (sy, p)) = mk_infix sy ty c (p + 1) p p  | 
|
| 35129 | 146  | 
| mfix_of (c, ty, Binder (sy, p, q)) =  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
147  | 
          [Syn_Ext.Mfix ("(3" ^ sy ^ "_./ _)", binder_typ c ty, (binder_name c), [0, p], q)]
 | 
| 
42130
 
e10e2cce85c8
removed unclear comments stemming from ed24ba6f69aa;
 
wenzelm 
parents: 
37216 
diff
changeset
 | 
148  | 
      | mfix_of (c, _, _) = error ("Bad mixfix declaration for " ^ quote c);
 | 
| 384 | 149  | 
|
| 
21534
 
68f805e9db0b
Binder: syntax const is determined by binder_name, not its syntax;
 
wenzelm 
parents: 
20892 
diff
changeset
 | 
150  | 
fun binder (c, _, Binder _) = SOME (binder_name c, c)  | 
| 15531 | 151  | 
| binder _ = NONE;  | 
| 384 | 152  | 
|
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19467 
diff
changeset
 | 
153  | 
val mfix = maps mfix_of const_decls;  | 
| 35129 | 154  | 
val xconsts = map #1 const_decls;  | 
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19467 
diff
changeset
 | 
155  | 
val binders = map_filter binder const_decls;  | 
| 35129 | 156  | 
val binder_trs = binders  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
157  | 
|> map (Syn_Ext.stamp_trfun binder_stamp o apsnd K o Syn_Trans.mk_binder_tr);  | 
| 35129 | 158  | 
val binder_trs' = binders  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
159  | 
|> map (Syn_Ext.stamp_trfun binder_stamp o  | 
| 
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
160  | 
apsnd (K o Syn_Trans.non_typed_tr') o Syn_Trans.mk_binder_tr' o swap);  | 
| 14903 | 161  | 
in  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
35412 
diff
changeset
 | 
162  | 
Syn_Ext.syn_ext' true is_logtype  | 
| 42268 | 163  | 
mfix xconsts ([], binder_trs, binder_trs', []) ([], [])  | 
| 14903 | 164  | 
end;  | 
| 384 | 165  | 
|
166  | 
end;  |