1274
|
1 |
(* TTITLEF/thy_ops.ML
|
|
2 |
ID: $Id$
|
|
3 |
Author: Tobias Mayr
|
|
4 |
|
|
5 |
Additional theory file section for HOLCF: ops
|
|
6 |
There's an elaborate but german description of this program,
|
|
7 |
write to mayrt@informatik.tu-muenchen.de.
|
|
8 |
For a short english description of the new sections
|
|
9 |
write to regensbu@informatik.tu-muenchen.de.
|
|
10 |
|
|
11 |
TODO: vielleicht AST-Darstellung mit "op name" statt _I_...
|
|
12 |
|
|
13 |
*)
|
|
14 |
|
|
15 |
signature THY_OPS =
|
|
16 |
sig
|
|
17 |
(* continuous mixfixes (extension of datatype PrivateSyntax.Mixfix.mixfix) *)
|
|
18 |
datatype cmixfix =
|
|
19 |
Mixfix of PrivateSyntax.Mixfix.mixfix |
|
|
20 |
CInfixl of int |
|
|
21 |
CInfixr of int |
|
|
22 |
CMixfix of string * int list *int;
|
|
23 |
|
|
24 |
exception CINFIX of cmixfix;
|
|
25 |
val cmixfix_to_mixfix : cmixfix -> PrivateSyntax.Mixfix.mixfix;
|
|
26 |
|
|
27 |
(* theory extenders : *)
|
|
28 |
val add_ops : {curried: bool, total: bool, strict: bool} ->
|
|
29 |
(string * string * cmixfix) list -> theory -> theory;
|
|
30 |
val add_ops_i : {curried: bool, total: bool, strict: bool} ->
|
|
31 |
(string * typ * cmixfix) list -> theory -> theory;
|
|
32 |
val ops_keywords : string list;
|
|
33 |
val ops_sections : (string * (ThyParse.token list ->
|
|
34 |
(string * string) * ThyParse.token list)) list;
|
|
35 |
val opt_cmixfix: ThyParse.token list -> (string * ThyParse.token list);
|
|
36 |
val const_name : string -> cmixfix -> string;
|
|
37 |
end;
|
|
38 |
|
|
39 |
structure ThyOps : THY_OPS =
|
|
40 |
struct
|
|
41 |
|
|
42 |
open HOLCFlogic;
|
|
43 |
|
|
44 |
(** library ******************************************************)
|
|
45 |
|
|
46 |
(* abbreviations *)
|
|
47 |
val internal = fst; (* cinfix-ops will have diffrent internal/external names *)
|
|
48 |
val external = snd;
|
|
49 |
|
|
50 |
fun apsnd_of_three f = fn (a,b,c) => (a,f b,c);
|
|
51 |
|
|
52 |
|
|
53 |
(******** ops ********************)
|
|
54 |
|
|
55 |
(* the extended copy of mixfix *)
|
|
56 |
datatype cmixfix =
|
|
57 |
Mixfix of PrivateSyntax.Mixfix.mixfix |
|
|
58 |
CInfixl of int |
|
|
59 |
CInfixr of int |
|
|
60 |
CMixfix of string * int list *int;
|
|
61 |
|
|
62 |
exception CINFIX of cmixfix;
|
|
63 |
|
|
64 |
fun cmixfix_to_mixfix (Mixfix x) = x
|
|
65 |
| cmixfix_to_mixfix x = raise CINFIX x;
|
|
66 |
|
|
67 |
|
|
68 |
(** theory extender : add_ops *)
|
|
69 |
|
|
70 |
(* generating the declarations of the new constants. *************
|
|
71 |
cinfix names x are internally non infix (renamed by mk_internal_name)
|
|
72 |
and externally non continous infix function names (changed by op_to_fun).
|
|
73 |
Thus the cinfix declaration is splitted in an 'oldstyle' decl,
|
|
74 |
which is NoSyn (non infix) and is added by add_consts_i,
|
|
75 |
and an syn(tactic) decl, which is an infix function (not operation)
|
|
76 |
added by add_syntax_i, so that it can appear in input strings, but
|
|
77 |
not in terms.
|
|
78 |
The interface between internal and external names is realized by
|
|
79 |
transrules A x B <=> _x ' A ' B (generated by xrules_of)
|
|
80 |
The boolean argument 'curried' distinguishes between curried and
|
|
81 |
tupeled syntax of operation application *)
|
|
82 |
|
|
83 |
local
|
|
84 |
fun strip ("'" :: c :: cs) = c :: strip cs
|
|
85 |
| strip ["'"] = []
|
|
86 |
| strip (c :: cs) = c :: strip cs
|
|
87 |
| strip [] = [];
|
|
88 |
|
|
89 |
val strip_esc = implode o strip o explode;
|
|
90 |
|
|
91 |
fun infix_name c = "op " ^ strip_esc c;
|
|
92 |
in
|
|
93 |
val mk_internal_name = infix_name;
|
|
94 |
(*
|
|
95 |
(* changing e.g. 'ab' to '_I_97_98'.
|
|
96 |
Called by oldstyle, xrules_of, strictness_axms and totality_axms. *)
|
|
97 |
fun mk_internal_name name =
|
|
98 |
let fun alphanum (s::ss) = "_"^(string_of_int (ord s))^(alphanum ss)
|
|
99 |
| alphanum [] = "";
|
|
100 |
in
|
|
101 |
"_I"^(alphanum o explode) name
|
|
102 |
end;
|
|
103 |
*)
|
|
104 |
(* extension of Pure/Syntax/mixfix.ML: SynExt.const_name *)
|
|
105 |
fun const_name c (CInfixl _) = mk_internal_name c
|
|
106 |
| const_name c (CInfixr _) = mk_internal_name c
|
|
107 |
| const_name c (CMixfix _) = c
|
|
108 |
| const_name c (Mixfix x) = Syntax.const_name c x;
|
|
109 |
end;
|
|
110 |
|
|
111 |
(* Changing a->b->c res. a*b->c to a=>b=>c. Called by syn_decls. *)
|
|
112 |
(*####*)
|
|
113 |
fun op_to_fun true sign (Type("->" ,[larg,t]))=
|
|
114 |
Type("fun",[larg,op_to_fun true sign t])
|
|
115 |
| op_to_fun false sign (Type("->",[args,res])) = let
|
|
116 |
fun otf (Type("*",[larg,rargs])) = Type("fun",[larg,otf rargs])
|
|
117 |
| otf t = Type("fun",[t,res]);
|
|
118 |
in otf args end
|
|
119 |
| op_to_fun _ sign t = t(*error("Wrong type for cinfix/cmixfix : "^
|
|
120 |
(Sign.string_of_typ sign t))*);
|
|
121 |
(*####*)
|
|
122 |
|
|
123 |
(* oldstyle is called by add_ext_axioms(_i) *)
|
|
124 |
(* the first part is just copying the homomorphic part of the structures *)
|
|
125 |
fun oldstyle ((name,typ,Mixfix(x))::tl) =
|
|
126 |
(name,typ,x)::(oldstyle tl)
|
|
127 |
| oldstyle ((name,typ,CInfixl(i))::tl) =
|
|
128 |
(mk_internal_name name,typ,PrivateSyntax.Mixfix.NoSyn)::
|
|
129 |
(oldstyle tl)
|
|
130 |
| oldstyle ((name,typ,CInfixr(i))::tl) =
|
|
131 |
(mk_internal_name name,typ,PrivateSyntax.Mixfix.NoSyn)::
|
|
132 |
(oldstyle tl)
|
|
133 |
| oldstyle ((name,typ,CMixfix(x))::tl) =
|
|
134 |
(name,typ,PrivateSyntax.Mixfix.NoSyn)::
|
|
135 |
(oldstyle tl)
|
|
136 |
| oldstyle [] = [];
|
|
137 |
|
|
138 |
(* generating the external purely syntactical infix functions.
|
|
139 |
Called by add_ext_axioms(_i) *)
|
|
140 |
fun syn_decls curried sign ((name,typ,CInfixl(i))::tl) =
|
|
141 |
(name,op_to_fun curried sign typ,PrivateSyntax.Mixfix.Infixl(i))::
|
|
142 |
(syn_decls curried sign tl)
|
|
143 |
| syn_decls curried sign ((name,typ,CInfixr(i))::tl) =
|
|
144 |
(name,op_to_fun curried sign typ,PrivateSyntax.Mixfix.Infixr(i))::
|
|
145 |
(syn_decls curried sign tl)
|
|
146 |
| syn_decls curried sign ((name,typ,CMixfix(x))::tl) =
|
|
147 |
(*####
|
|
148 |
("@"^name,op_to_fun curried sign typ,PrivateSyntax.Mixfix.Mixfix(x))::
|
|
149 |
####**)
|
|
150 |
(name,op_to_fun curried sign typ,PrivateSyntax.Mixfix.Mixfix(x))::
|
|
151 |
|
|
152 |
(syn_decls curried sign tl)
|
|
153 |
| syn_decls curried sign (_::tl) = syn_decls curried sign tl
|
|
154 |
| syn_decls _ _ [] = [];
|
|
155 |
|
|
156 |
(* generating the translation rules. Called by add_ext_axioms(_i) *)
|
|
157 |
local open PrivateSyntax.Ast in
|
|
158 |
fun xrules_of true ((name,typ,CInfixl(i))::tail) =
|
|
159 |
((mk_appl (Constant (mk_internal_name name)) [Variable "A",Variable "B"]) <->
|
|
160 |
(mk_appl (Constant "@fapp") [(mk_appl (Constant "@fapp") [
|
|
161 |
Constant (mk_internal_name name),Variable "A"]),Variable "B"]))
|
|
162 |
::xrules_of true tail
|
|
163 |
| xrules_of true ((name,typ,CInfixr(i))::tail) =
|
|
164 |
((mk_appl (Constant (mk_internal_name name)) [Variable "A",Variable "B"]) <->
|
|
165 |
(mk_appl (Constant "@fapp") [(mk_appl (Constant "@fapp") [
|
|
166 |
Constant (mk_internal_name name),Variable "A"]),Variable "B"]))
|
|
167 |
::xrules_of true tail
|
|
168 |
(*####*)
|
|
169 |
| xrules_of true ((name,typ,CMixfix(_))::tail) = let
|
|
170 |
fun argnames n (Type("->" ,[_,t]))= chr n :: argnames (n+1) t
|
|
171 |
| argnames _ _ = [];
|
|
172 |
val names = argnames (ord"A") typ;
|
|
173 |
in if names = [] then [] else [mk_appl (Constant name) (map Variable names)<->
|
|
174 |
foldl (fn (t,arg) => (mk_appl (Constant "@fapp") [t,Variable arg]))
|
|
175 |
(Constant name,names)] end
|
|
176 |
@xrules_of true tail
|
|
177 |
(*####*)
|
|
178 |
| xrules_of false ((name,typ,CInfixl(i))::tail) =
|
|
179 |
((mk_appl (Constant (mk_internal_name name)) [Variable "A",Variable "B"]) <->
|
|
180 |
(mk_appl (Constant "@fapp") [ Constant(mk_internal_name name),
|
|
181 |
(mk_appl (Constant "@ctuple") [Variable "A",Variable "B"])]))
|
|
182 |
::xrules_of false tail
|
|
183 |
| xrules_of false ((name,typ,CInfixr(i))::tail) =
|
|
184 |
((mk_appl (Constant (mk_internal_name name)) [Variable "A",Variable "B"]) <->
|
|
185 |
(mk_appl (Constant "@fapp") [ Constant(mk_internal_name name),
|
|
186 |
(mk_appl (Constant "@ctuple") [Variable "A",Variable "B"])]))
|
|
187 |
::xrules_of false tail
|
|
188 |
(*####*)
|
|
189 |
| xrules_of false ((name,typ,CMixfix(_))::tail) = let
|
|
190 |
fun foldr' f l =
|
|
191 |
let fun itr [] = raise LIST "foldr'"
|
|
192 |
| itr [a] = a
|
|
193 |
| itr (a::l) = f(a, itr l)
|
|
194 |
in itr l end;
|
|
195 |
fun argnames n (Type("*" ,[_,t]))= chr n :: argnames (n+1) t
|
|
196 |
| argnames n _ = [chr n];
|
|
197 |
val vars = map Variable (case typ of (Type("->" ,[t,_])) =>argnames (ord"A") t
|
|
198 |
| _ => []);
|
|
199 |
in if vars = [] then [] else [mk_appl (Constant name) vars <->
|
|
200 |
(mk_appl (Constant "@fapp") [Constant name, case vars of [v] => v
|
|
201 |
| args => mk_appl (Constant "@ctuple") [hd args,foldr' (fn (t,arg) =>
|
|
202 |
mk_appl (Constant "_args") [t,arg]) (tl args)]])]
|
|
203 |
end
|
|
204 |
@xrules_of false tail
|
|
205 |
(*####*)
|
|
206 |
| xrules_of c ((name,typ,Mixfix(_))::tail) = xrules_of c tail
|
|
207 |
| xrules_of _ [] = [];
|
|
208 |
end;
|
|
209 |
(**** producing the new axioms ****************)
|
|
210 |
|
|
211 |
datatype arguments = Curried_args of ((typ*typ) list) |
|
|
212 |
Tupeled_args of (typ list);
|
|
213 |
|
|
214 |
fun num_of_args (Curried_args l) = length l
|
|
215 |
| num_of_args (Tupeled_args l) = length l;
|
|
216 |
|
|
217 |
fun types_of (Curried_args l) = map fst l
|
|
218 |
| types_of (Tupeled_args l) = l;
|
|
219 |
|
|
220 |
fun mk_mkNotEqUU_vars (typ::tl) cnt = mkNotEqUU (Free("x"^(string_of_int cnt),typ))::
|
|
221 |
(mk_mkNotEqUU_vars tl (cnt+1))
|
|
222 |
| mk_mkNotEqUU_vars [] _ = [];
|
|
223 |
|
|
224 |
local
|
|
225 |
(* T1*...*Tn goes to [T1,...,Tn] *)
|
|
226 |
fun args_of_tupel (Type("*",[left,right])) = left::(args_of_tupel right)
|
|
227 |
| args_of_tupel T = [T];
|
|
228 |
|
|
229 |
(* A1->...->An->R goes to [(A1,B1),...,(An,Bn)] where Bi=Ai->...->An->R
|
|
230 |
Bi is the Type of the function that is applied to an Ai type argument *)
|
|
231 |
fun args_of_curried (typ as (Type("->",[S,T]))) =
|
|
232 |
(S,typ) :: args_of_curried T
|
|
233 |
| args_of_curried _ = [];
|
|
234 |
in
|
|
235 |
fun args_of_op true typ = Curried_args(rev(args_of_curried typ))
|
|
236 |
| args_of_op false (typ as (Type("->",[S,T]))) =
|
|
237 |
Tupeled_args(args_of_tupel S)
|
|
238 |
| args_of_op false _ = Tupeled_args([]);
|
|
239 |
end;
|
|
240 |
|
|
241 |
(* generates for the type t the type of the fapp constant
|
|
242 |
that will be applied to t *)
|
|
243 |
fun mk_fapp_typ (typ as Type("->",argl)) = Type("fun",[typ,Type("fun",argl)])
|
1284
|
244 |
| mk_fapp_typ t = (
|
1274
|
245 |
error("Internal error:mk_fapp_typ: wrong argument\n"));
|
|
246 |
|
|
247 |
fun mk_arg_tupel_UU uu_pos [typ] n =
|
|
248 |
if n<>uu_pos then Free("x"^(string_of_int n),typ)
|
|
249 |
else Const("UU",typ)
|
|
250 |
| mk_arg_tupel_UU uu_pos (typ::tail) n =
|
|
251 |
mkCPair
|
|
252 |
(if n<>uu_pos then Free("x"^(string_of_int n),typ)
|
|
253 |
else Const("UU",typ))
|
|
254 |
(mk_arg_tupel_UU uu_pos tail (n+1))
|
|
255 |
| mk_arg_tupel_UU _ [] _ = error("Internal error:mk_arg_tupel: empty list");
|
|
256 |
|
|
257 |
fun mk_app_UU cnt uu_pos fname (Curried_args((typ,ftyp)::tl)) =
|
|
258 |
Const("fapp",mk_fapp_typ ftyp) $
|
|
259 |
(mk_app_UU (cnt-1) uu_pos fname (Curried_args tl))$
|
|
260 |
(if cnt = uu_pos then Const("UU",typ)
|
|
261 |
else Free("x"^(string_of_int cnt),typ))
|
|
262 |
| mk_app_UU _ _ (name,typ) (Curried_args []) = Const(name,typ)
|
|
263 |
| mk_app_UU cnt uu_pos (name,typ) (Tupeled_args []) = Const(name,typ)
|
|
264 |
| mk_app_UU cnt uu_pos (name,typ) (Tupeled_args list) =
|
|
265 |
Const("fapp",mk_fapp_typ typ) $ Const(name,typ) $
|
|
266 |
mk_arg_tupel_UU uu_pos list 0;
|
|
267 |
|
|
268 |
fun mk_app cnt fname args = mk_app_UU cnt (~1) fname args;
|
|
269 |
|
|
270 |
(* producing the strictness axioms *)
|
|
271 |
local
|
|
272 |
fun s_axm_of curried name typ args num cnt =
|
|
273 |
if cnt = num then
|
|
274 |
error("Internal error: s_axm_of: arg is no operation "^(external name))
|
|
275 |
else
|
|
276 |
let val app = mk_app_UU (num-1) cnt (internal name,typ) args
|
|
277 |
val equation = HOLogic.mk_eq(app,Const("UU",fastype_of app))
|
|
278 |
in
|
|
279 |
if cnt = num-1 then equation
|
|
280 |
else And $ equation $
|
|
281 |
s_axm_of curried name typ args num (cnt+1)
|
|
282 |
end;
|
|
283 |
in
|
|
284 |
fun strictness_axms curried ((rawname,typ,cmixfix)::tail) =
|
|
285 |
let val name = case cmixfix of
|
|
286 |
(CInfixl _) => (mk_internal_name rawname,rawname)
|
|
287 |
| (CInfixr _) => (mk_internal_name rawname,rawname)
|
|
288 |
| _ => (rawname,rawname)
|
|
289 |
val args = args_of_op curried typ;
|
|
290 |
val num = num_of_args args;
|
|
291 |
in
|
|
292 |
((external name)^"_strict",
|
|
293 |
if num <> 0
|
|
294 |
then HOLogic.mk_Trueprop(s_axm_of curried name typ args num 0)
|
|
295 |
else HOLogic.mk_Trueprop(True)) :: strictness_axms curried tail
|
|
296 |
end
|
|
297 |
| strictness_axms _ [] = [];
|
|
298 |
end; (*local*)
|
|
299 |
|
|
300 |
(* producing the totality axioms *)
|
|
301 |
|
|
302 |
fun totality_axms curried ((rawname,typ,cmixfix)::tail) =
|
|
303 |
let val name = case cmixfix of
|
|
304 |
(CInfixl _) => (mk_internal_name rawname,rawname)
|
|
305 |
| (CInfixr _) => (mk_internal_name rawname,rawname)
|
|
306 |
| _ => (rawname,rawname)
|
|
307 |
val args = args_of_op curried typ;
|
|
308 |
val prems = mk_mkNotEqUU_vars (if curried then rev (types_of args)
|
|
309 |
else (types_of args)) 0;
|
|
310 |
val term = mk_app (num_of_args args - 1) (internal name,typ) args;
|
|
311 |
in
|
|
312 |
((external name)^"_total",
|
|
313 |
if num_of_args args <> 0
|
|
314 |
then Logic.list_implies (prems,mkNotEqUU term)
|
|
315 |
else HOLogic.mk_Trueprop(True)) :: totality_axms curried tail
|
|
316 |
end
|
|
317 |
| totality_axms _ [] = [];
|
|
318 |
|
|
319 |
|
|
320 |
|
|
321 |
(* the theory extenders ****************************)
|
|
322 |
|
|
323 |
fun add_ops {curried,strict,total} raw_decls thy =
|
|
324 |
let val {sign,...} = rep_theory thy;
|
|
325 |
val decls = map (apsnd_of_three (typ_of o read_ctyp sign)) raw_decls;
|
|
326 |
val oldstyledecls = oldstyle decls;
|
|
327 |
val syndecls = syn_decls curried sign decls;
|
|
328 |
val xrules = xrules_of curried decls;
|
|
329 |
val s_axms = (if strict then strictness_axms curried decls else []);
|
|
330 |
val t_axms = (if total then totality_axms curried decls else []);
|
|
331 |
in
|
|
332 |
add_trrules_i xrules (add_axioms_i (s_axms @ t_axms)
|
|
333 |
(add_syntax_i syndecls (add_consts_i oldstyledecls thy)))
|
|
334 |
end;
|
|
335 |
|
|
336 |
fun add_ops_i {curried,strict,total} decls thy =
|
|
337 |
let val {sign,...} = rep_theory thy;
|
|
338 |
val oldstyledecls = oldstyle decls;
|
|
339 |
val syndecls = syn_decls curried sign decls;
|
|
340 |
val xrules = xrules_of curried decls;
|
|
341 |
val s_axms = (if strict then strictness_axms curried decls else []);
|
|
342 |
val t_axms = (if total then totality_axms curried decls else []);
|
|
343 |
in
|
|
344 |
add_trrules_i xrules (add_axioms_i (s_axms @ t_axms)
|
|
345 |
(add_syntax_i syndecls (add_consts_i oldstyledecls thy)))
|
|
346 |
end;
|
|
347 |
|
|
348 |
|
|
349 |
(* parser: ops_decls ********************************)
|
|
350 |
|
|
351 |
local open ThyParse
|
|
352 |
in
|
|
353 |
(* the following is an adapted version of const_decls from thy_parse.ML *)
|
|
354 |
|
|
355 |
val names1 = list1 name;
|
|
356 |
|
|
357 |
val split_decls = flat o map (fn (xs, y) => map (rpair y) xs);
|
|
358 |
|
|
359 |
fun mk_triple2 (x, (y, z)) = mk_triple (x, y, z);
|
|
360 |
|
|
361 |
fun mk_strict_vals [] = ""
|
|
362 |
| mk_strict_vals [name] =
|
|
363 |
"get_axiom thy \""^name^"_strict\"\n"
|
|
364 |
| mk_strict_vals (name::tail) =
|
|
365 |
"get_axiom thy \""^name^"_strict\",\n"^
|
|
366 |
mk_strict_vals tail;
|
|
367 |
|
|
368 |
fun mk_total_vals [] = ""
|
|
369 |
| mk_total_vals [name] =
|
|
370 |
"get_axiom thy \""^name^"_total\"\n"
|
|
371 |
| mk_total_vals (name::tail) =
|
|
372 |
"get_axiom thy \""^name^"_total\",\n"^
|
|
373 |
mk_total_vals tail;
|
|
374 |
|
|
375 |
fun mk_ops_decls (((curried,strict),total),list) =
|
|
376 |
(* call for the theory extender *)
|
|
377 |
("|> ThyOps.add_ops \n"^
|
|
378 |
"{ curried = "^curried^" , strict = "^strict^
|
|
379 |
" , total = "^total^" } \n"^
|
|
380 |
(mk_big_list o map mk_triple2) list^";\n"^
|
|
381 |
"val strict_axms = []; val total_axms = [];\nval thy = thy\n",
|
|
382 |
(* additional declarations *)
|
|
383 |
(if strict="true" then "val strict_axms = strict_axms @ [\n"^
|
|
384 |
mk_strict_vals (map (strip_quotes o fst) list)^
|
|
385 |
"];\n"
|
|
386 |
else "")^
|
|
387 |
(if total="true" then "val total_axms = total_axms @ [\n"^
|
|
388 |
mk_total_vals (map (strip_quotes o fst) list)^
|
|
389 |
"];\n"
|
|
390 |
else ""));
|
|
391 |
|
|
392 |
(* mixfix annotations *)
|
|
393 |
|
|
394 |
fun cat_parens pre1 pre2 s = cat pre1 (parens (cat pre2 s));
|
|
395 |
|
|
396 |
val infxl = "infixl" $$-- !! nat >> cat_parens "ThyOps.Mixfix" "Infixl";
|
|
397 |
val infxr = "infixr" $$-- !! nat >> cat_parens "ThyOps.Mixfix" "Infixr";
|
|
398 |
|
|
399 |
val cinfxl = "cinfixl" $$-- !! nat >> cat "ThyOps.CInfixl";
|
|
400 |
val cinfxr = "cinfixr" $$-- !! nat >> cat "ThyOps.CInfixr";
|
|
401 |
|
|
402 |
val opt_pris = optional ("[" $$-- !! (list nat --$$ "]")) [] >> mk_list;
|
|
403 |
|
|
404 |
val cmixfx = "cmixfix" $$-- string -- !! (opt_pris -- optional nat "max_pri")
|
|
405 |
>> (cat "ThyOps.CMixfix" o mk_triple2);
|
|
406 |
|
|
407 |
val bindr = "binder" $$--
|
|
408 |
!! (string -- ( ("[" $$-- nat --$$ "]") -- nat
|
|
409 |
|| nat >> (fn n => (n,n))
|
|
410 |
) )
|
|
411 |
>> (cat_parens "ThyOps.Mixfix" "Binder" o mk_triple2);
|
|
412 |
|
|
413 |
val mixfx = string -- !! (opt_pris -- optional nat "max_pri")
|
|
414 |
>> (cat_parens "ThyOps.Mixfix" "Mixfix" o mk_triple2);
|
|
415 |
|
|
416 |
fun opt_syn fx = optional ("(" $$-- fx --$$ ")") "ThyOps.Mixfix NoSyn";
|
|
417 |
|
|
418 |
val opt_cmixfix = opt_syn (mixfx || infxl || infxr || bindr ||
|
|
419 |
cinfxl || cinfxr || cmixfx);
|
|
420 |
|
|
421 |
fun ops_decls toks=
|
|
422 |
(optional ($$ "curried" >> K "true") "false" --
|
|
423 |
optional ($$ "strict" >> K "true") "false" --
|
|
424 |
optional ($$ "total" >> K "true") "false" --
|
|
425 |
(repeat1 (names1 --$$ "::" -- !! (string -- opt_cmixfix))
|
|
426 |
>> split_decls)
|
|
427 |
>> mk_ops_decls) toks
|
|
428 |
|
|
429 |
end;
|
|
430 |
|
|
431 |
(*** new keywords and sections: ******************************************)
|
|
432 |
|
|
433 |
val ops_keywords = ["curried","strict","total","cinfixl","cinfixr","cmixfix"];
|
|
434 |
(* "::" is already a pure keyword *)
|
|
435 |
|
|
436 |
val ops_sections = [("ops" , ops_decls) ];
|
|
437 |
|
|
438 |
end; (* the structure ThyOps *)
|
|
439 |
|