| 256 |      1 | (*  Title:      Pure/type.ML
 | 
| 0 |      2 |     ID:         $Id$
 | 
| 416 |      3 |     Author:     Tobias Nipkow & Lawrence C Paulson
 | 
| 0 |      4 | 
 | 
| 416 |      5 | Type classes and sorts. Type signatures. Type unification and inference.
 | 
| 256 |      6 | 
 | 
|  |      7 | TODO:
 | 
| 416 |      8 |   move type unification and inference to type_unify.ML (TypeUnify) (?)
 | 
|  |      9 |   rename args -> tycons, coreg -> arities (?)
 | 
| 0 |     10 | *)
 | 
|  |     11 | 
 | 
|  |     12 | signature TYPE =
 | 
|  |     13 | sig
 | 
| 256 |     14 |   structure Symtab: SYMTAB
 | 
| 621 |     15 |   val no_tvars: typ -> typ
 | 
|  |     16 |   val varifyT: typ -> typ
 | 
|  |     17 |   val unvarifyT: typ -> typ
 | 
|  |     18 |   val varify: term * string list -> term
 | 
| 416 |     19 |   val str_of_sort: sort -> string
 | 
|  |     20 |   val str_of_arity: string * sort list * sort -> string
 | 
| 0 |     21 |   type type_sig
 | 
| 200 |     22 |   val rep_tsig: type_sig ->
 | 
| 256 |     23 |     {classes: class list,
 | 
|  |     24 |      subclass: (class * class list) list,
 | 
|  |     25 |      default: sort,
 | 
|  |     26 |      args: (string * int) list,
 | 
| 621 |     27 |      abbrs: (string * (string list * typ)) list,
 | 
| 256 |     28 |      coreg: (string * (class * sort list) list) list}
 | 
| 0 |     29 |   val defaultS: type_sig -> sort
 | 
| 416 |     30 |   val tsig0: type_sig
 | 
| 256 |     31 |   val logical_types: type_sig -> string list
 | 
| 621 |     32 |   val ext_tsig_classes: type_sig -> (class * class list) list -> type_sig
 | 
| 422 |     33 |   val ext_tsig_subclass: type_sig -> (class * class) list -> type_sig
 | 
|  |     34 |   val ext_tsig_defsort: type_sig -> sort -> type_sig
 | 
| 582 |     35 |   val ext_tsig_types: type_sig -> (string * int) list -> type_sig
 | 
| 621 |     36 |   val ext_tsig_abbrs: type_sig -> (string * string list * typ) list -> type_sig
 | 
|  |     37 |   val ext_tsig_arities: type_sig -> (string * sort list * sort) list -> type_sig
 | 
| 256 |     38 |   val merge_tsigs: type_sig * type_sig -> type_sig
 | 
| 416 |     39 |   val subsort: type_sig -> sort * sort -> bool
 | 
|  |     40 |   val norm_sort: type_sig -> sort -> sort
 | 
|  |     41 |   val rem_sorts: typ -> typ
 | 
| 256 |     42 |   val cert_typ: type_sig -> typ -> typ
 | 
|  |     43 |   val norm_typ: type_sig -> typ -> typ
 | 
| 0 |     44 |   val freeze: (indexname -> bool) -> term -> term
 | 
|  |     45 |   val freeze_vars: typ -> typ
 | 
| 565 |     46 |   val infer_types: type_sig * (string -> typ option) * (indexname -> typ option) *
 | 
| 256 |     47 |     (indexname -> sort option) * typ * term -> term * (indexname * typ) list
 | 
|  |     48 |   val inst_term_tvars: type_sig * (indexname * typ) list -> term -> term
 | 
| 0 |     49 |   val thaw_vars: typ -> typ
 | 
| 256 |     50 |   val typ_errors: type_sig -> typ * string list -> string list
 | 
| 0 |     51 |   val typ_instance: type_sig * typ * typ -> bool
 | 
| 256 |     52 |   val typ_match: type_sig -> (indexname * typ) list * (typ * typ)
 | 
|  |     53 |     -> (indexname * typ) list
 | 
|  |     54 |   val unify: type_sig -> (typ * typ) * (indexname * typ) list
 | 
|  |     55 |     -> (indexname * typ) list
 | 
| 450 |     56 |   val raw_unify: typ * typ -> bool
 | 
| 0 |     57 |   exception TUNIFY
 | 
| 256 |     58 |   exception TYPE_MATCH
 | 
| 0 |     59 | end;
 | 
|  |     60 | 
 | 
| 416 |     61 | functor TypeFun(structure Symtab: SYMTAB and Syntax: SYNTAX): TYPE =
 | 
| 0 |     62 | struct
 | 
|  |     63 | 
 | 
| 256 |     64 | structure Symtab = Symtab;
 | 
| 0 |     65 | 
 | 
|  |     66 | 
 | 
| 621 |     67 | (*** TFrees vs TVars ***)
 | 
|  |     68 | 
 | 
|  |     69 | (*disallow TVars*)
 | 
|  |     70 | fun no_tvars T =
 | 
|  |     71 |   if null (typ_tvars T) then T
 | 
|  |     72 |   else raise_type "Illegal schematic type variable(s)" [T] [];
 | 
|  |     73 | 
 | 
|  |     74 | (*turn TFrees into TVars to allow types & axioms to be written without "?"*)
 | 
|  |     75 | fun varifyT (Type (a, Ts)) = Type (a, map varifyT Ts)
 | 
|  |     76 |   | varifyT (TFree (a, S)) = TVar ((a, 0), S)
 | 
|  |     77 |   | varifyT T = T;
 | 
|  |     78 | 
 | 
|  |     79 | (*inverse of varifyT*)
 | 
|  |     80 | fun unvarifyT (Type (a, Ts)) = Type (a, map unvarifyT Ts)
 | 
|  |     81 |   | unvarifyT (TVar ((a, 0), S)) = TFree (a, S)
 | 
|  |     82 |   | unvarifyT T = T;
 | 
|  |     83 | 
 | 
|  |     84 | (*turn TFrees except those in fixed into new TVars*)
 | 
|  |     85 | fun varify (t, fixed) =
 | 
|  |     86 |   let
 | 
|  |     87 |     val fs = add_term_tfree_names (t, []) \\ fixed;
 | 
|  |     88 |     val ixns = add_term_tvar_ixns (t, []);
 | 
