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