author | wenzelm |
Fri, 10 Oct 1997 15:47:41 +0200 | |
changeset 3829 | d7333ef9e72c |
parent 3815 | 7e8847f8f3a4 |
child 4053 | c88d0d5ae806 |
permissions | -rw-r--r-- |
384 | 1 |
(* Title: Pure/Syntax/mixfix.ML |
2 |
ID: $Id$ |
|
551 | 3 |
Author: Tobias Nipkow and Markus Wenzel, TU Muenchen |
384 | 4 |
|
2199
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
5 |
Mixfix declarations, infixes, binders. Also parts of the Pure syntax. |
384 | 6 |
*) |
7 |
||
8 |
signature MIXFIX0 = |
|
2199
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
9 |
sig |
384 | 10 |
datatype mixfix = |
11 |
NoSyn | |
|
12 |
Mixfix of string * int list * int | |
|
13 |
Delimfix of string | |
|
2199
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
14 |
InfixlName of string * int | |
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
15 |
InfixrName of string * int | |
384 | 16 |
Infixl of int | |
17 |
Infixr of int | |
|
865 | 18 |
Binder of string * int * int |
384 | 19 |
val max_pri: int |
2199
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
20 |
end; |
384 | 21 |
|
22 |
signature MIXFIX1 = |
|
2199
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
23 |
sig |
384 | 24 |
include MIXFIX0 |
25 |
val type_name: string -> mixfix -> string |
|
26 |
val const_name: string -> mixfix -> string |
|
27 |
val pure_types: (string * int * mixfix) list |
|
28 |
val pure_syntax: (string * string * mixfix) list |
|
1181 | 29 |
val pure_appl_syntax: (string * string * mixfix) list |
30 |
val pure_applC_syntax: (string * string * mixfix) list |
|
2256 | 31 |
val pure_sym_syntax: (string * string * mixfix) list |
2199
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
32 |
end; |
384 | 33 |
|
34 |
signature MIXFIX = |
|
2199
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
35 |
sig |
384 | 36 |
include MIXFIX1 |
2199
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
37 |
val syn_ext_types: string list -> (string * int * mixfix) list |
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
38 |
-> SynExt.syn_ext |
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
39 |
val syn_ext_consts: string list -> (string * typ * mixfix) list |
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
40 |
-> SynExt.syn_ext |
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
41 |
end; |
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
42 |
|
384 | 43 |
|
1507 | 44 |
structure Mixfix : MIXFIX = |
384 | 45 |
struct |
46 |
||
1507 | 47 |
open Lexicon SynExt Ast SynTrans; |
384 | 48 |
|
2199
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
49 |
|
384 | 50 |
(** mixfix declarations **) |
51 |
||
52 |
datatype mixfix = |
|
53 |
NoSyn | |
|
54 |
Mixfix of string * int list * int | |
|
55 |
Delimfix of string | |
|
2199
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
56 |
InfixlName of string * int | |
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
57 |
InfixrName of string * int | |
384 | 58 |
Infixl of int | |
59 |
Infixr of int | |
|
887 | 60 |
Binder of string * int * int; |
384 | 61 |
|
62 |
||
63 |
(* type / const names *) |
|
64 |
||
65 |
fun strip ("'" :: c :: cs) = c :: strip cs |
|
66 |
| strip ["'"] = [] |
|
67 |
| strip (c :: cs) = c :: strip cs |
|
68 |
| strip [] = []; |
|
69 |
||
70 |
val strip_esc = implode o strip o explode; |
|
71 |
||
2199
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
72 |
fun type_name t (InfixlName _) = t |
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
73 |
| type_name t (InfixrName _) = t |
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
74 |
| type_name t (Infixl _) = strip_esc t |
384 | 75 |
| type_name t (Infixr _) = strip_esc t |
76 |
| type_name t _ = t; |
|
77 |
||
2199
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
78 |
fun const_name c (InfixlName _) = c |
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
79 |
| const_name c (InfixrName _) = c |
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
80 |
| const_name c (Infixl _) = "op " ^ strip_esc c |
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
81 |
| const_name c (Infixr _) = "op " ^ strip_esc c |
384 | 82 |
| const_name c _ = c; |
83 |
||
84 |
||
85 |
(* syn_ext_types *) |
|
86 |
||
764
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
clasohm
parents:
639
diff
changeset
|
87 |
fun syn_ext_types logtypes type_decls = |
384 | 88 |
let |
89 |
fun name_of (t, _, mx) = type_name t mx; |
|
90 |
||
2199
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
91 |
fun mk_infix sy t p1 p2 p3 = |
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
92 |
Mfix ("(_ " ^ sy ^ "/ _)", [typeT, typeT] ---> typeT, t, [p1, p2], p3); |
384 | 93 |
|
2199
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
94 |
fun mfix_of decl = |
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
95 |
let val t = name_of decl in |
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
96 |
(case decl of |
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
97 |
(_, _, NoSyn) => None |
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
98 |
| (_, 2, InfixlName (sy, p)) => Some (mk_infix t sy p (p + 1) p) |
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
99 |
| (_, 2, InfixrName (sy, p)) => Some (mk_infix t sy (p + 1) p p) |
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
100 |
| (sy, 2, Infixl p) => Some (mk_infix t sy p (p + 1) p) |
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
101 |
| (sy, 2, Infixr p) => Some (mk_infix t sy (p + 1) p p) |
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
102 |
| _ => error ("Bad mixfix declaration for type " ^ quote t)) |
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
103 |
end; |
384 | 104 |
|
105 |
val mfix = mapfilter mfix_of type_decls; |
|
106 |
val xconsts = map name_of type_decls; |
|
107 |
in |
|
2696 | 108 |
syn_ext logtypes mfix xconsts ([], [], [], []) [] ([], []) |
384 | 109 |
end; |
110 |
||
111 |
||
112 |
(* syn_ext_consts *) |
|
113 |
||
764
b60e77395d1a
changed Pure's grammar and the way types are converted to nonterminals
clasohm
parents:
639
diff
changeset
|
114 |
fun syn_ext_consts logtypes const_decls = |
384 | 115 |
let |
116 |
fun name_of (c, _, mx) = const_name c mx; |
|
117 |
||
118 |
fun mk_infix sy ty c p1 p2 p3 = |
|
3815 | 119 |
[Mfix ("op " ^ sy, ty, c, [], max_pri), |
120 |
Mfix ("(_ " ^ sy ^ "/ _)", ty, c, [p1, p2], p3)]; |
|
384 | 121 |
|
122 |
fun binder_typ _ (Type ("fun", [Type ("fun", [_, ty2]), ty3])) = |
|
123 |
[Type ("idts", []), ty2] ---> ty3 |
|
124 |
| binder_typ c _ = error ("Bad type of binder " ^ quote c); |
|
125 |
||
2199
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
126 |
fun mfix_of decl = |
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
127 |
let val c = name_of decl in |
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
128 |
(case decl of |
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
129 |
(_, _, NoSyn) => [] |
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
130 |
| (_, ty, Mixfix (sy, ps, p)) => [Mfix (sy, ty, c, ps, p)] |
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
131 |
| (_, ty, Delimfix sy) => [Mfix (sy, ty, c, [], max_pri)] |
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
132 |
| (_, ty, InfixlName (sy, p)) => mk_infix sy ty c p (p + 1) p |
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
133 |
| (_, ty, InfixrName (sy, p)) => mk_infix sy ty c (p + 1) p p |
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
134 |
| (sy, ty, Infixl p) => mk_infix sy ty c p (p + 1) p |
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
135 |
| (sy, ty, Infixr p) => mk_infix sy ty c (p + 1) p p |
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
136 |
| (_, ty, Binder (sy, p, q)) => |
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
137 |
[Mfix ("(3" ^ sy ^ "_./ _)", binder_typ c ty, sy, [0, p], q)]) |
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
138 |
end; |
384 | 139 |
|
865 | 140 |
fun binder (c, _, Binder (sy, _, _)) = Some (sy, c) |
384 | 141 |
| binder _ = None; |
142 |
||
143 |
val mfix = flat (map mfix_of const_decls); |
|
144 |
val xconsts = map name_of const_decls; |
|
145 |
val binders = mapfilter binder const_decls; |
|
146 |
val binder_trs = map mk_binder_tr binders; |
|
2381 | 147 |
val binder_trs' = map (apsnd fix_tr' o mk_binder_tr' o swap) binders; |
384 | 148 |
in |
2696 | 149 |
syn_ext logtypes mfix xconsts ([], binder_trs, binder_trs', []) [] ([], []) |
384 | 150 |
end; |
151 |
||
152 |
||
153 |
||
154 |
(** Pure syntax **) |
|
155 |
||
156 |
val pure_types = |
|
157 |
map (fn t => (t, 0, NoSyn)) |
|
1178 | 158 |
(terminals @ [logic, "type", "types", "sort", "classes", args, cargs, |
3690 | 159 |
"pttrn", "pttrns", "idt", "idts", "aprop", "asms", any, sprop, "dummy"]); |
384 | 160 |
|
161 |
val pure_syntax = |
|
3690 | 162 |
[("_lambda", "[pttrns, 'a] => logic", Mixfix ("(3%_./ _)", [0, 3], 3)), |
2696 | 163 |
("_abs", "'a", NoSyn), |
164 |
("", "'a => " ^ args, Delimfix "_"), |
|
165 |
("_args", "['a, " ^ args ^ "] => " ^ args, Delimfix "_,/ _"), |
|
166 |
("", "id => idt", Delimfix "_"), |
|
167 |
("_idtyp", "[id, type] => idt", Mixfix ("_::_", [], 0)), |
|
168 |
("", "idt => idt", Delimfix "'(_')"), |
|
3690 | 169 |
("", "idt => idts", Delimfix "_"), |
170 |
("_idts", "[idt, idts] => idts", Mixfix ("_/ _", [1, 0], 0)), |
|
2696 | 171 |
("", "idt => pttrn", Delimfix "_"), |
3690 | 172 |
("", "pttrn => pttrns", Delimfix "_"), |
173 |
("_pttrns", "[pttrn, pttrns] => pttrns", Mixfix ("_/ _", [1, 0], 0)), |
|
2696 | 174 |
("", "id => aprop", Delimfix "_"), |
3829 | 175 |
("", "longid => aprop", Delimfix "_"), |
2696 | 176 |
("", "var => aprop", Delimfix "_"), |
177 |
("_aprop", "aprop => prop", Delimfix "PROP _"), |
|
178 |
("", "prop => asms", Delimfix "_"), |
|
179 |
("_asms", "[prop, asms] => asms", Delimfix "_;/ _"), |
|
180 |
("_bigimpl", "[asms, prop] => prop", Mixfix ("((3[| _ |])/ ==> _)", [0, 1], 1)), |
|
181 |
("_ofclass", "[type, logic] => prop", Delimfix "(1OFCLASS/(1'(_,/ _')))"), |
|
182 |
("_mk_ofclass", "_", NoSyn), |
|
183 |
("_mk_ofclassS", "_", NoSyn), |
|
184 |
("_K", "_", NoSyn), |
|
185 |
("", "id => logic", Delimfix "_"), |
|
3829 | 186 |
("", "longid => logic", Delimfix "_"), |
2696 | 187 |
("", "var => logic", Delimfix "_")]; |
384 | 188 |
|
1181 | 189 |
val pure_appl_syntax = |
2696 | 190 |
[("_appl", "[('b => 'a), args] => logic", Mixfix ("(1_/(1'(_')))", [max_pri, 0], max_pri)), |
191 |
("_appl", "[('b => 'a), args] => aprop", Mixfix ("(1_/(1'(_')))", [max_pri, 0], max_pri))]; |
|
1181 | 192 |
|
193 |
val pure_applC_syntax = |
|
2696 | 194 |
[("", "'a => cargs", Delimfix "_"), |
195 |
("_cargs", "['a, cargs] => cargs", Mixfix ("_/ _", [max_pri, max_pri], max_pri)), |
|
196 |
("_applC", "[('b => 'a), cargs] => logic", Mixfix ("(1_/ _)", [max_pri, max_pri], max_pri - 1)), |
|
197 |
("_applC", "[('b => 'a), cargs] => aprop", Mixfix ("(1_/ _)", [max_pri, max_pri], max_pri - 1))]; |
|
2199
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
198 |
|
2256 | 199 |
val pure_sym_syntax = |
2696 | 200 |
[("fun", "[type, type] => type", Mixfix ("(_/ \\<Rightarrow> _)", [1, 0], 0)), |
201 |
("_bracket", "[types, type] => type", Mixfix ("([_]/ \\<Rightarrow> _)", [0, 0], 0)), |
|
2767
e1d15eabb64d
added \<Colon> syntax for constraints (more compact!);
wenzelm
parents:
2696
diff
changeset
|
202 |
("_ofsort", "[tid, sort] => type", Mixfix ("_\\<Colon>_", [max_pri, 0], max_pri)), |
e1d15eabb64d
added \<Colon> syntax for constraints (more compact!);
wenzelm
parents:
2696
diff
changeset
|
203 |
("_constrain", "['a, type] => 'a", Mixfix ("_\\<Colon>_", [4, 0], 3)), |
e1d15eabb64d
added \<Colon> syntax for constraints (more compact!);
wenzelm
parents:
2696
diff
changeset
|
204 |
("_idtyp", "[id, type] => idt", Mixfix ("_\\<Colon>_", [], 0)), |
3690 | 205 |
("_lambda", "[pttrns, 'a] => logic", Mixfix ("(3\\<lambda>_./ _)", [0, 3], 3)), |
2696 | 206 |
("==>", "[prop, prop] => prop", InfixrName ("\\<Midarrow>\\<Rightarrow>", 1)), |
207 |
("_bigimpl", "[asms, prop] => prop", Mixfix ("((3\\<lbrakk>_\\<rbrakk>)/ \\<Midarrow>\\<Rightarrow> _)", [0, 1], 1)), |
|
208 |
("==", "['a::{}, 'a] => prop", InfixrName ("\\<equiv>", 2)), |
|
209 |
("!!", "[idts, prop] => prop", Mixfix ("(3\\<And>_./ _)", [0, 0], 0))]; |
|
2199
bcb360f80dac
added Infixl/rName: specify infix name independently from syntax;
wenzelm
parents:
1952
diff
changeset
|
210 |
|
3829 | 211 |
|
384 | 212 |
end; |