src/Pure/sign.ML
changeset 3791 c5db2c87a646
parent 3552 f348e8a2db4b
child 3805 a50d0b5219dd
equal deleted inserted replaced
3790:95a47d8bcd69 3791:c5db2c87a646
     4 
     4 
     5 The abstract type "sg" of signatures.
     5 The abstract type "sg" of signatures.
     6 *)
     6 *)
     7 
     7 
     8 signature SIGN =
     8 signature SIGN =
     9   sig
     9 sig
    10   type sg
    10   type sg
    11   val rep_sg: sg ->
    11   val rep_sg: sg ->
    12    {tsig: Type.type_sig,
    12    {tsig: Type.type_sig,
    13     const_tab: typ Symtab.table,
    13     const_tab: typ Symtab.table,
    14     syn: Syntax.syntax,
    14     syn: Syntax.syntax,
       
    15     path: string list,
       
    16     spaces: (string * NameSpace.T) list,
    15     stamps: string ref list}
    17     stamps: string ref list}
    16   val subsig: sg * sg -> bool
    18   val subsig: sg * sg -> bool
    17   val eq_sg: sg * sg -> bool
    19   val eq_sg: sg * sg -> bool
    18   val same_sg: sg * sg -> bool
    20   val same_sg: sg * sg -> bool
    19   val is_draft: sg -> bool
    21   val is_draft: sg -> bool
    21   val classes: sg -> class list
    23   val classes: sg -> class list
    22   val subsort: sg -> sort * sort -> bool
    24   val subsort: sg -> sort * sort -> bool
    23   val nodup_Vars: term -> unit
    25   val nodup_Vars: term -> unit
    24   val norm_sort: sg -> sort -> sort
    26   val norm_sort: sg -> sort -> sort
    25   val nonempty_sort: sg -> sort list -> sort -> bool
    27   val nonempty_sort: sg -> sort list -> sort -> bool
       
    28   val long_names: bool ref
       
    29   val intern_class: sg -> class -> class
       
    30   val extern_class: sg -> class -> class
       
    31   val intern_sort: sg -> sort -> sort
       
    32   val extern_sort: sg -> sort -> sort
       
    33   val intern_typ: sg -> typ -> typ
       
    34   val extern_typ: sg -> typ -> typ
       
    35   val intern_term: sg -> term -> term
       
    36   val extern_term: sg -> term -> term
       
    37   val intern_tycons: sg -> typ -> typ
       
    38   val intern_tycon: sg -> string -> string
       
    39   val intern_const: sg -> string -> string
    26   val print_sg: sg -> unit
    40   val print_sg: sg -> unit
    27   val pretty_sg: sg -> Pretty.T
    41   val pretty_sg: sg -> Pretty.T
    28   val pprint_sg: sg -> pprint_args -> unit
    42   val pprint_sg: sg -> pprint_args -> unit
    29   val pretty_term: sg -> term -> Pretty.T
    43   val pretty_term: sg -> term -> Pretty.T
    30   val pretty_typ: sg -> typ -> Pretty.T
    44   val pretty_typ: sg -> typ -> Pretty.T
    31   val pretty_sort: sort -> Pretty.T
    45   val pretty_sort: sg -> sort -> Pretty.T
    32   val string_of_term: sg -> term -> string
    46   val string_of_term: sg -> term -> string
    33   val string_of_typ: sg -> typ -> string
    47   val string_of_typ: sg -> typ -> string
       
    48   val string_of_sort: sg -> sort -> string
    34   val pprint_term: sg -> term -> pprint_args -> unit
    49   val pprint_term: sg -> term -> pprint_args -> unit
    35   val pprint_typ: sg -> typ -> pprint_args -> unit
    50   val pprint_typ: sg -> typ -> pprint_args -> unit
    36   val certify_typ: sg -> typ -> typ
    51   val certify_typ: sg -> typ -> typ
    37   val certify_term: sg -> term -> term * typ * int
    52   val certify_term: sg -> term -> term * typ * int
    38   val read_typ: sg * (indexname -> sort option) -> string -> typ
    53   val read_typ: sg * (indexname -> sort option) -> string -> typ
    39   val infer_types: sg -> (indexname -> typ option) ->
    54   val infer_types: sg -> (indexname -> typ option) ->
    40     (indexname -> sort option) -> string list -> bool
    55     (indexname -> sort option) -> string list -> bool
    41     -> term list * typ -> int * term * (indexname * typ) list
    56     -> term list * typ -> int * term * (indexname * typ) list
    42   val add_classes: (class * class list) list -> sg -> sg
    57   val add_classes: (class * class list) list -> sg -> sg
       
    58   val add_classes_i: (class * class list) list -> sg -> sg
    43   val add_classrel: (class * class) list -> sg -> sg
    59   val add_classrel: (class * class) list -> sg -> sg
       
    60   val add_classrel_i: (class * class) list -> sg -> sg
    44   val add_defsort: sort -> sg -> sg
    61   val add_defsort: sort -> sg -> sg
       
    62   val add_defsort_i: sort -> sg -> sg
    45   val add_types: (string * int * mixfix) list -> sg -> sg
    63   val add_types: (string * int * mixfix) list -> sg -> sg
    46   val add_tyabbrs: (string * string list * string * mixfix) list -> sg -> sg
    64   val add_tyabbrs: (string * string list * string * mixfix) list -> sg -> sg
    47   val add_tyabbrs_i: (string * string list * typ * mixfix) list -> sg -> sg
    65   val add_tyabbrs_i: (string * string list * typ * mixfix) list -> sg -> sg
    48   val add_arities: (string * sort list * sort) list -> sg -> sg
    66   val add_arities: (string * sort list * sort) list -> sg -> sg
       
    67   val add_arities_i: (string * sort list * sort) list -> sg -> sg
    49   val add_consts: (string * string * mixfix) list -> sg -> sg
    68   val add_consts: (string * string * mixfix) list -> sg -> sg
    50   val add_consts_i: (string * typ * mixfix) list -> sg -> sg
    69   val add_consts_i: (string * typ * mixfix) list -> sg -> sg
    51   val add_syntax: (string * string * mixfix) list -> sg -> sg
    70   val add_syntax: (string * string * mixfix) list -> sg -> sg
    52   val add_syntax_i: (string * typ * mixfix) list -> sg -> sg
    71   val add_syntax_i: (string * typ * mixfix) list -> sg -> sg
    53   val add_modesyntax: (string * bool) * (string * string * mixfix) list -> sg -> sg
    72   val add_modesyntax: (string * bool) * (string * string * mixfix) list -> sg -> sg
    61     (string * (typ -> term list -> term)) list -> sg -> sg
    80     (string * (typ -> term list -> term)) list -> sg -> sg
    62   val add_tokentrfuns:
    81   val add_tokentrfuns:
    63     (string * string * (string -> string * int)) list -> sg -> sg
    82     (string * string * (string -> string * int)) list -> sg -> sg
    64   val add_trrules: (string * string) Syntax.trrule list -> sg -> sg
    83   val add_trrules: (string * string) Syntax.trrule list -> sg -> sg
    65   val add_trrules_i: ast Syntax.trrule list -> sg -> sg
    84   val add_trrules_i: ast Syntax.trrule list -> sg -> sg
       
    85   val add_path: string -> sg -> sg
       
    86   val add_space: string * string list -> sg -> sg
    66   val add_name: string -> sg -> sg
    87   val add_name: string -> sg -> sg
    67   val make_draft: sg -> sg
    88   val make_draft: sg -> sg
    68   val merge: sg * sg -> sg
    89   val merge: sg * sg -> sg
    69   val proto_pure: sg
    90   val proto_pure: sg
    70   val pure: sg
    91   val pure: sg
    71   val cpure: sg
    92   val cpure: sg
    72   val const_of_class: class -> string
    93   val const_of_class: class -> string
    73   val class_of_const: string -> class
    94   val class_of_const: string -> class
    74   end;
    95 end;
    75 
    96 
    76 structure Sign : SIGN =
    97 structure Sign : SIGN =
    77 struct
    98 struct
    78 
    99 
    79 
       
    80 (** datatype sg **)
   100 (** datatype sg **)
    81 
       
    82 (*the "ref" in stamps ensures that no two signatures are identical -- it is
       
    83   impossible to forge a signature*)
       
    84 
   101 
    85 datatype sg =
   102 datatype sg =
    86   Sg of {
   103   Sg of {
    87     tsig: Type.type_sig,            (*order-sorted signature of types*)
   104     tsig: Type.type_sig,                        (*order-sorted signature of types*)
    88     const_tab: typ Symtab.table,    (*types of constants*)
   105     const_tab: typ Symtab.table,                (*types of constants*)
    89     syn: Syntax.syntax,             (*syntax for parsing and printing*)
   106     syn: Syntax.syntax,                         (*syntax for parsing and printing*)
    90     stamps: string ref list};       (*unique theory indentifier*)
   107     path: string list,                     	(*current position in name space*)
       
   108     spaces: (string * NameSpace.T) list,   	(*name spaces for consts, types etc.*)
       
   109     stamps: string ref list};                   (*unique theory indentifier*)
       
   110 
       
   111     (*the "ref" in stamps ensures that no two signatures are identical
       
   112       -- it is impossible to forge a signature*)
    91 
   113 
    92 fun rep_sg (Sg args) = args;
   114 fun rep_sg (Sg args) = args;
    93 val tsig_of = #tsig o rep_sg;
   115 val tsig_of = #tsig o rep_sg;
    94 
   116 
    95 
   117 
   144 
   166 
   145 val subsort = Type.subsort o tsig_of;
   167 val subsort = Type.subsort o tsig_of;
   146 val norm_sort = Type.norm_sort o tsig_of;
   168 val norm_sort = Type.norm_sort o tsig_of;
   147 val nonempty_sort = Type.nonempty_sort o tsig_of;
   169 val nonempty_sort = Type.nonempty_sort o tsig_of;
   148 
   170 
   149 (* FIXME move to Sorts? *)
   171 
   150 fun pretty_sort [c] = Pretty.str c
   172 
   151   | pretty_sort cs = Pretty.str_list "{" "}" cs;
   173 (** name spaces **)
       
   174 
       
   175 (*prune names by default*)
       
   176 val long_names = ref false;
       
   177 
       
   178 
       
   179 (* kinds *)
       
   180 
       
   181 val classK = "class";
       
   182 val typeK = "type";
       
   183 val constK = "const";
       
   184 
       
   185 
       
   186 (* add and retrieve names *)
       
   187 
       
   188 fun space_of spaces kind =
       
   189   if_none (assoc (spaces, kind)) NameSpace.empty;
       
   190 
       
   191 (*input and output of qualified names*)
       
   192 fun intrn spaces kind = NameSpace.lookup (space_of spaces kind);
       
   193 fun extrn spaces kind = NameSpace.prune (space_of spaces kind);
       
   194 
       
   195 (*add names*)
       
   196 fun add_names spaces kind names =
       
   197   let val space' = NameSpace.extend (names, space_of spaces kind) in
       
   198     overwrite (spaces, (kind, space'))
       
   199   end;
       
   200 
       
   201 fun full_name path name = NameSpace.pack (path @ (NameSpace.unpack name));
       
   202 
       
   203 
       
   204 (* intern / extern names *)
       
   205 
       
   206 (*Note: These functions are not necessarily idempotent!*)	(* FIXME *)
       
   207 
       
   208 local
       
   209 
       
   210   (* FIXME move to term.ML? *)
       
   211 
       
   212   fun add_typ_classes (Type (_, Ts), cs) = foldr add_typ_classes (Ts, cs)
       
   213     | add_typ_classes (TFree (_, S), cs) = S union cs
       
   214     | add_typ_classes (TVar (_, S), cs) = S union cs;
       
   215 
       
   216   fun add_typ_tycons (Type (c, Ts), cs) = foldr add_typ_tycons (Ts, c ins cs)
       
   217     | add_typ_tycons (_, cs) = cs;
       
   218 
       
   219   val add_term_classes = it_term_types add_typ_classes;
       
   220   val add_term_tycons = it_term_types add_typ_tycons;
       
   221 
       
   222   fun add_term_consts (Const (c, _), cs) = c ins cs
       
   223     | add_term_consts (t $ u, cs) = add_term_consts (t, add_term_consts (u, cs))
       
   224     | add_term_consts (Abs (_, _, t), cs) = add_term_consts (t, cs)
       
   225     | add_term_consts (_, cs) = cs;
       
   226 
       
   227 
       
   228   (*map classes, tycons*)
       
   229   fun map_typ f g (Type (c, Ts)) = Type (g c, map (map_typ f g) Ts)
       
   230     | map_typ f _ (TFree (x, S)) = TFree (x, map f S)
       
   231     | map_typ f _ (TVar (xi, S)) = TVar (xi, map f S);
       
   232 
       
   233   (*map classes, tycons, consts*)
       
   234   fun map_term f g h (Const (c, T)) = Const (h c, map_typ f g T)
       
   235     | map_term f g _ (Free (x, T)) = Free (x, map_typ f g T)
       
   236     | map_term f g _ (Var (xi, T)) = Var (xi, map_typ f g T)
       
   237     | map_term _ _ _ (t as Bound _) = t
       
   238     | map_term f g h (Abs (x, T, t)) = Abs (x, map_typ f g T, map_term f g h t)
       
   239     | map_term f g h (t $ u) = map_term f g h t $ map_term f g h u;
       
   240 
       
   241   (*prepare mapping of names*)
       
   242   fun mapping f add_xs t =
       
   243     let
       
   244       fun f' x = let val y = f x in if x = y then None else Some (x, y) end;
       
   245       val table = mapfilter f' (add_xs (t, []));
       
   246       fun lookup x = if_none (assoc (table, x)) x;
       
   247     in lookup end;
       
   248 
       
   249   (*intern / extern typ*)
       
   250   fun trn_typ trn T =
       
   251     T |> map_typ
       
   252       (mapping (trn classK) add_typ_classes T)
       
   253       (mapping (trn typeK) add_typ_tycons T);
       
   254 
       
   255   (*intern / extern term*)
       
   256   fun trn_term trn t =
       
   257     t |> map_term
       
   258       (mapping (trn classK) add_term_classes t)
       
   259       (mapping (trn typeK) add_term_tycons t)
       
   260       (mapping (trn constK) add_term_consts t);
       
   261 
       
   262 
       
   263   fun spaces_of (Sg {spaces, ...}) = spaces;
       
   264 
       
   265 in
       
   266 
       
   267   fun intrn_class spaces = intrn spaces classK;
       
   268   fun extrn_class spaces = extrn spaces classK;
       
   269   val intrn_sort = map o intrn_class;
       
   270   val extrn_sort = map o extrn_class;
       
   271   val intrn_typ = trn_typ o intrn;
       
   272   val extrn_typ = trn_typ o extrn;
       
   273   val intrn_term = trn_term o intrn;
       
   274   val extrn_term = trn_term o extrn;
       
   275 
       
   276   val intern_class = intrn_class o spaces_of;
       
   277   val extern_class = extrn_class o spaces_of;
       
   278   val intern_sort = intrn_sort o spaces_of;
       
   279   val extern_sort = extrn_sort o spaces_of;
       
   280   val intern_typ = intrn_typ o spaces_of;
       
   281   val extern_typ = extrn_typ o spaces_of;
       
   282   val intern_term = intrn_term o spaces_of;
       
   283   val extern_term = extrn_term o spaces_of;
       
   284 
       
   285   fun intrn_tycons spaces T =
       
   286     map_typ I (mapping (intrn spaces typeK) add_typ_tycons T) T;
       
   287 
       
   288   fun intern_tycon sg = intrn (spaces_of sg) typeK;
       
   289   fun intern_const sg = intrn (spaces_of sg) constK;
       
   290   val intern_tycons = intrn_tycons o spaces_of;
       
   291 
       
   292 end;
   152 
   293 
   153 
   294 
   154 
   295 
   155 (** print signature **)
   296 (** print signature **)
   156 
   297 
   157 fun stamp_names stamps = rev (map ! stamps);
   298 fun stamp_names stamps = rev (map ! stamps);
   158 
   299 
   159 fun print_sg sg =
   300 fun print_sg sg =
   160   let
   301   let
   161     fun prt_typ syn ty = Pretty.quote (Syntax.pretty_typ syn ty);
   302     val Sg {syn, ...} = sg;
   162 
   303 
   163     fun pretty_classrel (cl, cls) = Pretty.block
   304     fun prt_typ ty = Pretty.quote (Syntax.pretty_typ syn ty);
   164       (Pretty.str (cl ^ " <") :: Pretty.brk 1 ::
   305     fun prt_cls c = Syntax.pretty_sort syn [c];
   165         Pretty.commas (map Pretty.str cls));
   306 
   166 
   307     fun pretty_space (kind, space) = Pretty.block (Pretty.breaks
   167     fun pretty_default cls = Pretty.block
   308       (map Pretty.str (kind ^ ":" :: map quote (NameSpace.dest space))));
   168       [Pretty.str "default:", Pretty.brk 1, pretty_sort cls];
   309 
       
   310     fun pretty_classes cs = Pretty.block
       
   311       (Pretty.breaks (Pretty.str "classes:" :: map prt_cls cs));
       
   312 
       
   313     fun pretty_classrel (c, cs) = Pretty.block
       
   314       (prt_cls c :: Pretty.str " <" :: Pretty.brk 1 ::
       
   315         Pretty.commas (map prt_cls cs));
       
   316 
       
   317     fun pretty_default S = Pretty.block
       
   318       [Pretty.str "default:", Pretty.brk 1, Syntax.pretty_sort syn S];
   169 
   319 
   170     fun pretty_ty (ty, n) = Pretty.str (ty ^ " " ^ string_of_int n);
   320     fun pretty_ty (ty, n) = Pretty.str (ty ^ " " ^ string_of_int n);
   171 
   321 
   172     fun pretty_abbr syn (ty, (vs, rhs)) = Pretty.block
   322     fun pretty_abbr (ty, (vs, rhs)) = Pretty.block
   173       [prt_typ syn (Type (ty, map (fn v => TVar ((v, 0), [])) vs)),
   323       [prt_typ (Type (ty, map (fn v => TVar ((v, 0), [])) vs)),
   174         Pretty.str " =", Pretty.brk 1, prt_typ syn rhs];
   324         Pretty.str " =", Pretty.brk 1, prt_typ rhs];
   175 
   325 
   176     fun pretty_arity ty (cl, []) = Pretty.block
   326     fun pretty_arity ty (c, []) = Pretty.block
   177           [Pretty.str (ty ^ " ::"), Pretty.brk 1, Pretty.str cl]
   327           [Pretty.str (ty ^ " ::"), Pretty.brk 1, prt_cls c]
   178       | pretty_arity ty (cl, srts) = Pretty.block
   328       | pretty_arity ty (c, Ss) = Pretty.block
   179           [Pretty.str (ty ^ " ::"), Pretty.brk 1,
   329           [Pretty.str (ty ^ " ::"), Pretty.brk 1,
   180             Pretty.list "(" ")" (map pretty_sort srts),
   330             Pretty.list "(" ")" (map (Syntax.pretty_sort syn) Ss),
   181             Pretty.brk 1, Pretty.str cl];
   331             Pretty.brk 1, prt_cls c];
   182 
   332 
   183     fun pretty_arities (ty, ars) = map (pretty_arity ty) ars;
   333     fun pretty_arities (ty, ars) = map (pretty_arity ty) ars;
   184 
   334 
   185     fun pretty_const syn (c, ty) = Pretty.block
   335     fun pretty_const (c, ty) = Pretty.block
   186       [Pretty.str (quote c ^ " ::"), Pretty.brk 1, prt_typ syn ty];
   336       [Pretty.str (quote c ^ " ::"), Pretty.brk 1, prt_typ ty];
   187 
   337 
   188 
   338     val Sg {tsig, const_tab, syn, path, spaces, stamps} = sg;
   189     val Sg {tsig, const_tab, syn, stamps} = sg;
       
   190     val {classes, classrel, default, tycons, abbrs, arities} =
   339     val {classes, classrel, default, tycons, abbrs, arities} =
   191       Type.rep_tsig tsig;
   340       Type.rep_tsig tsig;
   192   in
   341   in
   193     Pretty.writeln (Pretty.strs ("stamps:" :: stamp_names stamps));
   342     Pretty.writeln (Pretty.strs ("stamps:" :: stamp_names stamps));
   194     Pretty.writeln (Pretty.strs ("classes:" :: classes));
   343     Pretty.writeln (Pretty.strs ["name entry path:", NameSpace.pack path]);
   195     Pretty.writeln (Pretty.big_list "classrel:" (map pretty_classrel classrel));
   344     Pretty.writeln (Pretty.big_list "name spaces:" (map pretty_space spaces));
       
   345     Pretty.writeln (pretty_classes classes);
       
   346     Pretty.writeln (Pretty.big_list "class relation:" (map pretty_classrel classrel));
   196     Pretty.writeln (pretty_default default);
   347     Pretty.writeln (pretty_default default);
   197     Pretty.writeln (Pretty.big_list "types:" (map pretty_ty tycons));
   348     Pretty.writeln (Pretty.big_list "type constructors:" (map pretty_ty tycons));
   198     Pretty.writeln (Pretty.big_list "abbrs:" (map (pretty_abbr syn) abbrs));
   349     Pretty.writeln (Pretty.big_list "type abbreviations:" (map pretty_abbr abbrs));
   199     Pretty.writeln (Pretty.big_list "arities:"
   350     Pretty.writeln (Pretty.big_list "type arities:" (flat (map pretty_arities arities)));
   200       (List.concat (map pretty_arities arities)));
   351     Pretty.writeln (Pretty.big_list "consts:" (map pretty_const (Symtab.dest const_tab)))
   201     Pretty.writeln (Pretty.big_list "consts:"
       
   202       (map (pretty_const syn) (Symtab.dest const_tab)))
       
   203   end;
   352   end;
   204 
   353 
   205 
   354 
   206 fun pretty_sg (Sg {stamps, ...}) =
   355 fun pretty_sg (Sg {stamps, ...}) =
   207   Pretty.str_list "{" "}" (stamp_names stamps);
   356   Pretty.str_list "{" "}" (stamp_names stamps);
   210 
   359 
   211 
   360 
   212 
   361 
   213 (** pretty printing of terms and types **)
   362 (** pretty printing of terms and types **)
   214 
   363 
   215 fun pretty_term (Sg {syn, stamps, ...}) =
   364 fun pretty_term (Sg {syn, spaces, stamps, ...}) t =
   216   let val curried = "CPure" mem_string (map ! stamps);
   365   Syntax.pretty_term syn
   217   in Syntax.pretty_term curried syn end;
   366     ("CPure" mem_string (map ! stamps))
   218 
   367     (if ! long_names then t else extrn_term spaces t);
   219 fun pretty_typ (Sg {syn, ...}) = Syntax.pretty_typ syn;
   368 
   220 
   369 fun pretty_typ (Sg {syn, spaces, ...}) T =
   221 fun string_of_term (Sg {syn, stamps, ...}) =
   370   Syntax.pretty_typ syn
   222   let val curried = "CPure" mem_string (map ! stamps);
   371     (if ! long_names then T else extrn_typ spaces T);
   223   in Syntax.string_of_term curried syn end;
   372 
   224 
   373 fun pretty_sort (Sg {syn, spaces, ...}) S =
   225 fun string_of_typ (Sg {syn, ...}) = Syntax.string_of_typ syn;
   374   Syntax.pretty_sort syn
       
   375     (if ! long_names then S else extrn_sort spaces S);
       
   376 
       
   377 
       
   378 fun string_of_term sg t = Pretty.string_of (pretty_term sg t);
       
   379 fun string_of_typ sg T = Pretty.string_of (pretty_typ sg T);
       
   380 fun string_of_sort sg S = Pretty.string_of (pretty_sort sg S);
   226 
   381 
   227 fun pprint_term sg = Pretty.pprint o Pretty.quote o (pretty_term sg);
   382 fun pprint_term sg = Pretty.pprint o Pretty.quote o (pretty_term sg);
   228 fun pprint_typ sg = Pretty.pprint o Pretty.quote o (pretty_typ sg);
   383 fun pprint_typ sg = Pretty.pprint o Pretty.quote o (pretty_typ sg);
   229 
   384 
   230 
   385 
   232 (** read types **)  (*exception ERROR*)
   387 (** read types **)  (*exception ERROR*)
   233 
   388 
   234 fun err_in_type s =
   389 fun err_in_type s =
   235   error ("The error(s) above occurred in type " ^ quote s);
   390   error ("The error(s) above occurred in type " ^ quote s);
   236 
   391 
   237 fun read_raw_typ syn tsig sort_of str =
   392 fun read_raw_typ syn tsig spaces def_sort str =
   238   Syntax.read_typ syn (Type.eq_sort tsig) (Type.get_sort tsig sort_of) str
   393   intrn_tycons spaces
   239     handle ERROR => err_in_type str;
   394     (Syntax.read_typ syn (Type.get_sort tsig def_sort (intrn_sort spaces)) str
   240 
   395       handle ERROR => err_in_type str);
       
   396   
   241 (*read and certify typ wrt a signature*)
   397 (*read and certify typ wrt a signature*)
   242 fun read_typ (Sg {tsig, syn, ...}, sort_of) str =
   398 fun read_typ (Sg {tsig, syn, spaces, ...}, def_sort) str =
   243   Type.cert_typ tsig (read_raw_typ syn tsig sort_of str)
   399   Type.cert_typ tsig (read_raw_typ syn tsig spaces def_sort str)
   244     handle TYPE (msg, _, _) => (error_msg msg; err_in_type str);
   400     handle TYPE (msg, _, _) => (error_msg msg; err_in_type str);
   245 
   401 
   246 
   402 
   247 
   403 
   248 (** certify types and terms **)   (*exception TYPE*)
   404 (** certify types and terms **)   (*exception TYPE*)
   256   | TFree _ => tvars
   412   | TFree _ => tvars
   257   | TVar (v as (a, S)) =>
   413   | TVar (v as (a, S)) =>
   258       (case assoc_string_int (tvars, a) of
   414       (case assoc_string_int (tvars, a) of
   259         Some S' =>
   415         Some S' =>
   260           if S = S' then tvars
   416           if S = S' then tvars
   261           else raise_type ("Type variable " ^ Syntax.string_of_vname a ^
   417           else raise TYPE ("Type variable " ^ Syntax.string_of_vname a ^
   262             " has two distinct sorts") [TVar (a, S'), T] []
   418             " has two distinct sorts", [TVar (a, S'), T], [])
   263       | None => v :: tvars))
   419       | None => v :: tvars))
   264 (*equivalent to foldl nodup_TVars_list, but 3X faster under Poly/ML*)
   420 (*equivalent to foldl nodup_TVars_list, but 3X faster under Poly/ML*)
   265 and nodup_TVars_list (tvars, []) = tvars
   421 and nodup_TVars_list (tvars, []) = tvars
   266   | nodup_TVars_list (tvars, T :: Ts) =
   422   | nodup_TVars_list (tvars, T :: Ts) =
   267       nodup_TVars_list (nodup_TVars (tvars, T), Ts);
   423       nodup_TVars_list (nodup_TVars (tvars, T), Ts);
   275       | Free (a, T) => (vars, nodup_TVars (tvars, T))
   431       | Free (a, T) => (vars, nodup_TVars (tvars, T))
   276       | Var (v as (ixn, T)) =>
   432       | Var (v as (ixn, T)) =>
   277           (case assoc_string_int (vars, ixn) of
   433           (case assoc_string_int (vars, ixn) of
   278             Some T' =>
   434             Some T' =>
   279               if T = T' then (vars, nodup_TVars (tvars, T))
   435               if T = T' then (vars, nodup_TVars (tvars, T))
   280               else raise_type ("Variable " ^ Syntax.string_of_vname ixn ^
   436               else raise TYPE ("Variable " ^ Syntax.string_of_vname ixn ^
   281                 " has two distinct types") [T', T] []
   437                 " has two distinct types", [T', T], [])
   282           | None => (v :: vars, tvars))
   438           | None => (v :: vars, tvars))
   283       | Bound _ => (vars, tvars)
   439       | Bound _ => (vars, tvars)
   284       | Abs(_, T, t) => nodups vars (nodup_TVars (tvars, T)) t
   440       | Abs (_, T, t) => nodups vars (nodup_TVars (tvars, T)) t
   285       | s $ t =>
   441       | s $ t =>
   286           let val (vars',tvars') = nodups vars tvars s in
   442           let val (vars',tvars') = nodups vars tvars s in
   287             nodups vars' tvars' t
   443             nodups vars' tvars' t
   288           end);
   444           end);
   289   in nodups [] [] tm; () end;
   445   in nodups [] [] tm; () end;
   310       | atom_err _ = None;
   466       | atom_err _ = None;
   311 
   467 
   312     val norm_tm =
   468     val norm_tm =
   313       (case it_term_types (Type.typ_errors tsig) (tm, []) of
   469       (case it_term_types (Type.typ_errors tsig) (tm, []) of
   314         [] => map_term_types (Type.norm_typ tsig) tm
   470         [] => map_term_types (Type.norm_typ tsig) tm
   315       | errs => raise_type (cat_lines errs) [] [tm]);
   471       | errs => raise TYPE (cat_lines errs, [], [tm]));
   316     val _ = nodup_Vars norm_tm;
   472     val _ = nodup_Vars norm_tm;
   317   in
   473   in
   318     (case mapfilt_atoms atom_err norm_tm of
   474     (case mapfilt_atoms atom_err norm_tm of
   319       [] => (norm_tm, type_of norm_tm, maxidx_of_term norm_tm)
   475       [] => (norm_tm, type_of norm_tm, maxidx_of_term norm_tm)
   320     | errs => raise_type (cat_lines errs) [] [norm_tm])
   476     | errs => raise TYPE (cat_lines errs, [], [norm_tm]))
   321   end;
   477   end;
   322 
   478 
   323 
   479 
   324 
   480 
   325 (** infer_types **)         (*exception ERROR*)
   481 (** infer_types **)         (*exception ERROR*)
   337 fun infer_types sg def_type def_sort used freeze (ts, T) =
   493 fun infer_types sg def_type def_sort used freeze (ts, T) =
   338   let
   494   let
   339     val Sg {tsig, ...} = sg;
   495     val Sg {tsig, ...} = sg;
   340     val prt = setmp Syntax.show_brackets true (pretty_term sg);
   496     val prt = setmp Syntax.show_brackets true (pretty_term sg);
   341     val prT = pretty_typ sg;
   497     val prT = pretty_typ sg;
   342     val infer = Type.infer_types prt prT tsig (const_type sg)
   498     val infer = Type.infer_types prt prT tsig (const_type sg) def_type def_sort
   343       def_type def_sort used freeze;
   499       (intern_const sg) (intern_tycons sg) (intern_sort sg) used freeze;
   344 
   500 
   345     val T' = certify_typ sg T handle TYPE (msg, _, _) => error msg;
   501     val T' = certify_typ sg T handle TYPE (msg, _, _) => error msg;
   346 
   502 
   347     fun warn () =
   503     fun warn () =
   348       if length ts > 1 andalso length ts <= ! Syntax.ambiguity_level
   504       if length ts > 1 andalso length ts <= ! Syntax.ambiguity_level
   374           warning "Fortunately, only one parse tree is type correct.\n\
   530           warning "Fortunately, only one parse tree is type correct.\n\
   375             \It helps (speed!) if you disambiguate your grammar or your input."
   531             \It helps (speed!) if you disambiguate your grammar or your input."
   376         else (); res)
   532         else (); res)
   377     | Errs errs => (warn (); error (cat_lines errs))
   533     | Errs errs => (warn (); error (cat_lines errs))
   378     | Ambigs us =>
   534     | Ambigs us =>
   379         (warn (); error ("Error: More than one term is type correct:\n" ^
   535         (warn (); error ("More than one term is type correct:\n" ^
   380           (cat_lines (map (Pretty.string_of o prt) us)))))
   536           (cat_lines (map (Pretty.string_of o prt) us)))))
   381   end;
   537   end;
   382 
   538 
   383 
   539 
   384 
   540 
   385 (** extend signature **)    (*exception ERROR*)
   541 (** extend signature **)    (*exception ERROR*)
   386 
   542 
   387 (** signature extension functions **)  (*exception ERROR*)
   543 (** signature extension functions **)  (*exception ERROR*)
   388 
   544 
   389 fun decls_of name_of mfixs =
   545 fun decls_of path name_of mfixs =
   390   map (fn (x, y, mx) => (name_of x mx, y)) mfixs;
   546   map (fn (x, y, mx) => (full_name path (name_of x mx), y)) mfixs;
       
   547 
       
   548 fun no_read _ _ _ decl = decl;
   391 
   549 
   392 
   550 
   393 (* add default sort *)
   551 (* add default sort *)
   394 
   552 
   395 fun ext_defsort (syn, tsig, ctab) defsort =
   553 fun ext_defsort int (syn, tsig, ctab, (path, spaces)) S =
   396   (syn, Type.ext_tsig_defsort tsig defsort, ctab);
   554   (syn, Type.ext_tsig_defsort tsig (if int then intrn_sort spaces S else S),
       
   555     ctab, (path, spaces));
   397 
   556 
   398 
   557 
   399 (* add type constructors *)
   558 (* add type constructors *)
   400 
   559 
   401 fun ext_types (syn, tsig, ctab) types =
   560 fun ext_types (syn, tsig, ctab, (path, spaces)) types =
   402   (Syntax.extend_type_gram syn types,
   561   let val decls = decls_of path Syntax.type_name types in
   403     Type.ext_tsig_types tsig (decls_of Syntax.type_name types),
   562     (Syntax.extend_type_gram syn types,
   404     ctab);
   563       Type.ext_tsig_types tsig decls, ctab,
       
   564       (path, add_names spaces typeK (map fst decls)))
       
   565   end;
   405 
   566 
   406 
   567 
   407 (* add type abbreviations *)
   568 (* add type abbreviations *)
   408 
   569 
   409 fun read_abbr syn tsig (t, vs, rhs_src) =
   570 fun read_abbr syn tsig spaces (t, vs, rhs_src) =
   410   (t, vs, read_raw_typ syn tsig (K None) rhs_src)
   571   (t, vs, read_raw_typ syn tsig spaces (K None) rhs_src)
   411     handle ERROR => error ("in type abbreviation " ^ t);
   572     handle ERROR => error ("in type abbreviation " ^ t);
   412 
   573 
   413 fun ext_abbrs rd_abbr (syn, tsig, ctab) abbrs =
   574 fun ext_abbrs rd_abbr (syn, tsig, ctab, (path, spaces)) abbrs =
   414   let
   575   let
   415     fun mfix_of (t, vs, _, mx) = (t, length vs, mx);
   576     fun mfix_of (t, vs, _, mx) = (t, length vs, mx);
   416     val syn1 = Syntax.extend_type_gram syn (map mfix_of abbrs);
   577     val syn' = Syntax.extend_type_gram syn (map mfix_of abbrs);
   417 
   578 
   418     fun decl_of (t, vs, rhs, mx) =
   579     val abbrs' =
   419       rd_abbr syn1 tsig (Syntax.type_name t mx, vs, rhs);
   580       map (fn (t, vs, rhs, mx) =>
   420   in
   581         (full_name path (Syntax.type_name t mx), vs, rhs)) abbrs;
   421     (syn1, Type.ext_tsig_abbrs tsig (map decl_of abbrs), ctab)
   582     val space' = add_names spaces typeK (map #1 abbrs');
   422   end;
   583     val decls = map (rd_abbr syn' tsig space') abbrs';
   423 
   584   in
   424 fun ext_tyabbrs_i arg = ext_abbrs (K (K I)) arg;
   585     (syn', Type.ext_tsig_abbrs tsig decls, ctab, (path, spaces))
   425 fun ext_tyabbrs   arg = ext_abbrs read_abbr arg;
   586   end;
       
   587 
       
   588 fun ext_tyabbrs abbrs = ext_abbrs read_abbr abbrs;
       
   589 fun ext_tyabbrs_i abbrs = ext_abbrs no_read abbrs;
   426 
   590 
   427 
   591 
   428 (* add type arities *)
   592 (* add type arities *)
   429 
   593 
   430 fun ext_arities (syn, tsig, ctab) arities =
   594 fun ext_arities int (syn, tsig, ctab, (path, spaces)) arities =
   431   let
   595   let
   432     val tsig1 = Type.ext_tsig_arities tsig arities;
   596     fun intrn_arity (c, Ss, S) =
   433     val log_types = Type.logical_types tsig1;
   597       (intrn spaces typeK c, map (intrn_sort spaces) Ss, intrn_sort spaces S);
   434   in
   598     val intrn = if int then map intrn_arity else I;
   435     (Syntax.extend_log_types syn log_types, tsig1, ctab)
   599     val tsig' = Type.ext_tsig_arities tsig (intrn arities);
       
   600     val log_types = Type.logical_types tsig';
       
   601   in
       
   602     (Syntax.extend_log_types syn log_types, tsig', ctab, (path, spaces))
   436   end;
   603   end;
   437 
   604 
   438 
   605 
   439 (* add term constants and syntax *)
   606 (* add term constants and syntax *)
   440 
   607 
   443 
   610 
   444 fun err_dup_consts cs =
   611 fun err_dup_consts cs =
   445   error ("Duplicate declaration of constant(s) " ^ commas_quote cs);
   612   error ("Duplicate declaration of constant(s) " ^ commas_quote cs);
   446 
   613 
   447 
   614 
   448 fun read_const syn tsig (c, ty_src, mx) =
   615 fun read_const syn tsig spaces (c, ty_src, mx) =
   449   (c, read_raw_typ syn tsig (K None) ty_src, mx)
   616   (c, read_raw_typ syn tsig spaces (K None) ty_src, mx)
   450     handle ERROR => err_in_const (Syntax.const_name c mx);
   617     handle ERROR => err_in_const (Syntax.const_name c mx);
   451 
   618 
   452 fun ext_cnsts rd_const syn_only prmode (syn, tsig, ctab) raw_consts =
   619 fun ext_cnsts rd_const syn_only prmode (syn, tsig, ctab, (path, spaces)) raw_consts =
   453   let
   620   let
   454     fun prep_const (c, ty, mx) =
   621     fun prep_const (c, ty, mx) =
   455      (c, compress_type (Type.varifyT (Type.cert_typ tsig (Type.no_tvars ty))), mx)
   622       (c, compress_type (Type.varifyT (Type.cert_typ tsig (Type.no_tvars ty))), mx)
   456        handle TYPE (msg, _, _)
   623         handle TYPE (msg, _, _) =>
   457          => (error_msg msg; err_in_const (Syntax.const_name c mx));
   624           (error_msg msg; err_in_const (Syntax.const_name c mx));
   458 
   625 
   459     val consts = map (prep_const o rd_const syn tsig) raw_consts;
   626     val consts = map (prep_const o rd_const syn tsig spaces) raw_consts;
   460     val decls =
   627     val decls =
   461       if syn_only then []
   628       if syn_only then []
   462       else filter_out (equal "" o fst) (decls_of Syntax.const_name consts);
   629       else decls_of path Syntax.const_name consts;
   463   in
   630   in
   464     (Syntax.extend_const_gram syn prmode consts, tsig,
   631     (Syntax.extend_const_gram syn prmode consts, tsig,
   465       Symtab.extend_new (ctab, decls)
   632       Symtab.extend_new (ctab, decls)
   466         handle Symtab.DUPS cs => err_dup_consts cs)
   633         handle Symtab.DUPS cs => err_dup_consts cs,
   467   end;
   634       (path, add_names spaces constK (map fst decls)))
   468 
   635   end;
   469 val ext_consts_i = ext_cnsts (K (K I)) false ("", true);
   636 
       
   637 val ext_consts_i = ext_cnsts no_read false ("", true);
   470 val ext_consts = ext_cnsts read_const false ("", true);
   638 val ext_consts = ext_cnsts read_const false ("", true);
   471 val ext_syntax_i = ext_cnsts (K (K I)) true ("", true);
   639 val ext_syntax_i = ext_cnsts no_read true ("", true);
   472 val ext_syntax = ext_cnsts read_const true ("", true);
   640 val ext_syntax = ext_cnsts read_const true ("", true);
   473 fun ext_modesyntax_i sg (prmode, consts) = ext_cnsts (K (K I)) true prmode sg consts;
   641 fun ext_modesyntax_i sg (prmode, consts) = ext_cnsts no_read true prmode sg consts;
   474 fun ext_modesyntax sg (prmode, consts) = ext_cnsts read_const true prmode sg consts;
   642 fun ext_modesyntax sg (prmode, consts) = ext_cnsts read_const true prmode sg consts;
   475 
   643 
   476 
   644 
   477 (* add type classes *)
   645 (* add type classes *)
   478 
   646 
   479 fun const_of_class c = c ^ "_class";
   647 fun const_of_class c = c ^ "_class";
   480 
   648 
   481 fun class_of_const c_class =
   649 fun class_of_const c_class =
   482   let
   650   let
   483     val c = implode (take (size c_class - 6, explode c_class));
   651     val c = implode (take (size c_class - size "_class", explode c_class));
   484   in
   652   in
   485     if const_of_class c = c_class then c
   653     if const_of_class c = c_class then c
   486     else raise_term ("class_of_const: bad name " ^ quote c_class) []
   654     else raise TERM ("class_of_const: bad name " ^ quote c_class, [])
   487   end;
   655   end;
   488 
   656 
   489 
   657 
   490 fun ext_classes (syn, tsig, ctab) classes =
   658 fun ext_classes int (syn, tsig, ctab, (path, spaces)) classes =
   491   let
   659   let
   492     val names = map fst classes;
   660     val names = map fst classes;
   493     val consts =
   661     val consts =
   494       map (fn c => (const_of_class c, a_itselfT --> propT, NoSyn)) names;
   662       map (fn c => (const_of_class c, a_itselfT --> propT, NoSyn)) names;
       
   663 
       
   664     val full_names = map (full_name path) names;
       
   665     val spaces' = add_names spaces classK full_names;
       
   666     val intrn = if int then map (intrn_class spaces') else I;
       
   667     val classes' =
       
   668       ListPair.map (fn (c, (_, cs)) => (c, intrn cs)) (full_names, classes);
   495   in
   669   in
   496     ext_consts_i
   670     ext_consts_i
   497       (Syntax.extend_consts syn names, Type.ext_tsig_classes tsig classes, ctab)
   671       (Syntax.extend_consts syn names,
   498       consts
   672         Type.ext_tsig_classes tsig classes', ctab, (path, spaces'))
       
   673     consts
   499   end;
   674   end;
   500 
   675 
   501 
   676 
   502 (* add to classrel *)
   677 (* add to classrel *)
   503 
   678 
   504 fun ext_classrel (syn, tsig, ctab) pairs =
   679 fun ext_classrel int (syn, tsig, ctab, (path, spaces)) pairs =
   505   (syn, Type.ext_tsig_classrel tsig pairs, ctab);
   680   let val intrn = if int then map (pairself (intrn_class spaces)) else I in
       
   681     (syn, Type.ext_tsig_classrel tsig (intrn pairs), ctab, (path, spaces))
       
   682   end;
   506 
   683 
   507 
   684 
   508 (* add to syntax *)
   685 (* add to syntax *)
   509 
   686 
   510 fun ext_syn extfun (syn, tsig, ctab) args =
   687 fun ext_syn extfun (syn, tsig, ctab, names) args =
   511   (extfun syn args, tsig, ctab);
   688   (extfun syn args, tsig, ctab, names);
       
   689 
       
   690 
       
   691 (* add to path *)
       
   692 
       
   693 fun ext_path (syn, tsig, ctab, (path, spaces)) elem =
       
   694   let
       
   695     val path' =
       
   696       if elem = ".." andalso not (null path) then fst (split_last path)
       
   697       else if elem = "/" then []
       
   698       else path @ NameSpace.unpack elem;
       
   699   in
       
   700     (syn, tsig, ctab, (path', spaces))
       
   701   end;      
       
   702 
       
   703 
       
   704 (* add to name space *)
       
   705 
       
   706 fun ext_space (syn, tsig, ctab, (path, spaces)) (kind, names) =
       
   707   (syn, tsig, ctab, (path, add_names spaces kind names));
   512 
   708 
   513 
   709 
   514 (* build signature *)
   710 (* build signature *)
   515 
   711 
   516 fun ext_stamps stamps name =
   712 fun ext_stamps stamps name =
   520     if exists (equal name o !) stmps then
   716     if exists (equal name o !) stmps then
   521       error ("Theory already contains a " ^ quote name ^ " component")
   717       error ("Theory already contains a " ^ quote name ^ " component")
   522     else ref name :: stmps
   718     else ref name :: stmps
   523   end;
   719   end;
   524 
   720 
   525 fun make_sign (syn, tsig, ctab) stamps name =
   721 fun make_sign (syn, tsig, ctab, (path, spaces)) stamps name =
   526   Sg {tsig = tsig, const_tab = ctab, syn = syn,
   722   Sg {tsig = tsig, const_tab = ctab, syn = syn, path = path, spaces = spaces,
   527     stamps = ext_stamps stamps name};
   723     stamps = ext_stamps stamps name};
   528 
   724 
   529 fun extend_sign extfun name decls (Sg {tsig, const_tab, syn, stamps}) =
   725 fun extend_sign extfun name decls (Sg {tsig, const_tab, syn, path, spaces, stamps}) =
   530   make_sign (extfun (syn, tsig, const_tab) decls) stamps name;
   726   make_sign (extfun (syn, tsig, const_tab, (path, spaces)) decls) stamps name;
   531 
   727 
   532 
   728 
   533 (* the external interfaces *)
   729 (* the external interfaces *)
   534 
   730 
   535 val add_classes      = extend_sign ext_classes "#";
   731 val add_classes      = extend_sign (ext_classes true) "#";
   536 val add_classrel     = extend_sign ext_classrel "#";
   732 val add_classes_i    = extend_sign (ext_classes false) "#";
   537 val add_defsort      = extend_sign ext_defsort "#";
   733 val add_classrel     = extend_sign (ext_classrel true) "#";
       
   734 val add_classrel_i   = extend_sign (ext_classrel false) "#";
       
   735 val add_defsort      = extend_sign (ext_defsort true) "#";
       
   736 val add_defsort_i    = extend_sign (ext_defsort false) "#";
   538 val add_types        = extend_sign ext_types "#";
   737 val add_types        = extend_sign ext_types "#";
   539 val add_tyabbrs      = extend_sign ext_tyabbrs "#";
   738 val add_tyabbrs      = extend_sign ext_tyabbrs "#";
   540 val add_tyabbrs_i    = extend_sign ext_tyabbrs_i "#";
   739 val add_tyabbrs_i    = extend_sign ext_tyabbrs_i "#";
   541 val add_arities      = extend_sign ext_arities "#";
   740 val add_arities      = extend_sign (ext_arities true) "#";
       
   741 val add_arities_i    = extend_sign (ext_arities false) "#";
   542 val add_consts       = extend_sign ext_consts "#";
   742 val add_consts       = extend_sign ext_consts "#";
   543 val add_consts_i     = extend_sign ext_consts_i "#";
   743 val add_consts_i     = extend_sign ext_consts_i "#";
   544 val add_syntax       = extend_sign ext_syntax "#";
   744 val add_syntax       = extend_sign ext_syntax "#";
   545 val add_syntax_i     = extend_sign ext_syntax_i "#";
   745 val add_syntax_i     = extend_sign ext_syntax_i "#";
   546 val add_modesyntax   = extend_sign ext_modesyntax "#";
   746 val add_modesyntax   = extend_sign ext_modesyntax "#";
   548 val add_trfuns       = extend_sign (ext_syn Syntax.extend_trfuns) "#";
   748 val add_trfuns       = extend_sign (ext_syn Syntax.extend_trfuns) "#";
   549 val add_trfunsT      = extend_sign (ext_syn Syntax.extend_trfunsT) "#";
   749 val add_trfunsT      = extend_sign (ext_syn Syntax.extend_trfunsT) "#";
   550 val add_tokentrfuns  = extend_sign (ext_syn Syntax.extend_tokentrfuns) "#";
   750 val add_tokentrfuns  = extend_sign (ext_syn Syntax.extend_tokentrfuns) "#";
   551 val add_trrules      = extend_sign (ext_syn Syntax.extend_trrules) "#";
   751 val add_trrules      = extend_sign (ext_syn Syntax.extend_trrules) "#";
   552 val add_trrules_i    = extend_sign (ext_syn Syntax.extend_trrules_i) "#";
   752 val add_trrules_i    = extend_sign (ext_syn Syntax.extend_trrules_i) "#";
   553 
   753 val add_path         = extend_sign ext_path "#";
       
   754 val add_space        = extend_sign ext_space "#";
   554 fun add_name name sg = extend_sign K name () sg;
   755 fun add_name name sg = extend_sign K name () sg;
       
   756 
   555 val make_draft = add_name "#";
   757 val make_draft = add_name "#";
   556 
   758 
   557 
   759 
   558 
   760 
   559 (** merge signatures **)    (*exception TERM*) (*exception ERROR (* FIXME *)*)
   761 (** merge signatures **)    (*exception TERM*) (*exception ERROR (* FIXME *) handle_error? *)
   560 
   762 
   561 fun merge (sg1, sg2) =
   763 fun merge (sg1, sg2) =
   562   if fast_subsig (sg2, sg1) then sg1
   764   if fast_subsig (sg2, sg1) then sg1
   563   else if fast_subsig (sg1, sg2) then sg2
   765   else if fast_subsig (sg1, sg2) then sg2
   564   else if subsig (sg2, sg1) then sg1
   766   else if subsig (sg2, sg1) then sg1
   565   else if subsig (sg1, sg2) then sg2
   767   else if subsig (sg1, sg2) then sg2
   566   else if is_draft sg1 orelse is_draft sg2 then
   768   else if is_draft sg1 orelse is_draft sg2 then
   567     raise_term "Illegal merge of draft signatures" []
   769     raise TERM ("Illegal merge of draft signatures", [])
   568   else
   770   else
   569     (*neither is union already; must form union*)
   771     (*neither is union already; must form union*)
   570     let
   772     let
   571       val Sg {tsig = tsig1, const_tab = const_tab1, syn = syn1,
   773       val Sg {tsig = tsig1, const_tab = const_tab1, syn = syn1,
   572         stamps = stamps1} = sg1;
   774         path, spaces = spaces1, stamps = stamps1} = sg1;
   573       val Sg {tsig = tsig2, const_tab = const_tab2, syn = syn2,
   775       val Sg {tsig = tsig2, const_tab = const_tab2, syn = syn2,
   574         stamps = stamps2} = sg2;
   776         path = _, spaces = spaces2, stamps = stamps2} = sg2;
   575 
   777 
   576 
   778 
   577       val stamps = merge_rev_lists stamps1 stamps2;
   779       val stamps = merge_rev_lists stamps1 stamps2;
   578       val _ =
   780       val _ =
   579         (case duplicates (stamp_names stamps) of
   781         (case duplicates (stamp_names stamps) of
   580           [] => ()
   782           [] => ()
   581         | dups => raise_term ("Attempt to merge different versions of theories "
   783         | dups => raise TERM ("Attempt to merge different versions of theories "
   582             ^ commas_quote dups) []);
   784             ^ commas_quote dups, []));
   583 
   785 
   584       val tsig = Type.merge_tsigs (tsig1, tsig2);
   786       val tsig = Type.merge_tsigs (tsig1, tsig2);
   585 
   787 
   586       val const_tab = Symtab.merge (op =) (const_tab1, const_tab2)
   788       val const_tab = Symtab.merge (op =) (const_tab1, const_tab2)
   587         handle Symtab.DUPS cs =>
   789         handle Symtab.DUPS cs =>
   588           raise_term ("Incompatible types for constant(s) " ^ commas_quote cs) [];
   790           raise TERM ("Incompatible types for constant(s) " ^ commas_quote cs, []);
   589 
   791 
   590       val syn = Syntax.merge_syntaxes syn1 syn2;
   792       val syn = Syntax.merge_syntaxes syn1 syn2;
       
   793 
       
   794       val kinds = distinct (map fst (spaces1 @ spaces2));
       
   795       val spaces =
       
   796         kinds ~~
       
   797           ListPair.map NameSpace.merge
       
   798             (map (space_of spaces1) kinds, map (space_of spaces2) kinds);
   591     in
   799     in
   592       Sg {tsig = tsig, const_tab = const_tab, syn = syn, stamps = stamps}
   800       Sg {tsig = tsig, const_tab = const_tab, syn = syn,
       
   801         path = path, spaces = spaces, stamps = stamps}
   593     end;
   802     end;
   594 
   803 
   595 
   804 
   596 
   805 
   597 (** the Pure signature **)
   806 (** the Pure signature **)
   598 
   807 
   599 val proto_pure =
   808 val proto_pure =
   600   make_sign (Syntax.pure_syn, Type.tsig0, Symtab.null) [] "#"
   809   make_sign (Syntax.pure_syn, Type.tsig0, Symtab.null, ([], [])) [] "#"
   601   |> add_types
   810   |> add_types
   602    (("fun", 2, NoSyn) ::
   811    (("fun", 2, NoSyn) ::
   603     ("prop", 0, NoSyn) ::
   812     ("prop", 0, NoSyn) ::
   604     ("itself", 1, NoSyn) ::
   813     ("itself", 1, NoSyn) ::
   605     Syntax.pure_types)
   814     Syntax.pure_types)
   606   |> add_classes [(logicC, [])]
   815   |> add_classes_i [(logicC, [])]
   607   |> add_defsort logicS
   816   |> add_defsort_i logicS
   608   |> add_arities
   817   |> add_arities_i
   609    [("fun", [logicS, logicS], logicS),
   818    [("fun", [logicS, logicS], logicS),
   610     ("prop", [], logicS),
   819     ("prop", [], logicS),
   611     ("itself", [logicS], logicS)]
   820     ("itself", [logicS], logicS)]
   612   |> add_syntax Syntax.pure_syntax
   821   |> add_syntax Syntax.pure_syntax
   613   |> add_modesyntax (("symbols", true), Syntax.pure_sym_syntax)
   822   |> add_modesyntax (("symbols", true), Syntax.pure_sym_syntax)
   629 
   838 
   630 val cpure = proto_pure
   839 val cpure = proto_pure
   631   |> add_syntax Syntax.pure_applC_syntax
   840   |> add_syntax Syntax.pure_applC_syntax
   632   |> add_name "CPure";
   841   |> add_name "CPure";
   633 
   842 
       
   843 
   634 end;
   844 end;
       
   845 
       
   846 
       
   847 val long_names = Sign.long_names;