author | wenzelm |
Mon, 26 Sep 1994 17:55:45 +0100 | |
changeset 620 | f787eb061618 |
parent 583 | 550292083e66 |
child 623 | ca9f5dbab880 |
permissions | -rw-r--r-- |
19 | 1 |
(* Title: Pure/sign.ML |
0 | 2 |
ID: $Id$ |
251 | 3 |
Author: Lawrence C Paulson and Markus Wenzel |
0 | 4 |
|
251 | 5 |
The abstract type "sg" of signatures. |
0 | 6 |
*) |
7 |
||
19 | 8 |
signature SIGN = |
0 | 9 |
sig |
10 |
structure Symtab: SYMTAB |
|
11 |
structure Syntax: SYNTAX |
|
251 | 12 |
structure Type: TYPE |
143 | 13 |
sharing Symtab = Type.Symtab |
562 | 14 |
local open Type Syntax in |
251 | 15 |
type sg |
16 |
val rep_sg: sg -> |
|
17 |
{tsig: type_sig, |
|
18 |
const_tab: typ Symtab.table, |
|
386 | 19 |
syn: syntax, |
251 | 20 |
stamps: string ref list} |
21 |
val subsig: sg * sg -> bool |
|
22 |
val eq_sg: sg * sg -> bool |
|
386 | 23 |
val is_draft: sg -> bool |
24 |
val const_type: sg -> string -> typ option |
|
402 | 25 |
val classes: sg -> class list |
583 | 26 |
val subsort: sg -> sort * sort -> bool |
27 |
val norm_sort: sg -> sort -> sort |
|
251 | 28 |
val print_sg: sg -> unit |
562 | 29 |
val pretty_sg: sg -> Pretty.T |
251 | 30 |
val pprint_sg: sg -> pprint_args -> unit |
386 | 31 |
val pretty_term: sg -> term -> Pretty.T |
32 |
val pretty_typ: sg -> typ -> Pretty.T |
|
620 | 33 |
val pretty_sort: sort -> Pretty.T |
251 | 34 |
val string_of_term: sg -> term -> string |
35 |
val string_of_typ: sg -> typ -> string |
|
36 |
val pprint_term: sg -> term -> pprint_args -> unit |
|
37 |
val pprint_typ: sg -> typ -> pprint_args -> unit |
|
38 |
val certify_typ: sg -> typ -> typ |
|
39 |
val certify_term: sg -> term -> term * typ * int |
|
40 |
val read_typ: sg * (indexname -> sort option) -> string -> typ |
|
562 | 41 |
val infer_types: sg -> (indexname -> typ option) -> |
42 |
(indexname -> sort option) -> term * typ -> term * (indexname * typ) list |
|
43 |
val add_classes: (class * class list) list -> sg -> sg |
|
421 | 44 |
val add_classrel: (class * class) list -> sg -> sg |
386 | 45 |
val add_defsort: sort -> sg -> sg |
46 |
val add_types: (string * int * mixfix) list -> sg -> sg |
|
47 |
val add_tyabbrs: (string * string list * string * mixfix) list -> sg -> sg |
|
48 |
val add_tyabbrs_i: (string * string list * typ * mixfix) list -> sg -> sg |
|
49 |
val add_arities: (string * sort list * sort) list -> sg -> sg |
|
50 |
val add_consts: (string * string * mixfix) list -> sg -> sg |
|
51 |
val add_consts_i: (string * typ * mixfix) list -> sg -> sg |
|
52 |
val add_syntax: (string * string * mixfix) list -> sg -> sg |
|
53 |
val add_syntax_i: (string * typ * mixfix) list -> sg -> sg |
|
54 |
val add_trfuns: |
|
55 |
(string * (ast list -> ast)) list * |
|
56 |
(string * (term list -> term)) list * |
|
57 |
(string * (term list -> term)) list * |
|
58 |
(string * (ast list -> ast)) list -> sg -> sg |
|
59 |
val add_trrules: xrule list -> sg -> sg |
|
60 |
val add_name: string -> sg -> sg |
|
61 |
val make_draft: sg -> sg |
|
251 | 62 |
val merge: sg * sg -> sg |
63 |
val pure: sg |
|
386 | 64 |
val const_of_class: class -> string |
65 |
val class_of_const: string -> class |
|
251 | 66 |
end |
0 | 67 |
end; |
68 |
||
251 | 69 |
functor SignFun(structure Syntax: SYNTAX and Type: TYPE): SIGN = |
143 | 70 |
struct |
0 | 71 |
|
72 |
structure Symtab = Type.Symtab; |
|
73 |
structure Syntax = Syntax; |
|
386 | 74 |
structure BasicSyntax: BASIC_SYNTAX = Syntax; |
251 | 75 |
structure Pretty = Syntax.Pretty; |
76 |
structure Type = Type; |
|
562 | 77 |
open BasicSyntax Type; |
0 | 78 |
|
143 | 79 |
|
251 | 80 |
(** datatype sg **) |
81 |
||
82 |
(*the "ref" in stamps ensures that no two signatures are identical -- it is |
|
83 |
impossible to forge a signature*) |
|
143 | 84 |
|
19 | 85 |
datatype sg = |
143 | 86 |
Sg of { |
251 | 87 |
tsig: type_sig, (*order-sorted signature of types*) |
143 | 88 |
const_tab: typ Symtab.table, (*types of constants*) |
89 |
syn: Syntax.syntax, (*syntax for parsing and printing*) |
|
90 |
stamps: string ref list}; (*unique theory indentifier*) |
|
0 | 91 |
|
92 |
fun rep_sg (Sg args) = args; |
|
402 | 93 |
val tsig_of = #tsig o rep_sg; |
0 | 94 |
|
206
0d624d1ba9cc
added subsig: sg * sg -> bool to test if one signature is contained in another.
nipkow
parents:
200
diff
changeset
|
95 |
|
402 | 96 |
(* stamps *) |
97 |
||
386 | 98 |
fun is_draft (Sg {stamps = ref "#" :: _, ...}) = true |
99 |
| is_draft _ = false; |
|
100 |
||
251 | 101 |
fun subsig (Sg {stamps = s1, ...}, Sg {stamps = s2, ...}) = s1 subset s2; |
0 | 102 |
|
266 | 103 |
fun eq_sg (Sg {stamps = s1, ...}, Sg {stamps = s2, ...}) = eq_set (s1, s2); |
0 | 104 |
|
105 |
||
402 | 106 |
(* consts *) |
107 |
||
386 | 108 |
fun const_type (Sg {const_tab, ...}) c = |
109 |
Symtab.lookup (const_tab, c); |
|
110 |
||
111 |
||
620 | 112 |
(* classes and sorts *) |
402 | 113 |
|
114 |
val classes = #classes o Type.rep_tsig o tsig_of; |
|
115 |
||
116 |
val subsort = Type.subsort o tsig_of; |
|
117 |
val norm_sort = Type.norm_sort o tsig_of; |
|
118 |
||
620 | 119 |
fun pretty_sort [c] = Pretty.str c |
120 |
| pretty_sort cs = Pretty.str_list "{" "}" cs; |
|
121 |
||
402 | 122 |
|
0 | 123 |
|
251 | 124 |
(** print signature **) |
125 |
||
386 | 126 |
val stamp_names = rev o map !; |
127 |
||
251 | 128 |
fun print_sg sg = |
129 |
let |
|
130 |
fun prt_typ syn ty = Pretty.quote (Syntax.pretty_typ syn ty); |
|
131 |
||
132 |
fun pretty_subclass (cl, cls) = Pretty.block |
|
133 |
[Pretty.str (cl ^ " <"), Pretty.brk 1, pretty_sort cls]; |
|
134 |
||
135 |
fun pretty_default cls = Pretty.block |
|
136 |
[Pretty.str "default:", Pretty.brk 1, pretty_sort cls]; |
|
137 |
||
138 |
fun pretty_arg (ty, n) = Pretty.str (ty ^ " " ^ string_of_int n); |
|
139 |
||
140 |
fun pretty_abbr syn (ty, (vs, rhs)) = Pretty.block |
|
620 | 141 |
[prt_typ syn (Type (ty, map (fn v => TVar ((v, 0), [])) vs)), |
251 | 142 |
Pretty.str " =", Pretty.brk 1, prt_typ syn rhs]; |
143 |
||
144 |
fun pretty_arity ty (cl, []) = Pretty.block |
|
145 |
[Pretty.str (ty ^ " ::"), Pretty.brk 1, Pretty.str cl] |
|
146 |
| pretty_arity ty (cl, srts) = Pretty.block |
|
147 |
[Pretty.str (ty ^ " ::"), Pretty.brk 1, |
|
148 |
Pretty.list "(" ")" (map pretty_sort srts), |
|
149 |
Pretty.brk 1, Pretty.str cl]; |
|
150 |
||
151 |
fun pretty_coreg (ty, ars) = map (pretty_arity ty) ars; |
|
152 |
||
153 |
fun pretty_const syn (c, ty) = Pretty.block |
|
266 | 154 |
[Pretty.str (quote c ^ " ::"), Pretty.brk 1, prt_typ syn ty]; |
251 | 155 |
|
156 |
||
157 |
val Sg {tsig, const_tab, syn, stamps} = sg; |
|
158 |
val {classes, subclass, default, args, abbrs, coreg} = rep_tsig tsig; |
|
159 |
in |
|
386 | 160 |
Pretty.writeln (Pretty.strs ("stamps:" :: stamp_names stamps)); |
251 | 161 |
Pretty.writeln (Pretty.strs ("classes:" :: classes)); |
162 |
Pretty.writeln (Pretty.big_list "subclass:" (map pretty_subclass subclass)); |
|
163 |
Pretty.writeln (pretty_default default); |
|
164 |
Pretty.writeln (Pretty.big_list "types:" (map pretty_arg args)); |
|
165 |
Pretty.writeln (Pretty.big_list "abbrs:" (map (pretty_abbr syn) abbrs)); |
|
166 |
Pretty.writeln (Pretty.big_list "coreg:" (flat (map pretty_coreg coreg))); |
|
167 |
Pretty.writeln (Pretty.big_list "consts:" |
|
386 | 168 |
(map (pretty_const syn) (Symtab.dest const_tab))) |
251 | 169 |
end; |
170 |
||
171 |
||
562 | 172 |
fun pretty_sg (Sg {stamps, ...}) = |
173 |
Pretty.str_list "{" "}" (stamp_names stamps); |
|
174 |
||
175 |
val pprint_sg = Pretty.pprint o pretty_sg; |
|
251 | 176 |
|
177 |
||
178 |
||
179 |
(** pretty printing of terms and types **) |
|
180 |
||
181 |
fun pretty_term (Sg {syn, ...}) = Syntax.pretty_term syn; |
|
182 |
fun pretty_typ (Sg {syn, ...}) = Syntax.pretty_typ syn; |
|
183 |
||
184 |
fun string_of_term (Sg {syn, ...}) = Syntax.string_of_term syn; |
|
185 |
fun string_of_typ (Sg {syn, ...}) = Syntax.string_of_typ syn; |
|
186 |
||
187 |
fun pprint_term sg = Pretty.pprint o Pretty.quote o (pretty_term sg); |
|
188 |
fun pprint_typ sg = Pretty.pprint o Pretty.quote o (pretty_typ sg); |
|
0 | 189 |
|
190 |
||
169 | 191 |
|
562 | 192 |
(** read types **) (*exception ERROR*) |
193 |
||
194 |
fun err_in_type s = |
|
195 |
error ("The error(s) above occurred in type " ^ quote s); |
|
196 |
||
197 |
fun read_raw_typ syn tsig sort_of str = |
|
198 |
Syntax.read_typ syn (fn x => if_none (sort_of x) (defaultS tsig)) str |
|
199 |
handle ERROR => err_in_type str; |
|
200 |
||
201 |
(*read and certify typ wrt a signature*) |
|
202 |
fun read_typ (Sg {tsig, syn, ...}, sort_of) str = |
|
203 |
cert_typ tsig (read_raw_typ syn tsig sort_of str) |
|
204 |
handle TYPE (msg, _, _) => (writeln msg; err_in_type str); |
|
205 |
||
206 |
||
207 |
||
386 | 208 |
(** certify types and terms **) (*exception TYPE*) |
251 | 209 |
|
210 |
fun certify_typ (Sg {tsig, ...}) ty = cert_typ tsig ty; |
|
211 |
||
212 |
||
386 | 213 |
fun mapfilt_atoms f (Abs (_, _, t)) = mapfilt_atoms f t |
214 |
| mapfilt_atoms f (t $ u) = mapfilt_atoms f t @ mapfilt_atoms f u |
|
215 |
| mapfilt_atoms f a = (case f a of Some y => [y] | None => []); |
|
0 | 216 |
|
386 | 217 |
fun certify_term (sg as Sg {tsig, ...}) tm = |
251 | 218 |
let |
219 |
fun valid_const a T = |
|
386 | 220 |
(case const_type sg a of |
251 | 221 |
Some U => typ_instance (tsig, T, U) |
222 |
| _ => false); |
|
169 | 223 |
|
251 | 224 |
fun atom_err (Const (a, T)) = |
225 |
if valid_const a T then None |
|
226 |
else Some ("Illegal type for constant " ^ quote a ^ " :: " ^ |
|
227 |
quote (string_of_typ sg T)) |
|
228 |
| atom_err (Var ((x, i), _)) = |
|
229 |
if i < 0 then Some ("Negative index for Var " ^ quote x) else None |
|
230 |
| atom_err _ = None; |
|
231 |
||
232 |
val norm_tm = |
|
233 |
(case it_term_types (typ_errors tsig) (tm, []) of |
|
234 |
[] => map_term_types (norm_typ tsig) tm |
|
235 |
| errs => raise_type (cat_lines errs) [] [tm]); |
|
236 |
in |
|
386 | 237 |
(case mapfilt_atoms atom_err norm_tm of |
251 | 238 |
[] => (norm_tm, type_of norm_tm, maxidx_of_term norm_tm) |
239 |
| errs => raise_type (cat_lines errs) [] [norm_tm]) |
|
240 |
end; |
|
241 |
||
242 |
||
243 |
||
583 | 244 |
(** infer_types **) (*exception ERROR*) |
251 | 245 |
|
562 | 246 |
fun infer_types sg types sorts (t, T) = |
251 | 247 |
let |
562 | 248 |
val Sg {tsig, ...} = sg; |
249 |
val show_typ = string_of_typ sg; |
|
250 |
val show_term = string_of_term sg; |
|
169 | 251 |
|
562 | 252 |
fun term_err [] = "" |
253 |
| term_err [t] = "\nInvolving this term:\n" ^ show_term t |
|
254 |
| term_err ts = "\nInvolving these terms:\n" ^ cat_lines (map show_term ts); |
|
255 |
||
256 |
val T' = certify_typ sg T |
|
257 |
handle TYPE (msg, _, _) => error msg; |
|
258 |
val (t', tye) = Type.infer_types (tsig, const_type sg, types, sorts, T', t) |
|
259 |
handle TYPE (msg, Ts, ts) => error ("Type checking error: " ^ msg ^ "\n" |
|
260 |
^ cat_lines (map show_typ Ts) ^ term_err ts); |
|
261 |
in (t', tye) end; |
|
251 | 262 |
|
263 |
||
264 |
||
620 | 265 |
(** signature extension functions **) (*exception ERROR*) |
386 | 266 |
|
267 |
fun decls_of name_of mfixs = |
|
268 |
map (fn (x, y, mx) => (name_of x mx, y)) mfixs; |
|
269 |
||
270 |
||
271 |
(* add default sort *) |
|
272 |
||
273 |
fun ext_defsort (syn, tsig, ctab) defsort = |
|
274 |
(syn, ext_tsig_defsort tsig defsort, ctab); |
|
275 |
||
276 |
||
277 |
(* add type constructors *) |
|
278 |
||
279 |
fun ext_types (syn, tsig, ctab) types = |
|
280 |
(Syntax.extend_type_gram syn types, |
|
562 | 281 |
ext_tsig_types tsig (decls_of Syntax.type_name types), |
386 | 282 |
ctab); |
283 |
||
284 |
||
285 |
(* add type abbreviations *) |
|
286 |
||
287 |
fun read_abbr syn tsig (t, vs, rhs_src) = |
|
288 |
(t, vs, read_raw_typ syn tsig (K None) rhs_src) |
|
289 |
handle ERROR => error ("in type abbreviation " ^ t); |
|
290 |
||
291 |
fun ext_abbrs rd_abbr (syn, tsig, ctab) abbrs = |
|
292 |
let |
|
293 |
fun mfix_of (t, vs, _, mx) = (t, length vs, mx); |
|
294 |
val syn1 = Syntax.extend_type_gram syn (map mfix_of abbrs); |
|
295 |
||
296 |
fun decl_of (t, vs, rhs, mx) = |
|
562 | 297 |
rd_abbr syn1 tsig (Syntax.type_name t mx, vs, rhs); |
386 | 298 |
in |
299 |
(syn1, ext_tsig_abbrs tsig (map decl_of abbrs), ctab) |
|
300 |
end; |
|
301 |
||
302 |
val ext_tyabbrs_i = ext_abbrs (K (K I)); |
|
303 |
val ext_tyabbrs = ext_abbrs read_abbr; |
|
304 |
||
305 |
||
306 |
(* add type arities *) |
|
307 |
||
308 |
fun ext_arities (syn, tsig, ctab) arities = |
|
309 |
let |
|
310 |
val tsig1 = ext_tsig_arities tsig arities; |
|
311 |
val log_types = logical_types tsig1; |
|
312 |
in |
|
313 |
(Syntax.extend_log_types syn log_types, tsig1, ctab) |
|
314 |
end; |
|
315 |
||
316 |
||
317 |
(* add term constants and syntax *) |
|
318 |
||
319 |
fun err_in_const c = |
|
320 |
error ("in declaration of constant " ^ quote c); |
|
321 |
||
322 |
fun err_dup_consts cs = |
|
323 |
error ("Duplicate declaration of constant(s) " ^ commas_quote cs); |
|
324 |
||
251 | 325 |
|
386 | 326 |
fun read_const syn tsig (c, ty_src, mx) = |
327 |
(c, read_raw_typ syn tsig (K None) ty_src, mx) |
|
562 | 328 |
handle ERROR => err_in_const (Syntax.const_name c mx); |
386 | 329 |
|
330 |
fun ext_cnsts rd_const syn_only (syn, tsig, ctab) raw_consts = |
|
331 |
let |
|
620 | 332 |
fun prep_const (c, ty, mx) = (c, varifyT (cert_typ tsig (no_tvars ty)), mx) |
562 | 333 |
handle TYPE (msg, _, _) => (writeln msg; err_in_const (Syntax.const_name c mx)); |
386 | 334 |
|
335 |
val consts = map (prep_const o rd_const syn tsig) raw_consts; |
|
336 |
val decls = |
|
337 |
if syn_only then [] |
|
562 | 338 |
else filter_out (equal "" o fst) (decls_of Syntax.const_name consts); |
386 | 339 |
in |
340 |
(Syntax.extend_const_gram syn consts, tsig, |
|
341 |
Symtab.extend_new (ctab, decls) |
|
342 |
handle Symtab.DUPS cs => err_dup_consts cs) |
|
343 |
end; |
|
344 |
||
345 |
val ext_consts_i = ext_cnsts (K (K I)) false; |
|
346 |
val ext_consts = ext_cnsts read_const false; |
|
347 |
val ext_syntax_i = ext_cnsts (K (K I)) true; |
|
348 |
val ext_syntax = ext_cnsts read_const true; |
|
349 |
||
350 |
||
351 |
(* add type classes *) |
|
352 |
||
353 |
fun const_of_class c = c ^ "_class"; |
|
354 |
||
355 |
fun class_of_const c_class = |
|
356 |
let |
|
357 |
val c = implode (take (size c_class - 6, explode c_class)); |
|
358 |
in |
|
359 |
if const_of_class c = c_class then c |
|
360 |
else raise_term ("class_of_const: bad name " ^ quote c_class) [] |
|
361 |
end; |
|
362 |
||
363 |
||
364 |
fun ext_classes (syn, tsig, ctab) classes = |
|
365 |
let |
|
562 | 366 |
val names = map fst classes; |
386 | 367 |
val consts = |
368 |
map (fn c => (const_of_class c, a_itselfT --> propT, NoSyn)) names; |
|
369 |
in |
|
370 |
ext_consts_i |
|
371 |
(Syntax.extend_consts syn names, ext_tsig_classes tsig classes, ctab) |
|
372 |
consts |
|
373 |
end; |
|
374 |
||
375 |
||
421 | 376 |
(* add to subclass relation *) |
377 |
||
378 |
fun ext_classrel (syn, tsig, ctab) pairs = |
|
379 |
(syn, ext_tsig_subclass tsig pairs, ctab); |
|
380 |
||
381 |
||
386 | 382 |
(* add syntactic translations *) |
383 |
||
384 |
fun ext_trfuns (syn, tsig, ctab) trfuns = |
|
385 |
(Syntax.extend_trfuns syn trfuns, tsig, ctab); |
|
386 |
||
387 |
fun ext_trrules (syn, tsig, ctab) xrules = |
|
388 |
(Syntax.extend_trrules syn xrules, tsig, ctab); |
|
389 |
||
390 |
||
391 |
(* build signature *) |
|
392 |
||
393 |
fun ext_stamps stamps name = |
|
394 |
let |
|
395 |
val stmps = (case stamps of ref "#" :: ss => ss | ss => ss); |
|
396 |
in |
|
397 |
if exists (equal name o !) stmps then |
|
398 |
error ("Theory already contains a " ^ quote name ^ " component") |
|
402 | 399 |
else ref name :: stmps |
386 | 400 |
end; |
401 |
||
402 |
fun make_sign (syn, tsig, ctab) stamps name = |
|
403 |
Sg {tsig = tsig, const_tab = ctab, syn = syn, |
|
404 |
stamps = ext_stamps stamps name}; |
|
405 |
||
406 |
fun extend_sign extfun name decls (Sg {tsig, const_tab, syn, stamps}) = |
|
407 |
make_sign (extfun (syn, tsig, const_tab) decls) stamps name; |
|
408 |
||
409 |
||
410 |
(* the external interfaces *) |
|
411 |
||
412 |
val add_classes = extend_sign ext_classes "#"; |
|
421 | 413 |
val add_classrel = extend_sign ext_classrel "#"; |
386 | 414 |
val add_defsort = extend_sign ext_defsort "#"; |
415 |
val add_types = extend_sign ext_types "#"; |
|
416 |
val add_tyabbrs = extend_sign ext_tyabbrs "#"; |
|
417 |
val add_tyabbrs_i = extend_sign ext_tyabbrs_i "#"; |
|
418 |
val add_arities = extend_sign ext_arities "#"; |
|
419 |
val add_consts = extend_sign ext_consts "#"; |
|
420 |
val add_consts_i = extend_sign ext_consts_i "#"; |
|
421 |
val add_syntax = extend_sign ext_syntax "#"; |
|
422 |
val add_syntax_i = extend_sign ext_syntax_i "#"; |
|
423 |
val add_trfuns = extend_sign ext_trfuns "#"; |
|
424 |
val add_trrules = extend_sign ext_trrules "#"; |
|
425 |
||
426 |
fun add_name name sg = extend_sign K name () sg; |
|
427 |
val make_draft = add_name "#"; |
|
428 |
||
429 |
||
430 |
||
620 | 431 |
(** merge signatures **) (*exception TERM*) (*exception ERROR (* FIXME *)*) |
143 | 432 |
|
251 | 433 |
fun merge (sg1, sg2) = |
434 |
if subsig (sg2, sg1) then sg1 |
|
435 |
else if subsig (sg1, sg2) then sg2 |
|
386 | 436 |
else if is_draft sg1 orelse is_draft sg2 then |
437 |
raise_term "Illegal merge of draft signatures" [] |
|
251 | 438 |
else |
439 |
(*neither is union already; must form union*) |
|
440 |
let |
|
441 |
val Sg {tsig = tsig1, const_tab = const_tab1, syn = syn1, |
|
442 |
stamps = stamps1} = sg1; |
|
443 |
val Sg {tsig = tsig2, const_tab = const_tab2, syn = syn2, |
|
444 |
stamps = stamps2} = sg2; |
|
445 |
||
386 | 446 |
|
402 | 447 |
val stamps = merge_rev_lists stamps1 stamps2; |
448 |
val _ = |
|
449 |
(case duplicates (stamp_names stamps) of |
|
450 |
[] => () |
|
451 |
| dups => raise_term ("Attempt to merge different versions of theories " |
|
452 |
^ commas_quote dups) []); |
|
453 |
||
251 | 454 |
val tsig = merge_tsigs (tsig1, tsig2); |
455 |
||
456 |
val const_tab = Symtab.merge (op =) (const_tab1, const_tab2) |
|
386 | 457 |
handle Symtab.DUPS cs => |
458 |
raise_term ("Incompatible types for constant(s) " ^ commas_quote cs) []; |
|
459 |
||
460 |
val syn = Syntax.merge_syntaxes syn1 syn2; |
|
251 | 461 |
in |
462 |
Sg {tsig = tsig, const_tab = const_tab, syn = syn, stamps = stamps} |
|
463 |
end; |
|
143 | 464 |
|
0 | 465 |
|
466 |
||
251 | 467 |
(** the Pure signature **) |
0 | 468 |
|
386 | 469 |
val pure = |
470 |
make_sign (Syntax.type_syn, tsig0, Symtab.null) [] "#" |
|
410 | 471 |
|> add_types |
386 | 472 |
(("fun", 2, NoSyn) :: |
473 |
("prop", 0, NoSyn) :: |
|
474 |
("itself", 1, NoSyn) :: |
|
562 | 475 |
Syntax.pure_types) |
476 |
|> add_classes [(logicC, [])] |
|
410 | 477 |
|> add_defsort logicS |
478 |
|> add_arities |
|
386 | 479 |
[("fun", [logicS, logicS], logicS), |
480 |
("prop", [], logicS), |
|
481 |
("itself", [logicS], logicS)] |
|
562 | 482 |
|> add_syntax Syntax.pure_syntax |
410 | 483 |
|> add_trfuns Syntax.pure_trfuns |
484 |
|> add_consts |
|
386 | 485 |
[("==", "['a::{}, 'a] => prop", Mixfix ("(_ ==/ _)", [3, 2], 2)), |
486 |
("=?=", "['a::{}, 'a] => prop", Mixfix ("(_ =?=/ _)", [3, 2], 2)), |
|
487 |
("==>", "[prop, prop] => prop", Mixfix ("(_ ==>/ _)", [2, 1], 1)), |
|
488 |
("all", "('a => prop) => prop", Binder ("!!", 0)), |
|
573 | 489 |
("TYPE", "'a itself", NoSyn)] |
410 | 490 |
|> add_name "Pure"; |
0 | 491 |
|
492 |
||
493 |
end; |
|
143 | 494 |