|  |     89 |     val fmap = fs ~~ variantlist (fs, map #1 ixns)
 | 
|  |     90 |     fun thaw (Type(a, Ts)) = Type (a, map thaw Ts)
 | 
|  |     91 |       | thaw (T as TVar _) = T
 | 
|  |     92 |       | thaw (T as TFree(a, S)) =
 | 
|  |     93 |           (case assoc (fmap, a) of None => T | Some b => TVar((b, 0), S))
 | 
|  |     94 |   in
 | 
|  |     95 |     map_term_types thaw t
 | 
|  |     96 |   end;
 | 
|  |     97 | 
 | 
|  |     98 | 
 | 
|  |     99 | 
 | 
| 416 |    100 | (*** type classes and sorts ***)
 | 
|  |    101 | 
 | 
|  |    102 | (*
 | 
|  |    103 |   Classes denote (possibly empty) collections of types (e.g. sets of types)
 | 
|  |    104 |   and are partially ordered by 'inclusion'. They are represented by strings.
 | 
|  |    105 | 
 | 
|  |    106 |   Sorts are intersections of finitely many classes. They are represented by
 | 
|  |    107 |   lists of classes.
 | 
|  |    108 | *)
 | 
| 0 |    109 | 
 | 
|  |    110 | type domain = sort list;
 | 
| 416 |    111 | 
 | 
|  |    112 | 
 | 
|  |    113 | (* print sorts and arities *)
 | 
| 0 |    114 | 
 | 
| 416 |    115 | fun str_of_sort [c] = c
 | 
| 565 |    116 |   | str_of_sort cs = enclose "{" "}" (commas cs);
 | 
| 416 |    117 | 
 | 
| 565 |    118 | fun str_of_dom dom = enclose "(" ")" (commas (map str_of_sort dom));
 | 
| 416 |    119 | 
 | 
|  |    120 | fun str_of_arity (t, [], S) = t ^ " :: " ^ str_of_sort S
 | 
|  |    121 |   | str_of_arity (t, SS, S) =
 | 
|  |    122 |       t ^ " :: " ^ str_of_dom SS ^ " " ^ str_of_sort S;
 | 
| 256 |    123 | 
 | 
|  |    124 | 
 | 
|  |    125 | 
 | 
| 416 |    126 | (*** type signatures ***)
 | 
| 256 |    127 | 
 | 
|  |    128 | (*
 | 
|  |    129 |   classes:
 | 
|  |    130 |     a list of all declared classes;
 | 
| 0 |    131 | 
 | 
| 256 |    132 |   subclass:
 | 
| 416 |    133 |     an association list representing the subclass relation; (c, cs) is
 | 
| 256 |    134 |     interpreted as "c is a proper subclass of all elemenst of cs"; note that
 | 
|  |    135 |     c itself is not a memeber of cs;
 | 
|  |    136 | 
 | 
|  |    137 |   default:
 | 
|  |    138 |     the default sort attached to all unconstrained type vars;
 | 
|  |    139 | 
 | 
|  |    140 |   args:
 | 
|  |    141 |     an association list of all declared types with the number of their
 | 
|  |    142 |     arguments;
 | 
|  |    143 | 
 | 
|  |    144 |   abbrs:
 | 
|  |    145 |     an association list of type abbreviations;
 | 
|  |    146 | 
 | 
|  |    147 |   coreg:
 | 
|  |    148 |     a two-fold association list of all type arities; (t, al) means that type
 | 
|  |    149 |     constructor t has the arities in al; an element (c, ss) of al represents
 | 
|  |    150 |     the arity (ss)c;
 | 
| 0 |    151 | *)
 | 
|  |    152 | 
 | 
| 256 |    153 | datatype type_sig =
 | 
|  |    154 |   TySg of {
 | 
|  |    155 |     classes: class list,
 | 
|  |    156 |     subclass: (class * class list) list,
 | 
|  |    157 |     default: sort,
 | 
|  |    158 |     args: (string * int) list,
 | 
| 621 |    159 |     abbrs: (string * (string list * typ)) list,
 | 
| 256 |    160 |     coreg: (string * (class * domain) list) list};
 | 
|  |    161 | 
 | 
| 189 |    162 | fun rep_tsig (TySg comps) = comps;
 | 
| 0 |    163 | 
 | 
| 256 |    164 | fun defaultS (TySg {default, ...}) = default;
 | 
|  |    165 | 
 | 
|  |    166 | 
 | 
| 582 |    167 | (* error messages *)
 | 
| 256 |    168 | 
 | 
| 416 |    169 | fun undcl_class c = "Undeclared class " ^ quote c;
 | 
| 256 |    170 | val err_undcl_class = error o undcl_class;
 | 
| 0 |    171 | 
 | 
| 422 |    172 | fun err_dup_classes cs =
 | 
|  |    173 |   error ("Duplicate declaration of class(es) " ^ commas_quote cs);
 | 
| 416 |    174 | 
 | 
|  |    175 | fun undcl_type c = "Undeclared type constructor " ^ quote c;
 | 
| 256 |    176 | val err_undcl_type = error o undcl_type;
 | 
|  |    177 | 
 | 
| 582 |    178 | fun err_neg_args c =
 | 
|  |    179 |   error ("Negative number of arguments of type constructor " ^ quote c);
 | 
|  |    180 | 
 | 
| 416 |    181 | fun err_dup_tycon c =
 | 
|  |    182 |   error ("Duplicate declaration of type constructor " ^ quote c);
 | 
|  |    183 | 
 | 
| 621 |    184 | fun dup_tyabbrs ts =
 | 
|  |    185 |   "Duplicate declaration of type abbreviation(s) " ^ commas_quote ts;
 | 
| 416 |    186 | 
 | 
|  |    187 | fun ty_confl c = "Conflicting type constructor and abbreviation " ^ quote c;
 | 
|  |    188 | val err_ty_confl = error o ty_confl;
 | 
| 0 |    189 | 
 | 
|  |    190 | 
 | 
|  |    191 | (* 'leq' checks the partial order on classes according to the
 | 
| 621 |    192 |    statements in the association list 'a' (i.e. 'subclass')
 | 
| 0 |    193 | *)
 | 
|  |    194 | 
 | 
| 256 |    195 | fun less a (C, D) = case assoc (a, C) of
 | 
| 621 |    196 |      Some ss => D mem ss
 | 
|  |    197 |    | None => err_undcl_class C;
 | 
| 0 |    198 | 
 | 
| 256 |    199 | fun leq a (C, D)  =  C = D orelse less a (C, D);
 | 
| 0 |    200 | 
 | 
|  |    201 | 
 | 
| 416 |    202 | (* logical_types *)
 | 
| 0 |    203 | 
 | 
| 416 |    204 | (*return all logical types of tsig, i.e. all types t with some arity t::(ss)c
 | 
|  |    205 |   and c <= logic*)
 | 
| 0 |    206 | 
 | 
| 416 |    207 | fun logical_types tsig =
 | 
|  |    208 |   let
 | 
|  |    209 |     val TySg {subclass, coreg, args, ...} = tsig;
 | 
|  |    210 | 
 | 
|  |    211 |     fun log_class c = leq subclass (c, logicC);
 | 
|  |    212 |     fun log_type t = exists (log_class o #1) (assocs coreg t);
 | 
|  |    213 |   in
 | 
|  |    214 |     filter log_type (map #1 args)
 | 
| 0 |    215 |   end;
 | 
|  |    216 | 
 | 
| 162 |    217 | 
 | 
| 256 |    218 | (* 'sortorder' checks the ordering on sets of classes, i.e. on sorts:
 | 
|  |    219 |    S1 <= S2 , iff for every class C2 in S2 there exists a class C1 in S1
 | 
| 0 |    220 |    with C1 <= C2 (according to an association list 'a')
 | 
|  |    221 | *)
 | 
|  |    222 | 
 | 
| 256 |    223 | fun sortorder a (S1, S2) =
 | 
|  |    224 |   forall  (fn C2 => exists  (fn C1 => leq a (C1, C2))  S1)  S2;
 | 
| 0 |    225 | 
 | 
|  |    226 | 
 | 
|  |    227 | (* 'inj' inserts a new class C into a given class set S (i.e.sort) only if
 | 
|  |    228 |   there exists no class in S which is <= C;
 | 
|  |    229 |   the resulting set is minimal if S was minimal
 | 
|  |    230 | *)
 | 
|  |    231 | 
 | 
| 256 |    232 | fun inj a (C, S) =
 | 
| 0 |    233 |   let fun inj1 [] = [C]
 | 
| 256 |    234 |         | inj1 (D::T) = if leq a (D, C) then D::T
 | 
|  |    235 |                         else if leq a (C, D) then inj1 T
 | 
| 0 |    236 |                              else D::(inj1 T)
 | 
|  |    237 |   in inj1 S end;
 | 
|  |    238 | 
 | 
|  |    239 | 
 | 
|  |    240 | (* 'union_sort' forms the minimal union set of two sorts S1 and S2
 | 
|  |    241 |    under the assumption that S2 is minimal *)
 | 
| 256 |    242 | (* FIXME rename to inter_sort (?) *)
 | 
| 0 |    243 | 
 | 
|  |    244 | fun union_sort a = foldr (inj a);
 | 
|  |    245 | 
 | 
|  |    246 | 
 | 
|  |    247 | (* 'elementwise_union' forms elementwise the minimal union set of two
 | 
|  |    248 |    sort lists under the assumption that the two lists have the same length
 | 
| 256 |    249 | *)
 | 
| 0 |    250 | 
 | 
| 256 |    251 | fun elementwise_union a (Ss1, Ss2) = map (union_sort a) (Ss1~~Ss2);
 | 
|  |    252 | 
 | 
| 0 |    253 | 
 | 
|  |    254 | (* 'lew' checks for two sort lists the ordering for all corresponding list
 | 
|  |    255 |    elements (i.e. sorts) *)
 | 
|  |    256 | 
 | 
| 256 |    257 | fun lew a (w1, w2) = forall (sortorder a)  (w1~~w2);
 | 
|  |    258 | 
 | 
| 0 |    259 | 
 | 
| 256 |    260 | (* 'is_min' checks if a class C is minimal in a given sort S under the
 | 
|  |    261 |    assumption that S contains C *)
 | 
| 0 |    262 | 
 | 
| 256 |    263 | fun is_min a S C = not (exists (fn (D) => less a (D, C)) S);
 | 
| 0 |    264 | 
 | 
|  |    265 | 
 | 
|  |    266 | (* 'min_sort' reduces a sort to its minimal classes *)
 | 
|  |    267 | 
 | 
|  |    268 | fun min_sort a S = distinct(filter (is_min a S) S);
 | 
|  |    269 | 
 | 
|  |    270 | 
 | 
|  |    271 | (* 'min_domain' minimizes the domain sorts of type declarationsl;
 | 
| 256 |    272 |    the function will be applied on the type declarations in extensions *)
 | 
| 0 |    273 | 
 | 
|  |    274 | fun min_domain subclass =
 | 
| 256 |    275 |   let fun one_min (f, (doms, ran)) = (f, (map (min_sort subclass) doms, ran))
 | 
| 0 |    276 |   in map one_min end;
 | 
|  |    277 | 
 | 
|  |    278 | 
 | 
|  |    279 | (* 'min_filter' filters a list 'ars' consisting of arities (domain * class)
 | 
| 256 |    280 |    and gives back a list of those range classes whose domains meet the
 | 
| 0 |    281 |    predicate 'pred' *)
 | 
| 256 |    282 | 
 | 
| 0 |    283 | fun min_filter a pred ars =
 | 
| 256 |    284 |   let fun filt ([], l) = l
 | 
|  |    285 |         | filt ((c, x)::xs, l) = if pred(x) then filt (xs, inj a (c, l))
 | 
|  |    286 |                                else filt (xs, l)
 | 
|  |    287 |   in filt (ars, []) end;
 | 
| 0 |    288 | 
 | 
|  |    289 | 
 | 
|  |    290 | (* 'cod_above' filters all arities whose domains are elementwise >= than
 | 
| 256 |    291 |    a given domain 'w' and gives back a list of the corresponding range
 | 
| 0 |    292 |    classes *)
 | 
|  |    293 | 
 | 
| 256 |    294 | fun cod_above (a, w, ars) = min_filter a (fn w' => lew a (w, w')) ars;
 | 
|  |    295 | 
 | 
|  |    296 | 
 | 
| 0 |    297 | 
 | 
| 200 |    298 | (*Instantiation of type variables in types*)
 | 
|  |    299 | (*Pre: instantiations obey restrictions! *)
 | 
|  |    300 | fun inst_typ tye =
 | 
| 256 |    301 |   let fun inst(Type(a, Ts)) = Type(a, map inst Ts)
 | 
| 200 |    302 |         | inst(T as TFree _) = T
 | 
| 256 |    303 |         | inst(T as TVar(v, _)) =
 | 
|  |    304 |             (case assoc(tye, v) of Some U => inst U | None => T)
 | 
| 200 |    305 |   in inst end;
 | 
| 0 |    306 | 
 | 
|  |    307 | (* 'least_sort' returns for a given type its maximum sort:
 | 
|  |    308 |    - type variables, free types: the sort brought with
 | 
|  |    309 |    - type constructors: recursive determination of the maximum sort of the
 | 
| 256 |    310 |                     arguments if the type is declared in 'coreg' of the
 | 
|  |    311 |                     given type signature  *)
 | 
| 0 |    312 | 
 | 
| 256 |    313 | fun least_sort (tsig as TySg{subclass, coreg, ...}) =
 | 
|  |    314 |   let fun ls(T as Type(a, Ts)) =
 | 
|  |    315 |                  (case assoc (coreg, a) of
 | 
|  |    316 |                           Some(ars) => cod_above(subclass, map ls Ts, ars)
 | 
|  |    317 |                         | None => raise TYPE(undcl_type a, [T], []))
 | 
|  |    318 |         | ls(TFree(a, S)) = S
 | 
|  |    319 |         | ls(TVar(a, S)) = S
 | 
| 0 |    320 |   in ls end;
 | 
|  |    321 | 
 | 
|  |    322 | 
 | 
| 256 |    323 | fun check_has_sort(tsig as TySg{subclass, coreg, ...}, T, S) =
 | 
|  |    324 |   if sortorder subclass ((least_sort tsig T), S) then ()
 | 
|  |    325 |   else raise TYPE("Type not of sort " ^ (str_of_sort S), [T], [])
 | 
| 0 |    326 | 
 | 
|  |    327 | 
 | 
|  |    328 | (*Instantiation of type variables in types *)
 | 
| 256 |    329 | fun inst_typ_tvars(tsig, tye) =
 | 
|  |    330 |   let fun inst(Type(a, Ts)) = Type(a, map inst Ts)
 | 
|  |    331 |         | inst(T as TFree _) = T
 | 
|  |    332 |         | inst(T as TVar(v, S)) = (case assoc(tye, v) of
 | 
|  |    333 |                 None => T | Some(U) => (check_has_sort(tsig, U, S); U))
 | 
| 0 |    334 |   in inst end;
 | 
|  |    335 | 
 | 
|  |    336 | (*Instantiation of type variables in terms *)
 | 
| 256 |    337 | fun inst_term_tvars(tsig, tye) = map_term_types (inst_typ_tvars(tsig, tye));
 | 
| 200 |    338 | 
 | 
|  |    339 | 
 | 
|  |    340 | (* expand_typ *)
 | 
|  |    341 | 
 | 
| 256 |    342 | fun expand_typ (TySg {abbrs, ...}) ty =
 | 
|  |    343 |   let
 | 
| 621 |    344 |     val idx = maxidx_of_typ ty + 1;
 | 
|  |    345 | 
 | 
|  |    346 |     fun expand (Type (a, Ts)) =
 | 
| 256 |    347 |           (case assoc (abbrs, a) of
 | 
| 621 |    348 |             Some (vs, U) =>
 | 
|  |    349 |               expand (inst_typ (map (rpair idx) vs ~~ Ts) (incr_tvar idx U))
 | 
|  |    350 |           | None => Type (a, map expand Ts))
 | 
|  |    351 |       | expand T = T
 | 
| 256 |    352 |   in
 | 
| 621 |    353 |     expand ty
 | 
| 256 |    354 |   end;
 | 
|  |    355 | 
 | 
|  |    356 | val norm_typ = expand_typ;
 | 
|  |    357 | 
 | 
|  |    358 | 
 | 
|  |    359 | 
 | 
|  |    360 | (** type matching **)
 | 
| 200 |    361 | 
 | 
| 0 |    362 | exception TYPE_MATCH;
 | 
|  |    363 | 
 | 
| 256 |    364 | (*typ_match (s, (U, T)) = s' <==> s'(U) = T and s' is an extension of s*)
 | 
|  |    365 | fun typ_match tsig =
 | 
|  |    366 |   let
 | 
|  |    367 |     fun match (subs, (TVar (v, S), T)) =
 | 
|  |    368 |           (case assoc (subs, v) of
 | 
|  |    369 |             None => ((v, (check_has_sort (tsig, T, S); T)) :: subs
 | 
|  |    370 |               handle TYPE _ => raise TYPE_MATCH)
 | 
| 422 |    371 |           | Some U => if U = T then subs else raise TYPE_MATCH)
 | 
| 256 |    372 |       | match (subs, (Type (a, Ts), Type (b, Us))) =
 | 
|  |    373 |           if a <> b then raise TYPE_MATCH
 | 
|  |    374 |           else foldl match (subs, Ts ~~ Us)
 | 
| 422 |    375 |       | match (subs, (TFree x, TFree y)) =
 | 
| 256 |    376 |           if x = y then subs else raise TYPE_MATCH
 | 
|  |    377 |       | match _ = raise TYPE_MATCH;
 | 
|  |    378 |   in match end;
 | 
| 0 |    379 | 
 | 
|  |    380 | 
 | 
| 256 |    381 | fun typ_instance (tsig, T, U) =
 | 
|  |    382 |   (typ_match tsig ([], (U, T)); true) handle TYPE_MATCH => false;
 | 
|  |    383 | 
 | 
|  |    384 | 
 | 
|  |    385 | 
 | 
|  |    386 | (** build type signatures **)
 | 
|  |    387 | 
 | 
| 416 |    388 | fun make_tsig (classes, subclass, default, args, abbrs, coreg) =
 | 
|  |    389 |   TySg {classes = classes, subclass = subclass, default = default,
 | 
|  |    390 |     args = args, abbrs = abbrs, coreg = coreg};
 | 
|  |    391 | 
 | 
|  |    392 | val tsig0 = make_tsig ([], [], [], [], [], []);
 | 
| 256 |    393 | 
 | 
| 0 |    394 | 
 | 
| 401 |    395 | (* sorts *)
 | 
|  |    396 | 
 | 
| 416 |    397 | fun subsort (TySg {subclass, ...}) (S1, S2) =
 | 
|  |    398 |   sortorder subclass (S1, S2);
 | 
|  |    399 | 
 | 
| 401 |    400 | fun norm_sort (TySg {subclass, ...}) S =
 | 
|  |    401 |   sort_strings (min_sort subclass S);
 | 
|  |    402 | 
 | 
| 416 |    403 | fun rem_sorts (Type (a, tys)) = Type (a, map rem_sorts tys)
 | 
|  |    404 |   | rem_sorts (TFree (x, _)) = TFree (x, [])
 | 
|  |    405 |   | rem_sorts (TVar (xi, _)) = TVar (xi, []);
 | 
| 401 |    406 | 
 | 
|  |    407 | 
 | 
| 416 |    408 | (* typ_errors *)
 | 
| 256 |    409 | 
 | 
| 416 |    410 | (*check validity of (not necessarily normal) type; accumulate error messages*)
 | 
| 256 |    411 | 
 | 
| 416 |    412 | fun typ_errors tsig (typ, errors) =
 | 
| 256 |    413 |   let
 | 
| 416 |    414 |     val TySg {classes, args, abbrs, ...} = tsig;
 | 
|  |    415 | 
 | 
|  |    416 |     fun class_err (errs, c) =
 | 
|  |    417 |       if c mem classes then errs
 | 
|  |    418 |       else undcl_class c ins errs;
 | 
| 256 |    419 | 
 | 
|  |    420 |     val sort_err = foldl class_err;
 | 
| 0 |    421 | 
 | 
| 256 |    422 |     fun typ_errs (Type (c, Us), errs) =
 | 
|  |    423 |           let
 | 
|  |    424 |             val errs' = foldr typ_errs (Us, errs);
 | 
|  |    425 |             fun nargs n =
 | 
|  |    426 |               if n = length Us then errs'
 | 
| 416 |    427 |               else ("Wrong number of arguments: " ^ quote c) ins errs';
 | 
| 256 |    428 |           in
 | 
|  |    429 |             (case assoc (args, c) of
 | 
|  |    430 |               Some n => nargs n
 | 
|  |    431 |             | None =>
 | 
|  |    432 |                 (case assoc (abbrs, c) of
 | 
|  |    433 |                   Some (vs, _) => nargs (length vs)
 | 
| 416 |    434 |                 | None => undcl_type c ins errs))
 | 
| 256 |    435 |           end
 | 
|  |    436 |     | typ_errs (TFree (_, S), errs) = sort_err (errs, S)
 | 
| 416 |    437 |     | typ_errs (TVar ((x, i), S), errs) =
 | 
|  |    438 |         if i < 0 then
 | 
|  |    439 |           ("Negative index for TVar " ^ quote x) ins sort_err (errs, S)
 | 
|  |    440 |         else sort_err (errs, S);
 | 
| 256 |    441 |   in
 | 
| 416 |    442 |     typ_errs (typ, errors)
 | 
| 256 |    443 |   end;
 | 
|  |    444 | 
 | 
|  |    445 | 
 | 
|  |    446 | (* cert_typ *)
 | 
|  |    447 | 
 | 
|  |    448 | (*check and normalize typ wrt. tsig; errors are indicated by exception TYPE*)
 | 
|  |    449 | 
 | 
|  |    450 | fun cert_typ tsig ty =
 | 
|  |    451 |   (case typ_errors tsig (ty, []) of
 | 
|  |    452 |     [] => norm_typ tsig ty
 | 
|  |    453 |   | errs => raise_type (cat_lines errs) [ty] []);
 | 
|  |    454 | 
 | 
|  |    455 | 
 | 
|  |    456 | 
 | 
| 422 |    457 | (** merge type signatures **)
 | 
| 256 |    458 | 
 | 
| 422 |    459 | (*'assoc_union' merges two association lists if the contents associated
 | 
|  |    460 |   the keys are lists*)
 | 
| 0 |    461 | 
 | 
| 422 |    462 | fun assoc_union (as1, []) = as1
 | 
|  |    463 |   | assoc_union (as1, (key, l2) :: as2) =
 | 
|  |    464 |       (case assoc (as1, key) of
 | 
|  |    465 |         Some l1 => assoc_union (overwrite (as1, (key, l1 union l2)), as2)
 | 
|  |    466 |       | None => assoc_union ((key, l2) :: as1, as2));
 | 
| 0 |    467 | 
 | 
|  |    468 | 
 | 
| 422 |    469 | (* merge subclass *)
 | 
| 0 |    470 | 
 | 
| 422 |    471 | fun merge_subclass (subclass1, subclass2) =
 | 
|  |    472 |   let val subclass = transitive_closure (assoc_union (subclass1, subclass2)) in
 | 
|  |    473 |     if exists (op mem) subclass then
 | 
|  |    474 |       error ("Cyclic class structure!")   (* FIXME improve msg, raise TERM *)
 | 
|  |    475 |     else subclass
 | 
| 416 |    476 |   end;
 | 
|  |    477 | 
 | 
|  |    478 | 
 | 
| 422 |    479 | (* coregularity *)
 | 
| 0 |    480 | 
 | 
|  |    481 | (* 'is_unique_decl' checks if there exists just one declaration t:(Ss)C *)
 | 
|  |    482 | 
 | 
| 256 |    483 | fun is_unique_decl coreg (t, (s, w)) = case assoc2 (coreg, (t, s)) of
 | 
| 0 |    484 |       Some(w1) => if w = w1 then () else
 | 
| 256 |    485 |         error("There are two declarations\n" ^
 | 
| 416 |    486 |               str_of_arity(t, w, [s]) ^ " and\n" ^
 | 
|  |    487 |               str_of_arity(t, w1, [s]) ^ "\n" ^
 | 
| 0 |    488 |               "with the same result class.")
 | 
|  |    489 |     | None => ();
 | 
|  |    490 | 
 | 
|  |    491 | (* 'restr2' checks if there are two declarations t:(Ss1)C1 and t:(Ss2)C2
 | 
|  |    492 |    such that C1 >= C2 then Ss1 >= Ss2 (elementwise) *)
 | 
|  |    493 | 
 | 
| 256 |    494 | fun subs (classes, subclass) C =
 | 
|  |    495 |   let fun sub (rl, l) = if leq subclass (l, C) then l::rl else rl
 | 
|  |    496 |   in foldl sub ([], classes) end;
 | 
| 0 |    497 | 
 | 
| 256 |    498 | fun coreg_err(t, (w1, C), (w2, D)) =
 | 
| 416 |    499 |     error("Declarations " ^ str_of_arity(t, w1, [C]) ^ " and "
 | 
|  |    500 |                           ^ str_of_arity(t, w2, [D]) ^ " are in conflict");
 | 
| 0 |    501 | 
 | 
| 256 |    502 | fun restr2 classes (subclass, coreg) (t, (s, w)) =
 | 
|  |    503 |   let fun restr ([], test) = ()
 | 
| 416 |    504 |         | restr (s1::Ss, test) =
 | 
|  |    505 |             (case assoc2 (coreg, (t, s1)) of
 | 
|  |    506 |               Some dom =>
 | 
|  |    507 |                 if lew subclass (test (w, dom))
 | 
|  |    508 |                 then restr (Ss, test)
 | 
|  |    509 |                 else coreg_err (t, (w, s), (dom, s1))
 | 
| 256 |    510 |             | None => restr (Ss, test))
 | 
|  |    511 |       fun forward (t, (s, w)) =
 | 
|  |    512 |         let val s_sups = case assoc (subclass, s) of
 | 
|  |    513 |                    Some(s_sups) => s_sups | None => err_undcl_class(s);
 | 
|  |    514 |         in restr (s_sups, I) end
 | 
|  |    515 |       fun backward (t, (s, w)) =
 | 
|  |    516 |         let val s_subs = subs (classes, subclass) s
 | 
|  |    517 |         in restr (s_subs, fn (x, y) => (y, x)) end
 | 
|  |    518 |   in (backward (t, (s, w)); forward (t, (s, w))) end;
 | 
| 0 |    519 | 
 | 
|  |    520 | 
 | 
| 256 |    521 | fun varying_decls t =
 | 
|  |    522 |   error ("Type constructor " ^ quote t ^ " has varying number of arguments");
 | 
| 0 |    523 | 
 | 
|  |    524 | 
 | 
| 422 |    525 | (* 'merge_coreg' builds the union of two 'coreg' lists;
 | 
|  |    526 |    it only checks the two restriction conditions and inserts afterwards
 | 
|  |    527 |    all elements of the second list into the first one *)
 | 
|  |    528 | 
 | 
|  |    529 | fun merge_coreg classes subclass1 =
 | 
|  |    530 |   let fun test_ar classes (t, ars1) (coreg1, (s, w)) =
 | 
|  |    531 |         (is_unique_decl coreg1 (t, (s, w));
 | 
|  |    532 |          restr2 classes (subclass1, coreg1) (t, (s, w));
 | 
|  |    533 |          overwrite (coreg1, (t, (s, w) ins ars1)));
 | 
|  |    534 | 
 | 
|  |    535 |       fun merge_c (coreg1, (c as (t, ars2))) = case assoc (coreg1, t) of
 | 
|  |    536 |           Some(ars1) => foldl (test_ar classes (t, ars1)) (coreg1, ars2)
 | 
|  |    537 |         | None => c::coreg1
 | 
|  |    538 |   in foldl merge_c end;
 | 
|  |    539 | 
 | 
|  |    540 | fun merge_args (args, (t, n)) =
 | 
|  |    541 |   (case assoc (args, t) of
 | 
|  |    542 |     Some m => if m = n then args else varying_decls t
 | 
|  |    543 |   | None => (t, n) :: args);
 | 
|  |    544 | 
 | 
|  |    545 | fun merge_abbrs (abbrs1, abbrs2) =
 | 
| 621 |    546 |   let val abbrs = abbrs1 union abbrs2 in
 | 
|  |    547 |     (case gen_duplicates eq_fst abbrs of
 | 
| 422 |    548 |       [] => abbrs
 | 
| 621 |    549 |     | dups => raise_term (dup_tyabbrs (map fst dups)) [])
 | 
| 422 |    550 |   end;
 | 
|  |    551 | 
 | 
|  |    552 | 
 | 
|  |    553 | (* 'merge_tsigs' takes the above declared functions to merge two type
 | 
|  |    554 |   signatures *)
 | 
|  |    555 | 
 | 
|  |    556 | fun merge_tsigs(TySg{classes=classes1, default=default1, subclass=subclass1, args=args1,
 | 
|  |    557 |            coreg=coreg1, abbrs=abbrs1},
 | 
|  |    558 |           TySg{classes=classes2, default=default2, subclass=subclass2, args=args2,
 | 
|  |    559 |            coreg=coreg2, abbrs=abbrs2}) =
 | 
|  |    560 |   let val classes' = classes1 union classes2;
 | 
|  |    561 |       val subclass' = merge_subclass (subclass1, subclass2);
 | 
|  |    562 |       val args' = foldl merge_args (args1, args2)
 | 
|  |    563 |       val coreg' = merge_coreg classes' subclass' (coreg1, coreg2);
 | 
|  |    564 |       val default' = min_sort subclass' (default1 @ default2);
 | 
|  |    565 |       val abbrs' = merge_abbrs(abbrs1, abbrs2);
 | 
|  |    566 |   in TySg{classes=classes' , default=default', subclass=subclass', args=args',
 | 
|  |    567 |           coreg=coreg' , abbrs = abbrs' }
 | 
|  |    568 |   end;
 | 
|  |    569 | 
 | 
|  |    570 | 
 | 
|  |    571 | 
 | 
|  |    572 | (*** extend type signatures ***)
 | 
|  |    573 | 
 | 
| 621 |    574 | (** add classes and subclass relations**)
 | 
| 422 |    575 | 
 | 
|  |    576 | fun add_classes classes cs =
 | 
|  |    577 |   (case cs inter classes of
 | 
|  |    578 |     [] => cs @ classes
 | 
|  |    579 |   | dups => err_dup_classes cs);
 | 
|  |    580 | 
 | 
|  |    581 | 
 | 
|  |    582 | (*'add_subclass' adds a tuple consisting of a new class (the new class has
 | 
|  |    583 |   already been inserted into the 'classes' list) and its superclasses (they
 | 
|  |    584 |   must be declared in 'classes' too) to the 'subclass' list of the given type
 | 
|  |    585 |   signature; furthermore all inherited superclasses according to the
 | 
|  |    586 |   superclasses brought with are inserted and there is a check that there are
 | 
|  |    587 |   no cycles (i.e. C <= D <= C, with C <> D);*)
 | 
|  |    588 | 
 | 
|  |    589 | fun add_subclass classes (subclass, (s, ges)) =
 | 
| 621 |    590 |   let
 | 
|  |    591 |     fun upd (subclass, s') =
 | 
|  |    592 |       if s' mem classes then
 | 
| 422 |    593 |         let val ges' = the (assoc (subclass, s))
 | 
|  |    594 |         in case assoc (subclass, s') of
 | 
|  |    595 |              Some sups => if s mem sups
 | 
|  |    596 |                            then error(" Cycle :" ^ s^" <= "^ s'^" <= "^ s )
 | 
|  |    597 |                            else overwrite (subclass, (s, sups union ges'))
 | 
|  |    598 |            | None => subclass
 | 
| 621 |    599 |         end
 | 
|  |    600 |       else err_undcl_class s'
 | 
|  |    601 |   in foldl upd (subclass @ [(s, ges)], ges) end;
 | 
| 422 |    602 | 
 | 
|  |    603 | 
 | 
|  |    604 | (* 'extend_classes' inserts all new classes into the corresponding
 | 
|  |    605 |    lists ('classes', 'subclass') if possible *)
 | 
|  |    606 | 
 | 
| 621 |    607 | fun extend_classes (classes, subclass, new_classes) =
 | 
|  |    608 |   let
 | 
|  |    609 |     val classes' = add_classes classes (map fst new_classes);
 | 
|  |    610 |     val subclass' = foldl (add_subclass classes') (subclass, new_classes);
 | 
| 422 |    611 |   in (classes', subclass') end;
 | 
|  |    612 | 
 | 
|  |    613 | 
 | 
| 621 |    614 | (* ext_tsig_classes *)
 | 
|  |    615 | 
 | 
|  |    616 | fun ext_tsig_classes tsig new_classes =
 | 
|  |    617 |   let
 | 
|  |    618 |     val TySg {classes, subclass, default, args, abbrs, coreg} = tsig;
 | 
|  |    619 |     val (classes', subclass') = extend_classes (classes, subclass, new_classes);
 | 
|  |    620 |   in
 | 
|  |    621 |     make_tsig (classes', subclass', default, args, abbrs, coreg)
 | 
|  |    622 |   end;
 | 
|  |    623 | 
 | 
|  |    624 | 
 | 
| 422 |    625 | (* ext_tsig_subclass *)
 | 
|  |    626 | 
 | 
|  |    627 | fun ext_tsig_subclass tsig pairs =
 | 
|  |    628 |   let
 | 
|  |    629 |     val TySg {classes, subclass, default, args, abbrs, coreg} = tsig;
 | 
|  |    630 | 
 | 
|  |    631 |     (* FIXME clean! *)
 | 
|  |    632 |     val subclass' =
 | 
|  |    633 |       merge_subclass (subclass, map (fn (c1, c2) => (c1, [c2])) pairs);
 | 
|  |    634 |   in
 | 
|  |    635 |     make_tsig (classes, subclass', default, args, abbrs, coreg)
 | 
|  |    636 |   end;
 | 
|  |    637 | 
 | 
|  |    638 | 
 | 
|  |    639 | (* ext_tsig_defsort *)
 | 
|  |    640 | 
 | 
|  |    641 | fun ext_tsig_defsort (TySg {classes, subclass, args, abbrs, coreg, ...}) default =
 | 
|  |    642 |   make_tsig (classes, subclass, default, args, abbrs, coreg);
 | 
|  |    643 | 
 | 
|  |    644 | 
 | 
|  |    645 | 
 | 
| 621 |    646 | (** add types **)
 | 
| 582 |    647 | 
 | 
|  |    648 | fun ext_tsig_types (TySg {classes, subclass, default, args, abbrs, coreg}) ts =
 | 
|  |    649 |   let
 | 
|  |    650 |     fun check_type (c, n) =
 | 
|  |    651 |       if n < 0 then err_neg_args c
 | 
|  |    652 |       else if is_some (assoc (args, c)) then err_dup_tycon c
 | 
|  |    653 |       else if is_some (assoc (abbrs, c)) then err_ty_confl c
 | 
|  |    654 |       else ();
 | 
|  |    655 |   in
 | 
|  |    656 |     seq check_type ts;
 | 
|  |    657 |     make_tsig (classes, subclass, default, ts @ args, abbrs,
 | 
| 621 |    658 |       map (rpair [] o #1) ts @ coreg)
 | 
| 582 |    659 |   end;
 | 
|  |    660 | 
 | 
|  |    661 | 
 | 
|  |    662 | 
 | 
|  |    663 | (** add type abbreviations **)
 | 
|  |    664 | 
 | 
|  |    665 | fun abbr_errors tsig (a, (lhs_vs, rhs)) =
 | 
|  |    666 |   let
 | 
|  |    667 |     val TySg {args, abbrs, ...} = tsig;
 | 
| 621 |    668 |     val rhs_vs = map (#1 o #1) (typ_tvars rhs);
 | 
| 582 |    669 | 
 | 
|  |    670 |     val dup_lhs_vars =
 | 
|  |    671 |       (case duplicates lhs_vs of
 | 
|  |    672 |         [] => []
 | 
| 621 |    673 |       | vs => ["Duplicate variables on lhs: " ^ commas_quote vs]);
 | 
| 582 |    674 | 
 | 
|  |    675 |     val extra_rhs_vars =
 | 
|  |    676 |       (case gen_rems (op =) (rhs_vs, lhs_vs) of
 | 
|  |    677 |         [] => []
 | 
| 621 |    678 |       | vs => ["Extra variables on rhs: " ^ commas_quote vs]);
 | 
| 582 |    679 | 
 | 
|  |    680 |     val tycon_confl =
 | 
|  |    681 |       if is_none (assoc (args, a)) then []
 | 
|  |    682 |       else [ty_confl a];
 | 
|  |    683 | 
 | 
|  |    684 |     val dup_abbr =
 | 
|  |    685 |       if is_none (assoc (abbrs, a)) then []
 | 
|  |    686 |       else ["Duplicate declaration of abbreviation"];
 | 
|  |    687 |   in
 | 
|  |    688 |     dup_lhs_vars @ extra_rhs_vars @ tycon_confl @ dup_abbr @
 | 
|  |    689 |       typ_errors tsig (rhs, [])
 | 
|  |    690 |   end;
 | 
|  |    691 | 
 | 
| 621 |    692 | fun prep_abbr tsig (a, vs, raw_rhs) =
 | 
|  |    693 |   let
 | 
|  |    694 |     fun err msgs = (seq writeln msgs;
 | 
|  |    695 |       error ("The error(s) above occurred in type abbreviation " ^ quote a));
 | 
|  |    696 | 
 | 
|  |    697 |     val rhs = rem_sorts (varifyT (no_tvars raw_rhs))
 | 
|  |    698 |       handle TYPE (msg, _, _) => err [msg];
 | 
|  |    699 |     val abbr = (a, (vs, rhs));
 | 
|  |    700 |   in
 | 
| 582 |    701 |     (case abbr_errors tsig abbr of
 | 
| 621 |    702 |       [] => abbr
 | 
|  |    703 |     | msgs => err msgs)
 | 
| 582 |    704 |   end;
 | 
|  |    705 | 
 | 
| 621 |    706 | fun add_abbr (tsig as TySg {classes, subclass, default, args, coreg, abbrs}, abbr) =
 | 
|  |    707 |   make_tsig
 | 
|  |    708 |     (classes, subclass, default, args, prep_abbr tsig abbr :: abbrs, coreg);
 | 
|  |    709 | 
 | 
|  |    710 | fun ext_tsig_abbrs tsig raw_abbrs = foldl add_abbr (tsig, raw_abbrs);
 | 
| 582 |    711 | 
 | 
|  |    712 | 
 | 
|  |    713 | 
 | 
| 422 |    714 | (** add arities **)
 | 
|  |    715 | 
 | 
| 0 |    716 | (* 'coregular' checks
 | 
|  |    717 |    - the two restriction conditions 'is_unique_decl' and 'restr2'
 | 
| 256 |    718 |    - if the classes in the new type declarations are known in the
 | 
| 0 |    719 |      given type signature
 | 
|  |    720 |    - if one type constructor has always the same number of arguments;
 | 
| 256 |    721 |    if one type declaration has passed all checks it is inserted into
 | 
| 0 |    722 |    the 'coreg' association list of the given type signatrure  *)
 | 
|  |    723 | 
 | 
| 256 |    724 | fun coregular (classes, subclass, args) =
 | 
|  |    725 |   let fun ex C = if C mem classes then () else err_undcl_class(C);
 | 
| 0 |    726 | 
 | 
| 256 |    727 |       fun addar(w, C) (coreg, t) = case assoc(args, t) of
 | 
| 0 |    728 |             Some(n) => if n <> length w then varying_decls(t) else
 | 
| 256 |    729 |                      (is_unique_decl coreg (t, (C, w));
 | 
|  |    730 |                       (seq o seq) ex w;
 | 
|  |    731 |                       restr2 classes (subclass, coreg) (t, (C, w));
 | 
| 416 |    732 |                       let val ars = the (assoc(coreg, t))
 | 
| 256 |    733 |                       in overwrite(coreg, (t, (C, w) ins ars)) end)
 | 
|  |    734 |           | None => err_undcl_type(t);
 | 
| 0 |    735 | 
 | 
| 256 |    736 |       fun addts(coreg, (ts, ar)) = foldl (addar ar) (coreg, ts)
 | 
| 0 |    737 | 
 | 
|  |    738 |   in addts end;
 | 
|  |    739 | 
 | 
|  |    740 | 
 | 
|  |    741 | (* 'close' extends the 'coreg' association list after all new type
 | 
|  |    742 |    declarations have been inserted successfully:
 | 
|  |    743 |    for every declaration t:(Ss)C , for all classses D with C <= D:
 | 
|  |    744 |       if there is no declaration t:(Ss')C' with C < C' and C' <= D
 | 
|  |    745 |       then insert the declaration t:(Ss)D into 'coreg'
 | 
|  |    746 |    this means, if there exists a declaration t:(Ss)C and there is
 | 
|  |    747 |    no declaration t:(Ss')D with C <=D then the declaration holds
 | 
| 256 |    748 |    for all range classes more general than C *)
 | 
|  |    749 | 
 | 
| 621 |    750 | fun close subclass coreg =
 | 
| 256 |    751 |   let fun check sl (l, (s, dom)) = case assoc (subclass, s) of
 | 
| 621 |    752 |           Some sups =>
 | 
| 256 |    753 |             let fun close_sup (l, sup) =
 | 
|  |    754 |                   if exists (fn s'' => less subclass (s, s'') andalso
 | 
|  |    755 |                                        leq subclass (s'', sup)) sl
 | 
| 0 |    756 |                   then l
 | 
| 256 |    757 |                   else (sup, dom)::l
 | 
|  |    758 |             in foldl close_sup (l, sups) end
 | 
| 0 |    759 |         | None => l;
 | 
| 256 |    760 |       fun ext (s, l) = (s, foldl (check (map #1 l)) (l, l));
 | 
| 0 |    761 |   in map ext coreg end;
 | 
|  |    762 | 
 | 
| 422 |    763 | 
 | 
| 621 |    764 | (* ext_tsig_arities *)
 | 
| 256 |    765 | 
 | 
| 621 |    766 | fun ext_tsig_arities tsig sarities =
 | 
| 416 |    767 |   let
 | 
| 621 |    768 |     val TySg {classes, subclass, default, args, coreg, abbrs} = tsig;
 | 
|  |    769 |     val arities =
 | 
|  |    770 |       flat (map (fn (t, ss, cs) => map (fn c => ([t], (ss, c))) cs) sarities);
 | 
|  |    771 |     val coreg' =
 | 
|  |    772 |       foldl (coregular (classes, subclass, args))
 | 
|  |    773 |         (coreg, min_domain subclass arities)
 | 
|  |    774 |       |> close subclass;
 | 
| 416 |    775 |   in
 | 
| 621 |    776 |     make_tsig (classes, subclass, default, args, abbrs, coreg')
 | 
| 416 |    777 |   end;
 | 
| 0 |    778 | 
 | 
|  |    779 | 
 | 
| 416 |    780 | 
 | 
|  |    781 | (*** type unification and inference ***)
 | 
| 0 |    782 | 
 | 
|  |    783 | (*
 | 
| 621 |    784 |   Input:
 | 
|  |    785 |     - a 'raw' term which contains only dummy types and some explicit type
 | 
|  |    786 |       constraints encoded as terms.
 | 
|  |    787 |     - the expected type of the term.
 | 
| 0 |    788 | 
 | 
| 621 |    789 |   Output:
 | 
|  |    790 |     - the correctly typed term
 | 
|  |    791 |     - the substitution needed to unify the actual type of the term with its
 | 
|  |    792 |       expected type; only the TVars in the expected type are included.
 | 
| 0 |    793 | 
 | 
| 621 |    794 |   During type inference all TVars in the term have negative index. This keeps
 | 
|  |    795 |   them apart from normal TVars, which is essential, because at the end the
 | 
|  |    796 |   type of the term is unified with the expected type, which contains normal
 | 
|  |    797 |   TVars.
 | 
| 0 |    798 | 
 | 
| 621 |    799 |   1. Add initial type information to the term (attach_types).
 | 
|  |    800 |      This freezes (freeze_vars) TVars in explicitly provided types (eg
 | 
|  |    801 |      constraints or defaults) by turning them into TFrees.
 | 
|  |    802 |   2. Carry out type inference, possibly introducing new negative TVars.
 | 
|  |    803 |   3. Unify actual and expected type.
 | 
|  |    804 |   4. Turn all (negative) TVars into unique new TFrees (freeze).
 | 
|  |    805 |   5. Thaw all TVars frozen in step 1 (thaw_vars).
 | 
| 0 |    806 | *)
 | 
|  |    807 | 
 | 
|  |    808 | (*Raised if types are not unifiable*)
 | 
|  |    809 | exception TUNIFY;
 | 
|  |    810 | 
 | 
| 621 |    811 | val tyvar_count = ref ~1;
 | 
| 0 |    812 | 
 | 
|  |    813 | fun tyinit() = (tyvar_count := ~1);
 | 
|  |    814 | 
 | 
| 621 |    815 | fun new_tvar_inx () = (tyvar_count := ! tyvar_count - 1; ! tyvar_count)
 | 
| 0 |    816 | 
 | 
|  |    817 | (*
 | 
|  |    818 | Generate new TVar.  Index is < ~1 to distinguish it from TVars generated from
 | 
|  |    819 | variable names (see id_type).  Name is arbitrary because index is new.
 | 
|  |    820 | *)
 | 
|  |    821 | 
 | 
| 256 |    822 | fun gen_tyvar(S) = TVar(("'a", new_tvar_inx()), S);
 | 
| 0 |    823 | 
 | 
|  |    824 | (*Occurs check: type variable occurs in type?*)
 | 
|  |    825 | fun occ v tye =
 | 
| 256 |    826 |   let fun occ(Type(_, Ts)) = exists occ Ts
 | 
| 0 |    827 |         | occ(TFree _) = false
 | 
| 256 |    828 |         | occ(TVar(w, _)) = v=w orelse
 | 
|  |    829 |                            (case assoc(tye, w) of
 | 
| 0 |    830 |                               None   => false
 | 
|  |    831 |                             | Some U => occ U);
 | 
|  |    832 |   in occ end;
 | 
|  |    833 | 
 | 
| 256 |    834 | (*Chase variable assignments in tye.
 | 
|  |    835 |   If devar (T, tye) returns a type var then it must be unassigned.*)
 | 
|  |    836 | fun devar (T as TVar(v, _), tye) = (case  assoc(tye, v)  of
 | 
|  |    837 |           Some U =>  devar (U, tye)
 | 
| 0 |    838 |         | None   =>  T)
 | 
| 256 |    839 |   | devar (T, tye) = T;
 | 
| 0 |    840 | 
 | 
|  |    841 | 
 | 
|  |    842 | (* 'dom' returns for a type constructor t the list of those domains
 | 
|  |    843 |    which deliver a given range class C *)
 | 
|  |    844 | 
 | 
| 256 |    845 | fun dom coreg t C = case assoc2 (coreg, (t, C)) of
 | 
| 0 |    846 |     Some(Ss) => Ss
 | 
|  |    847 |   | None => raise TUNIFY;
 | 
|  |    848 | 
 | 
|  |    849 | 
 | 
|  |    850 | (* 'Dom' returns the union of all domain lists of 'dom' for a given sort S
 | 
|  |    851 |    (i.e. a set of range classes ); the union is carried out elementwise
 | 
|  |    852 |    for the seperate sorts in the domains *)
 | 
|  |    853 | 
 | 
| 256 |    854 | fun Dom (subclass, coreg) (t, S) =
 | 
| 0 |    855 |   let val domlist = map (dom coreg t) S;
 | 
|  |    856 |   in if null domlist then []
 | 
| 256 |    857 |      else foldl (elementwise_union subclass) (hd domlist, tl domlist)
 | 
| 0 |    858 |   end;
 | 
|  |    859 | 
 | 
|  |    860 | 
 | 
| 256 |    861 | fun W ((T, S), tsig as TySg{subclass, coreg, ...}, tye) =
 | 
|  |    862 |   let fun Wd ((T, S), tye) = W ((devar (T, tye), S), tsig, tye)
 | 
|  |    863 |       fun Wk(T as TVar(v, S')) =
 | 
|  |    864 |               if sortorder subclass (S', S) then tye
 | 
|  |    865 |               else (v, gen_tyvar(union_sort subclass (S', S)))::tye
 | 
|  |    866 |         | Wk(T as TFree(v, S')) = if sortorder subclass (S', S) then tye
 | 
|  |    867 |                                  else raise TUNIFY
 | 
|  |    868 |         | Wk(T as Type(f, Ts)) =
 | 
|  |    869 |            if null S then tye
 | 
|  |    870 |            else foldr Wd (Ts~~(Dom (subclass, coreg) (f, S)) , tye)
 | 
| 0 |    871 |   in Wk(T) end;
 | 
|  |    872 | 
 | 
|  |    873 | 
 | 
|  |    874 | (* Order-sorted Unification of Types (U)  *)
 | 
|  |    875 | 
 | 
|  |    876 | (* Precondition: both types are well-formed w.r.t. type constructor arities *)
 | 
| 256 |    877 | fun unify (tsig as TySg{subclass, coreg, ...}) =
 | 
|  |    878 |   let fun unif ((T, U), tye) =
 | 
|  |    879 |         case (devar(T, tye), devar(U, tye)) of
 | 
|  |    880 |           (T as TVar(v, S1), U as TVar(w, S2)) =>
 | 
| 0 |    881 |              if v=w then tye else
 | 
| 256 |    882 |              if sortorder subclass (S1, S2) then (w, T)::tye else
 | 
|  |    883 |              if sortorder subclass (S2, S1) then (v, U)::tye
 | 
|  |    884 |              else let val nu = gen_tyvar (union_sort subclass (S1, S2))
 | 
|  |    885 |                   in (v, nu)::(w, nu)::tye end
 | 
|  |    886 |         | (T as TVar(v, S), U) =>
 | 
|  |    887 |              if occ v tye U then raise TUNIFY else W ((U, S), tsig, (v, U)::tye)
 | 
|  |    888 |         | (U, T as TVar (v, S)) =>
 | 
|  |    889 |              if occ v tye U then raise TUNIFY else W ((U, S), tsig, (v, U)::tye)
 | 
|  |    890 |         | (Type(a, Ts), Type(b, Us)) =>
 | 
|  |    891 |              if a<>b then raise TUNIFY else foldr unif (Ts~~Us, tye)
 | 
|  |    892 |         | (T, U) => if T=U then tye else raise TUNIFY
 | 
| 0 |    893 |   in unif end;
 | 
|  |    894 | 
 | 
|  |    895 | 
 | 
| 450 |    896 | (* raw_unify (ignores sorts) *)
 | 
|  |    897 | 
 | 
|  |    898 | fun raw_unify (ty1, ty2) =
 | 
|  |    899 |   (unify tsig0 ((rem_sorts ty1, rem_sorts ty2), []); true)
 | 
|  |    900 |     handle TUNIFY => false;
 | 
|  |    901 | 
 | 
|  |    902 | 
 | 
| 0 |    903 | (*Type inference for polymorphic term*)
 | 
|  |    904 | fun infer tsig =
 | 
| 256 |    905 |   let fun inf(Ts, Const (_, T), tye) = (T, tye)
 | 
|  |    906 |         | inf(Ts, Free  (_, T), tye) = (T, tye)
 | 
|  |    907 |         | inf(Ts, Bound i, tye) = ((nth_elem(i, Ts) , tye)
 | 
| 0 |    908 |           handle LIST _=> raise TYPE ("loose bound variable", [], [Bound i]))
 | 
| 256 |    909 |         | inf(Ts, Var (_, T), tye) = (T, tye)
 | 
|  |    910 |         | inf(Ts, Abs (_, T, body), tye) =
 | 
|  |    911 |             let val (U, tye') = inf(T::Ts, body, tye) in  (T-->U, tye')  end
 | 
| 0 |    912 |         | inf(Ts, f$u, tye) =
 | 
| 256 |    913 |             let val (U, tyeU) = inf(Ts, u, tye);
 | 
|  |    914 |                 val (T, tyeT) = inf(Ts, f, tyeU);
 | 
| 0 |    915 |                 fun err s =
 | 
|  |    916 |                   raise TYPE(s, [inst_typ tyeT T, inst_typ tyeT U], [f$u])
 | 
| 256 |    917 |             in case T of
 | 
|  |    918 |                  Type("fun", [T1, T2]) =>
 | 
|  |    919 |                    ( (T2, unify tsig ((T1, U), tyeT))
 | 
| 0 |    920 |                      handle TUNIFY => err"type mismatch in application" )
 | 
| 256 |    921 |                | TVar _ =>
 | 
| 0 |    922 |                    let val T2 = gen_tyvar([])
 | 
|  |    923 |                    in (T2, unify tsig ((T, U-->T2), tyeT))
 | 
|  |    924 |                       handle TUNIFY => err"type mismatch in application"
 | 
|  |    925 |                    end
 | 
|  |    926 |                | _ => err"rator must have function type"
 | 
|  |    927 |            end
 | 
|  |    928 |   in inf end;
 | 
|  |    929 | 
 | 
| 256 |    930 | fun freeze_vars(Type(a, Ts)) = Type(a, map freeze_vars Ts)
 | 
| 0 |    931 |   | freeze_vars(T as TFree _) = T
 | 
| 256 |    932 |   | freeze_vars(TVar(v, S)) = TFree(Syntax.string_of_vname v, S);
 | 
| 0 |    933 | 
 | 
|  |    934 | (* Attach a type to a constant *)
 | 
| 256 |    935 | fun type_const (a, T) = Const(a, incr_tvar (new_tvar_inx()) T);
 | 
| 0 |    936 | 
 | 
|  |    937 | (*Find type of ident.  If not in table then use ident's name for tyvar
 | 
|  |    938 |   to get consistent typing.*)
 | 
| 256 |    939 | fun new_id_type a = TVar(("'"^a, new_tvar_inx()), []);
 | 
|  |    940 | fun type_of_ixn(types, ixn as (a, _)) =
 | 
| 565 |    941 |   case types ixn of Some T => freeze_vars T | None => TVar(("'"^a, ~1), []);
 | 
|  |    942 | 
 | 
|  |    943 | fun constrain (term, T) = Const (Syntax.constrainC, T --> T) $ term;
 | 
| 0 |    944 | 
 | 
| 565 |    945 | fun constrainAbs (Abs (a, _, body), T) = Abs (a, T, body)
 | 
|  |    946 |   | constrainAbs _ = sys_error "constrainAbs";
 | 
| 256 |    947 | 
 | 
| 0 |    948 | 
 | 
| 565 |    949 | (* attach_types *)
 | 
|  |    950 | 
 | 
| 0 |    951 | (*
 | 
| 256 |    952 |   Attach types to a term. Input is a "parse tree" containing dummy types.
 | 
|  |    953 |   Type constraints are translated and checked for validity wrt tsig. TVars in
 | 
|  |    954 |   constraints are frozen.
 | 
| 0 |    955 | 
 | 
| 256 |    956 |   The atoms in the resulting term satisfy the following spec:
 | 
| 0 |    957 | 
 | 
| 256 |    958 |   Const (a, T):
 | 
|  |    959 |     T is a renamed copy of the generic type of a; renaming decreases index of
 | 
|  |    960 |     all TVars by new_tvar_inx(), which is less than ~1. The index of all
 | 
|  |    961 |     TVars in the generic type must be 0 for this to work!
 | 
| 0 |    962 | 
 | 
| 256 |    963 |   Free (a, T), Var (ixn, T):
 | 
|  |    964 |     T is either the frozen default type of a or TVar (("'"^a, ~1), [])
 | 
| 0 |    965 | 
 | 
| 256 |    966 |   Abs (a, T, _):
 | 
|  |    967 |     T is either a type constraint or TVar (("'" ^ a, i), []), where i is
 | 
|  |    968 |     generated by new_tvar_inx(). Thus different abstractions can have the
 | 
|  |    969 |     bound variables of the same name but different types.
 | 
| 0 |    970 | *)
 | 
|  |    971 | 
 | 
| 565 |    972 | (* FIXME consitency of sort_env / sorts (!?) *)
 | 
| 256 |    973 | 
 | 
| 565 |    974 | fun attach_types (tsig, const_type, types, sorts) tm =
 | 
| 256 |    975 |   let
 | 
| 565 |    976 |     val sort_env = Syntax.raw_term_sorts tm;
 | 
|  |    977 |     fun def_sort xi = if_none (sorts xi) (defaultS tsig);
 | 
| 256 |    978 | 
 | 
| 565 |    979 |     fun prepareT t =
 | 
|  |    980 |       freeze_vars (cert_typ tsig (Syntax.typ_of_term sort_env def_sort t));
 | 
| 256 |    981 | 
 | 
|  |    982 |     fun add (Const (a, _)) =
 | 
| 565 |    983 |           (case const_type a of
 | 
| 256 |    984 |             Some T => type_const (a, T)
 | 
|  |    985 |           | None => raise_type ("No such constant: " ^ quote a) [] [])
 | 
|  |    986 |       | add (Free (a, _)) =
 | 
| 565 |    987 |           (case const_type a of
 | 
| 256 |    988 |             Some T => type_const (a, T)
 | 
|  |    989 |           | None => Free (a, type_of_ixn (types, (a, ~1))))
 | 
|  |    990 |       | add (Var (ixn, _)) = Var (ixn, type_of_ixn (types, ixn))
 | 
| 565 |    991 |       | add (Bound i) = Bound i
 | 
| 256 |    992 |       | add (Abs (a, _, body)) = Abs (a, new_id_type a, add body)
 | 
|  |    993 |       | add ((f as Const (a, _) $ t1) $ t2) =
 | 
|  |    994 |           if a = Syntax.constrainC then
 | 
|  |    995 |             constrain (add t1, prepareT t2)
 | 
|  |    996 |           else if a = Syntax.constrainAbsC then
 | 
|  |    997 |             constrainAbs (add t1, prepareT t2)
 | 
|  |    998 |           else add f $ add t2
 | 
|  |    999 |       | add (f $ t) = add f $ add t;
 | 
| 565 |   1000 |   in add tm end;
 | 
| 0 |   1001 | 
 | 
|  |   1002 | 
 | 
|  |   1003 | (* Post-Processing *)
 | 
|  |   1004 | 
 | 
|  |   1005 | (*Instantiation of type variables in terms*)
 | 
|  |   1006 | fun inst_types tye = map_term_types (inst_typ tye);
 | 
|  |   1007 | 
 | 
|  |   1008 | (*Delete explicit constraints -- occurrences of "_constrain" *)
 | 
| 256 |   1009 | fun unconstrain (Abs(a, T, t)) = Abs(a, T, unconstrain t)
 | 
|  |   1010 |   | unconstrain ((f as Const(a, _)) $ t) =
 | 
| 0 |   1011 |       if a=Syntax.constrainC then unconstrain t
 | 
|  |   1012 |       else unconstrain f $ unconstrain t
 | 
|  |   1013 |   | unconstrain (f$t) = unconstrain f $ unconstrain t
 | 
|  |   1014 |   | unconstrain (t) = t;
 | 
|  |   1015 | 
 | 
|  |   1016 | 
 | 
|  |   1017 | (* Turn all TVars which satisfy p into new TFrees *)
 | 
|  |   1018 | fun freeze p t =
 | 
| 256 |   1019 |   let val fs = add_term_tfree_names(t, []);
 | 
|  |   1020 |       val inxs = filter p (add_term_tvar_ixns(t, []));
 | 
| 0 |   1021 |       val vmap = inxs ~~ variantlist(map #1 inxs, fs);
 | 
| 256 |   1022 |       fun free(Type(a, Ts)) = Type(a, map free Ts)
 | 
|  |   1023 |         | free(T as TVar(v, S)) =
 | 
|  |   1024 |             (case assoc(vmap, v) of None => T | Some(a) => TFree(a, S))
 | 
| 0 |   1025 |         | free(T as TFree _) = T
 | 
|  |   1026 |   in map_term_types free t end;
 | 
|  |   1027 | 
 | 
|  |   1028 | (* Thaw all TVars that were frozen in freeze_vars *)
 | 
| 256 |   1029 | fun thaw_vars(Type(a, Ts)) = Type(a, map thaw_vars Ts)
 | 
|  |   1030 |   | thaw_vars(T as TFree(a, S)) = (case explode a of
 | 
|  |   1031 |           "?"::"'"::vn => let val ((b, i), _) = Syntax.scan_varname vn
 | 
|  |   1032 |                           in TVar(("'"^b, i), S) end
 | 
|  |   1033 |         | _ => T)
 | 
| 0 |   1034 |   | thaw_vars(T) = T;
 | 
|  |   1035 | 
 | 
|  |   1036 | 
 | 
|  |   1037 | fun restrict tye =
 | 
| 256 |   1038 |   let fun clean(tye1, ((a, i), T)) =
 | 
|  |   1039 |         if i < 0 then tye1 else ((a, i), inst_typ tye T) :: tye1
 | 
|  |   1040 |   in foldl clean ([], tye) end
 | 
| 0 |   1041 | 
 | 
|  |   1042 | 
 | 
|  |   1043 | (*Infer types for term t using tables. Check that t's type and T unify *)
 | 
|  |   1044 | 
 | 
| 565 |   1045 | fun infer_term (tsig, const_type, types, sorts, T, t) =
 | 
|  |   1046 |   let
 | 
|  |   1047 |     val u = attach_types (tsig, const_type, types, sorts) t;
 | 
|  |   1048 |     val (U, tye) = infer tsig ([], u, []);
 | 
|  |   1049 |     val uu = unconstrain u;
 | 
|  |   1050 |     val tye' = unify tsig ((T, U), tye) handle TUNIFY => raise TYPE
 | 
|  |   1051 |       ("Term does not have expected type", [T, U], [inst_types tye uu])
 | 
|  |   1052 |     val Ttye = restrict tye' (*restriction to TVars in T*)
 | 
|  |   1053 |     val all = Const("", Type("", map snd Ttye)) $ (inst_types tye' uu)
 | 
|  |   1054 |       (*all is a dummy term which contains all exported TVars*)
 | 
|  |   1055 |     val Const(_, Type(_, Ts)) $ u'' =
 | 
|  |   1056 |       map_term_types thaw_vars (freeze (fn (_, i) => i < 0) all)
 | 
|  |   1057 |       (*turn all internally generated TVars into TFrees
 | 
|  |   1058 |         and thaw all initially frozen TVars*)
 | 
|  |   1059 |   in
 | 
|  |   1060 |     (u'', (map fst Ttye) ~~ Ts)
 | 
|  |   1061 |   end;
 | 
| 0 |   1062 | 
 | 
| 621 |   1063 | fun infer_types args = (tyinit (); infer_term args);
 | 
| 0 |   1064 | 
 | 
|  |   1065 | 
 | 
|  |   1066 | end;
 | 
| 256 |   1067 | 
 |