| author | nipkow | 
| Thu, 31 Jul 2003 14:01:04 +0200 | |
| changeset 14138 | ca5029d391d1 | 
| parent 13666 | a2730043029b | 
| child 14790 | 0d984ee030a1 | 
| permissions | -rw-r--r-- | 
| 256 | 1  | 
(* Title: Pure/type.ML  | 
| 0 | 2  | 
ID: $Id$  | 
| 416 | 3  | 
Author: Tobias Nipkow & Lawrence C Paulson  | 
| 0 | 4  | 
|
| 2964 | 5  | 
Type signatures, unification of types, interface to type inference.  | 
| 0 | 6  | 
*)  | 
7  | 
||
8  | 
signature TYPE =  | 
|
| 2964 | 9  | 
sig  | 
| 7641 | 10  | 
(*TFrees and TVars*)  | 
| 621 | 11  | 
val no_tvars: typ -> typ  | 
12  | 
val varifyT: typ -> typ  | 
|
13  | 
val unvarifyT: typ -> typ  | 
|
| 12501 | 14  | 
val varify: term * string list -> term * (string * indexname) list  | 
| 10495 | 15  | 
val freeze_thaw_type : typ -> typ * (typ -> typ)  | 
| 
3411
 
163f8f4a42d7
Removal of freeze_vars and thaw_vars.  New freeze_thaw
 
paulson 
parents: 
3175 
diff
changeset
 | 
16  | 
val freeze_thaw : term -> term * (term -> term)  | 
| 2964 | 17  | 
|
18  | 
(*type signatures*)  | 
|
| 0 | 19  | 
type type_sig  | 
| 200 | 20  | 
val rep_tsig: type_sig ->  | 
| 256 | 21  | 
    {classes: class list,
 | 
| 7069 | 22  | 
classrel: Sorts.classrel,  | 
| 256 | 23  | 
default: sort,  | 
| 7069 | 24  | 
tycons: int Symtab.table,  | 
| 7641 | 25  | 
log_types: string list,  | 
26  | 
univ_witness: (typ * sort) option,  | 
|
| 7069 | 27  | 
abbrs: (string list * typ) Symtab.table,  | 
28  | 
arities: Sorts.arities}  | 
|
| 7641 | 29  | 
val classes: type_sig -> class list  | 
| 0 | 30  | 
val defaultS: type_sig -> sort  | 
| 2964 | 31  | 
val logical_types: type_sig -> string list  | 
| 7641 | 32  | 
val univ_witness: type_sig -> (typ * sort) option  | 
| 2964 | 33  | 
val subsort: type_sig -> sort * sort -> bool  | 
34  | 
val eq_sort: type_sig -> sort * sort -> bool  | 
|
35  | 
val norm_sort: type_sig -> sort -> sort  | 
|
| 8899 | 36  | 
val cert_class: type_sig -> class -> class  | 
37  | 
val cert_sort: type_sig -> sort -> sort  | 
|
| 7641 | 38  | 
val witness_sorts: type_sig -> sort list -> sort list -> (typ * sort) list  | 
| 2964 | 39  | 
val rem_sorts: typ -> typ  | 
| 416 | 40  | 
val tsig0: type_sig  | 
| 621 | 41  | 
val ext_tsig_classes: type_sig -> (class * class list) list -> type_sig  | 
| 2964 | 42  | 
val ext_tsig_classrel: type_sig -> (class * class) list -> type_sig  | 
| 422 | 43  | 
val ext_tsig_defsort: type_sig -> sort -> type_sig  | 
| 582 | 44  | 
val ext_tsig_types: type_sig -> (string * int) list -> type_sig  | 
| 621 | 45  | 
val ext_tsig_abbrs: type_sig -> (string * string list * typ) list -> type_sig  | 
| 963 | 46  | 
val ext_tsig_arities: type_sig -> (string * sort list * sort)list -> type_sig  | 
| 256 | 47  | 
val merge_tsigs: type_sig * type_sig -> type_sig  | 
| 2964 | 48  | 
val typ_errors: type_sig -> typ * string list -> string list  | 
| 256 | 49  | 
val cert_typ: type_sig -> typ -> typ  | 
| 9504 | 50  | 
val cert_typ_no_norm: type_sig -> typ -> typ  | 
| 256 | 51  | 
val norm_typ: type_sig -> typ -> typ  | 
| 7069 | 52  | 
val norm_term: type_sig -> term -> term  | 
| 256 | 53  | 
val inst_term_tvars: type_sig * (indexname * typ) list -> term -> term  | 
| 
8406
 
a217b0cd304d
Type.unify and Type.typ_match now use Vartab instead of association lists.
 
berghofe 
parents: 
7641 
diff
changeset
 | 
54  | 
val inst_typ_tvars: type_sig * (indexname * typ) list -> typ -> typ  | 
| 2964 | 55  | 
|
56  | 
(*type matching*)  | 
|
57  | 
exception TYPE_MATCH  | 
|
| 
8406
 
a217b0cd304d
Type.unify and Type.typ_match now use Vartab instead of association lists.
 
berghofe 
parents: 
7641 
diff
changeset
 | 
58  | 
val typ_match: type_sig -> typ Vartab.table * (typ * typ)  | 
| 
 
a217b0cd304d
Type.unify and Type.typ_match now use Vartab instead of association lists.
 
berghofe 
parents: 
7641 
diff
changeset
 | 
59  | 
-> typ Vartab.table  | 
| 2964 | 60  | 
val typ_instance: type_sig * typ * typ -> bool  | 
| 3175 | 61  | 
val of_sort: type_sig -> typ * sort -> bool  | 
| 2964 | 62  | 
|
63  | 
(*type unification*)  | 
|
64  | 
exception TUNIFY  | 
|
| 12528 | 65  | 
val unify: type_sig -> typ Vartab.table * int -> typ * typ -> typ Vartab.table * int  | 
| 450 | 66  | 
val raw_unify: typ * typ -> bool  | 
| 0 | 67  | 
|
| 2964 | 68  | 
(*type inference*)  | 
| 3790 | 69  | 
val get_sort: type_sig -> (indexname -> sort option) -> (sort -> sort)  | 
70  | 
-> (indexname * sort) list -> indexname -> sort  | 
|
| 2964 | 71  | 
val constrain: term -> typ -> term  | 
| 12528 | 72  | 
val param: int -> string * sort -> typ  | 
73  | 
val paramify_dummies: int * typ -> int * typ  | 
|
| 2979 | 74  | 
val infer_types: (term -> Pretty.T) -> (typ -> Pretty.T)  | 
75  | 
-> type_sig -> (string -> typ option) -> (indexname -> typ option)  | 
|
| 3790 | 76  | 
-> (indexname -> sort option) -> (string -> string) -> (typ -> typ)  | 
77  | 
-> (sort -> sort) -> string list -> bool -> typ list -> term list  | 
|
| 2964 | 78  | 
-> term list * (indexname * typ) list  | 
79  | 
end;  | 
|
80  | 
||
81  | 
structure Type: TYPE =  | 
|
| 0 | 82  | 
struct  | 
83  | 
||
| 2964 | 84  | 
|
| 7641 | 85  | 
(*** TFrees and TVars ***)  | 
| 621 | 86  | 
|
87  | 
fun no_tvars T =  | 
|
| 12501 | 88  | 
(case typ_tvars T of [] => T  | 
89  | 
  | vs => raise TYPE ("Illegal schematic type variable(s): " ^
 | 
|
90  | 
commas (map (Syntax.string_of_vname o #1) vs), [T], []));  | 
|
| 621 | 91  | 
|
| 7641 | 92  | 
|
| 2964 | 93  | 
(* varify, unvarify *)  | 
| 621 | 94  | 
|
| 2964 | 95  | 
val varifyT = map_type_tfree (fn (a, S) => TVar ((a, 0), S));  | 
96  | 
||
| 621 | 97  | 
fun unvarifyT (Type (a, Ts)) = Type (a, map unvarifyT Ts)  | 
98  | 
| unvarifyT (TVar ((a, 0), S)) = TFree (a, S)  | 
|
99  | 
| unvarifyT T = T;  | 
|
100  | 
||
101  | 
fun varify (t, fixed) =  | 
|
102  | 
let  | 
|
103  | 
val fs = add_term_tfree_names (t, []) \\ fixed;  | 
|
104  | 
val ixns = add_term_tvar_ixns (t, []);  | 
|
| 12501 | 105  | 
val fmap = fs ~~ map (rpair 0) (variantlist (fs, map #1 ixns))  | 
| 2964 | 106  | 
fun thaw (f as (a, S)) =  | 
107  | 
(case assoc (fmap, a) of  | 
|
108  | 
None => TFree f  | 
|
| 12501 | 109  | 
| Some b => TVar (b, S));  | 
110  | 
in (map_term_types (map_type_tfree thaw) t, fmap) end;  | 
|
| 2964 | 111  | 
|
112  | 
||
| 7641 | 113  | 
(* freeze_thaw: freeze TVars in a term; return the "thaw" inverse *)  | 
| 
3411
 
163f8f4a42d7
Removal of freeze_vars and thaw_vars.  New freeze_thaw
 
paulson 
parents: 
3175 
diff
changeset
 | 
114  | 
|
| 7641 | 115  | 
local  | 
116  | 
||
117  | 
fun new_name (ix, (pairs,used)) =  | 
|
| 
3411
 
163f8f4a42d7
Removal of freeze_vars and thaw_vars.  New freeze_thaw
 
paulson 
parents: 
3175 
diff
changeset
 | 
118  | 
let val v = variant used (string_of_indexname ix)  | 
| 
 
163f8f4a42d7
Removal of freeze_vars and thaw_vars.  New freeze_thaw
 
paulson 
parents: 
3175 
diff
changeset
 | 
119  | 
in ((ix,v)::pairs, v::used) end;  | 
| 621 | 120  | 
|
| 7641 | 121  | 
fun freeze_one alist (ix,sort) =  | 
| 3790 | 122  | 
TFree (the (assoc (alist, ix)), sort)  | 
| 4142 | 123  | 
handle OPTION =>  | 
| 3790 | 124  | 
      raise TYPE ("Failure during freezing of ?" ^ string_of_indexname ix, [], []);
 | 
| 2964 | 125  | 
|
| 7641 | 126  | 
fun thaw_one alist (a,sort) = TVar (the (assoc (alist,a)), sort)  | 
| 4142 | 127  | 
handle OPTION => TFree(a,sort);  | 
| 416 | 128  | 
|
| 10495 | 129  | 
in  | 
130  | 
||
131  | 
(*this sort of code could replace unvarifyT*)  | 
|
| 7641 | 132  | 
fun freeze_thaw_type T =  | 
133  | 
let  | 
|
134  | 
val used = add_typ_tfree_names (T, [])  | 
|
135  | 
and tvars = map #1 (add_typ_tvars (T, []));  | 
|
136  | 
val (alist, _) = foldr new_name (tvars, ([], used));  | 
|
137  | 
in (map_type_tvar (freeze_one alist) T, map_type_tfree (thaw_one (map swap alist))) end;  | 
|
138  | 
||
| 
3411
 
163f8f4a42d7
Removal of freeze_vars and thaw_vars.  New freeze_thaw
 
paulson 
parents: 
3175 
diff
changeset
 | 
139  | 
fun freeze_thaw t =  | 
| 7641 | 140  | 
let  | 
141  | 
val used = it_term_types add_typ_tfree_names (t, [])  | 
|
142  | 
and tvars = map #1 (it_term_types add_typ_tvars (t, []));  | 
|
143  | 
val (alist, _) = foldr new_name (tvars, ([], used));  | 
|
144  | 
in  | 
|
145  | 
(case alist of  | 
|
146  | 
[] => (t, fn x => x) (*nothing to do!*)  | 
|
147  | 
| _ => (map_term_types (map_type_tvar (freeze_one alist)) t,  | 
|
148  | 
map_term_types (map_type_tfree (thaw_one (map swap alist)))))  | 
|
149  | 
end;  | 
|
150  | 
||
151  | 
end;  | 
|
152  | 
||
| 256 | 153  | 
|
154  | 
||
| 416 | 155  | 
(*** type signatures ***)  | 
| 256 | 156  | 
|
| 2964 | 157  | 
(* type type_sig *)  | 
158  | 
||
| 256 | 159  | 
(*  | 
| 7641 | 160  | 
classes: list of all declared classes;  | 
161  | 
classrel: (see Pure/sorts.ML)  | 
|
162  | 
default: default sort attached to all unconstrained type vars;  | 
|
163  | 
tycons: table of all declared types with the number of their arguments;  | 
|
164  | 
log_types: list of logical type constructors sorted by number of arguments;  | 
|
165  | 
univ_witness: type witnessing non-emptiness of least sort  | 
|
166  | 
abbrs: table of type abbreviations;  | 
|
167  | 
arities: (see Pure/sorts.ML)  | 
|
| 0 | 168  | 
*)  | 
169  | 
||
| 256 | 170  | 
datatype type_sig =  | 
171  | 
  TySg of {
 | 
|
172  | 
classes: class list,  | 
|
| 7069 | 173  | 
classrel: Sorts.classrel,  | 
| 256 | 174  | 
default: sort,  | 
| 7069 | 175  | 
tycons: int Symtab.table,  | 
| 7641 | 176  | 
log_types: string list,  | 
177  | 
univ_witness: (typ * sort) option,  | 
|
| 7069 | 178  | 
abbrs: (string list * typ) Symtab.table,  | 
179  | 
arities: Sorts.arities};  | 
|
| 256 | 180  | 
|
| 189 | 181  | 
fun rep_tsig (TySg comps) = comps;  | 
| 0 | 182  | 
|
| 7641 | 183  | 
fun classes (TySg {classes = cs, ...}) = cs;
 | 
| 256 | 184  | 
fun defaultS (TySg {default, ...}) = default;
 | 
| 7641 | 185  | 
fun logical_types (TySg {log_types, ...}) = log_types;
 | 
186  | 
fun univ_witness (TySg {univ_witness, ...}) = univ_witness;
 | 
|
| 8899 | 187  | 
|
188  | 
||
189  | 
(* error messages *)  | 
|
190  | 
||
191  | 
fun undeclared_class c = "Undeclared class: " ^ quote c;  | 
|
192  | 
fun undeclared_classes cs = "Undeclared class(es): " ^ commas_quote cs;  | 
|
193  | 
||
194  | 
fun err_undeclared_class s = error (undeclared_class s);  | 
|
195  | 
||
196  | 
fun err_dup_classes cs =  | 
|
197  | 
  error ("Duplicate declaration of class(es): " ^ commas_quote cs);
 | 
|
198  | 
||
199  | 
fun undeclared_type c = "Undeclared type constructor: " ^ quote c;  | 
|
200  | 
||
201  | 
fun err_neg_args c =  | 
|
202  | 
  error ("Negative number of arguments of type constructor: " ^ quote c);
 | 
|
203  | 
||
204  | 
fun err_dup_tycon c =  | 
|
205  | 
  error ("Duplicate declaration of type constructor: " ^ quote c);
 | 
|
206  | 
||
207  | 
fun dup_tyabbrs ts =  | 
|
208  | 
"Duplicate declaration of type abbreviation(s): " ^ commas_quote ts;  | 
|
209  | 
||
210  | 
fun ty_confl c = "Conflicting type constructor and abbreviation: " ^ quote c;  | 
|
| 2964 | 211  | 
|
212  | 
||
213  | 
(* sorts *)  | 
|
214  | 
||
215  | 
fun subsort (TySg {classrel, ...}) = Sorts.sort_le classrel;
 | 
|
216  | 
fun eq_sort (TySg {classrel, ...}) = Sorts.sort_eq classrel;
 | 
|
217  | 
fun norm_sort (TySg {classrel, ...}) = Sorts.norm_sort classrel;
 | 
|
218  | 
||
| 8899 | 219  | 
fun cert_class (TySg {classes, ...}) c =
 | 
220  | 
if c mem_string classes then c else raise TYPE (undeclared_class c, [], []);  | 
|
221  | 
||
222  | 
fun cert_sort tsig S = norm_sort tsig (map (cert_class tsig) S);  | 
|
223  | 
||
| 7641 | 224  | 
fun witness_sorts (tsig as TySg {classrel, arities, log_types, ...}) =
 | 
225  | 
Sorts.witness_sorts (classrel, arities, log_types);  | 
|
| 2964 | 226  | 
|
227  | 
fun rem_sorts (Type (a, tys)) = Type (a, map rem_sorts tys)  | 
|
228  | 
| rem_sorts (TFree (x, _)) = TFree (x, [])  | 
|
229  | 
| rem_sorts (TVar (xi, _)) = TVar (xi, []);  | 
|
230  | 
||
| 256 | 231  | 
|
| 8899 | 232  | 
(* FIXME err_undeclared_class! *)  | 
| 0 | 233  | 
(* 'leq' checks the partial order on classes according to the  | 
| 7641 | 234  | 
statements in classrel 'a'  | 
| 0 | 235  | 
*)  | 
236  | 
||
| 7069 | 237  | 
fun less a (C, D) = case Symtab.lookup (a, C) of  | 
| 
2182
 
29e56f003599
Removal of polymorphic equality via mem, subset, eq_set, etc
 
paulson 
parents: 
2148 
diff
changeset
 | 
238  | 
Some ss => D mem_string ss  | 
| 8899 | 239  | 
| None => err_undeclared_class C;  | 
| 0 | 240  | 
|
| 256 | 241  | 
fun leq a (C, D) = C = D orelse less a (C, D);  | 
| 0 | 242  | 
|
243  | 
||
244  | 
||
| 2964 | 245  | 
(* FIXME *)  | 
| 200 | 246  | 
(*Instantiation of type variables in types*)  | 
247  | 
(*Pre: instantiations obey restrictions! *)  | 
|
248  | 
fun inst_typ tye =  | 
|
| 
949
 
83c588d6fee9
Changed treatment of during type inference internally generated type
 
nipkow 
parents: 
621 
diff
changeset
 | 
249  | 
let fun inst(var as (v, _)) = case assoc(tye, v) of  | 
| 
 
83c588d6fee9
Changed treatment of during type inference internally generated type
 
nipkow 
parents: 
621 
diff
changeset
 | 
250  | 
Some U => inst_typ tye U  | 
| 
 
83c588d6fee9
Changed treatment of during type inference internally generated type
 
nipkow 
parents: 
621 
diff
changeset
 | 
251  | 
| None => TVar(var)  | 
| 
 
83c588d6fee9
Changed treatment of during type inference internally generated type
 
nipkow 
parents: 
621 
diff
changeset
 | 
252  | 
in map_type_tvar inst end;  | 
| 0 | 253  | 
|
254  | 
||
| 3175 | 255  | 
|
| 7641 | 256  | 
fun of_sort (TySg {classrel, arities, ...}) = Sorts.of_sort (classrel, arities);
 | 
| 3175 | 257  | 
|
258  | 
fun check_has_sort (tsig, T, S) =  | 
|
259  | 
if of_sort tsig (T, S) then ()  | 
|
| 3790 | 260  | 
  else raise TYPE ("Type not of sort " ^ Sorts.str_of_sort S, [T], []);
 | 
| 0 | 261  | 
|
262  | 
||
263  | 
(*Instantiation of type variables in types *)  | 
|
| 256 | 264  | 
fun inst_typ_tvars(tsig, tye) =  | 
| 
949
 
83c588d6fee9
Changed treatment of during type inference internally generated type
 
nipkow 
parents: 
621 
diff
changeset
 | 
265  | 
let fun inst(var as (v, S)) = case assoc(tye, v) of  | 
| 
 
83c588d6fee9
Changed treatment of during type inference internally generated type
 
nipkow 
parents: 
621 
diff
changeset
 | 
266  | 
Some U => (check_has_sort(tsig, U, S); U)  | 
| 
 
83c588d6fee9
Changed treatment of during type inference internally generated type
 
nipkow 
parents: 
621 
diff
changeset
 | 
267  | 
| None => TVar(var)  | 
| 
 
83c588d6fee9
Changed treatment of during type inference internally generated type
 
nipkow 
parents: 
621 
diff
changeset
 | 
268  | 
in map_type_tvar inst end;  | 
| 0 | 269  | 
|
270  | 
(*Instantiation of type variables in terms *)  | 
|
| 2617 | 271  | 
fun inst_term_tvars (_,[]) t = t  | 
272  | 
| inst_term_tvars arg t = map_term_types (inst_typ_tvars arg) t;  | 
|
| 200 | 273  | 
|
274  | 
||
| 7069 | 275  | 
(* norm_typ, norm_term *)  | 
| 200 | 276  | 
|
| 2991 | 277  | 
(*expand abbreviations and normalize sorts*)  | 
278  | 
fun norm_typ (tsig as TySg {abbrs, ...}) ty =
 | 
|
| 256 | 279  | 
let  | 
| 621 | 280  | 
val idx = maxidx_of_typ ty + 1;  | 
281  | 
||
| 2991 | 282  | 
fun norm (Type (a, Ts)) =  | 
| 7069 | 283  | 
(case Symtab.lookup (abbrs, a) of  | 
| 2991 | 284  | 
Some (vs, U) => norm (inst_typ (map (rpair idx) vs ~~ Ts) (incr_tvar idx U))  | 
285  | 
| None => Type (a, map norm Ts))  | 
|
286  | 
| norm (TFree (x, S)) = TFree (x, norm_sort tsig S)  | 
|
287  | 
| norm (TVar (xi, S)) = TVar (xi, norm_sort tsig S);  | 
|
| 7069 | 288  | 
|
289  | 
val ty' = norm ty;  | 
|
290  | 
in if ty = ty' then ty else ty' end; (*dumb tuning to avoid copying*)  | 
|
| 256 | 291  | 
|
| 7069 | 292  | 
fun norm_term tsig t =  | 
293  | 
let val t' = map_term_types (norm_typ tsig) t  | 
|
294  | 
in if t = t' then t else t' end; (*dumb tuning to avoid copying*)  | 
|
| 256 | 295  | 
|
296  | 
||
297  | 
||
298  | 
(** build type signatures **)  | 
|
299  | 
||
| 7641 | 300  | 
fun make_tsig (classes, classrel, default, tycons, log_types, univ_witness, abbrs, arities) =  | 
301  | 
  TySg {classes = classes, classrel = classrel, default = default, tycons = tycons,
 | 
|
302  | 
log_types = log_types, univ_witness = univ_witness, abbrs = abbrs, arities = arities};  | 
|
303  | 
||
304  | 
fun rebuild_tsig (TySg {classes, classrel, default, tycons, log_types = _, univ_witness = _, abbrs, arities}) =
 | 
|
305  | 
let  | 
|
306  | 
fun log_class c = Sorts.class_le classrel (c, logicC);  | 
|
307  | 
fun log_type (t, _) = exists (log_class o #1) (Symtab.lookup_multi (arities, t));  | 
|
308  | 
val ts = filter log_type (Symtab.dest tycons);  | 
|
| 416 | 309  | 
|
| 7641 | 310  | 
val log_types = map #1 (Library.sort (Library.int_ord o pairself #2) ts);  | 
311  | 
val univ_witness =  | 
|
312  | 
(case Sorts.witness_sorts (classrel, arities, log_types) [] [classes] of  | 
|
313  | 
[w] => Some w | _ => None);  | 
|
314  | 
in make_tsig (classes, classrel, default, tycons, log_types, univ_witness, abbrs, arities) end;  | 
|
| 256 | 315  | 
|
| 7641 | 316  | 
val tsig0 =  | 
317  | 
make_tsig ([], Symtab.empty, [], Symtab.empty, [], None, Symtab.empty, Symtab.empty)  | 
|
318  | 
|> rebuild_tsig;  | 
|
| 
1215
 
a206f722bef9
added nonempty_sort (a somewhat braindead version!);
 
wenzelm 
parents: 
963 
diff
changeset
 | 
319  | 
|
| 
 
a206f722bef9
added nonempty_sort (a somewhat braindead version!);
 
wenzelm 
parents: 
963 
diff
changeset
 | 
320  | 
|
| 416 | 321  | 
(* typ_errors *)  | 
| 256 | 322  | 
|
| 416 | 323  | 
(*check validity of (not necessarily normal) type; accumulate error messages*)  | 
| 256 | 324  | 
|
| 416 | 325  | 
fun typ_errors tsig (typ, errors) =  | 
| 256 | 326  | 
let  | 
| 7069 | 327  | 
    val {classes, tycons, abbrs, ...} = rep_tsig tsig;
 | 
| 416 | 328  | 
|
329  | 
fun class_err (errs, c) =  | 
|
| 
2182
 
29e56f003599
Removal of polymorphic equality via mem, subset, eq_set, etc
 
paulson 
parents: 
2148 
diff
changeset
 | 
330  | 
if c mem_string classes then errs  | 
| 8899 | 331  | 
else undeclared_class c ins_string errs;  | 
| 256 | 332  | 
|
333  | 
val sort_err = foldl class_err;  | 
|
| 0 | 334  | 
|
| 7069 | 335  | 
fun typ_errs (errs, Type (c, Us)) =  | 
| 256 | 336  | 
let  | 
| 7069 | 337  | 
val errs' = foldl typ_errs (errs, Us);  | 
| 256 | 338  | 
fun nargs n =  | 
339  | 
if n = length Us then errs'  | 
|
| 
2182
 
29e56f003599
Removal of polymorphic equality via mem, subset, eq_set, etc
 
paulson 
parents: 
2148 
diff
changeset
 | 
340  | 
              else ("Wrong number of arguments: " ^ quote c) ins_string errs';
 | 
| 256 | 341  | 
in  | 
| 7069 | 342  | 
(case Symtab.lookup (tycons, c) of  | 
| 256 | 343  | 
Some n => nargs n  | 
344  | 
| None =>  | 
|
| 7069 | 345  | 
(case Symtab.lookup (abbrs, c) of  | 
| 256 | 346  | 
Some (vs, _) => nargs (length vs)  | 
| 8899 | 347  | 
| None => undeclared_type c ins_string errs))  | 
| 256 | 348  | 
end  | 
| 7069 | 349  | 
| typ_errs (errs, TFree (_, S)) = sort_err (errs, S)  | 
350  | 
| typ_errs (errs, TVar ((x, i), S)) =  | 
|
| 416 | 351  | 
if i < 0 then  | 
| 
2182
 
29e56f003599
Removal of polymorphic equality via mem, subset, eq_set, etc
 
paulson 
parents: 
2148 
diff
changeset
 | 
352  | 
          ("Negative index for TVar " ^ quote x) ins_string sort_err (errs, S)
 | 
| 416 | 353  | 
else sort_err (errs, S);  | 
| 7069 | 354  | 
in typ_errs (errors, typ) end;  | 
| 256 | 355  | 
|
356  | 
||
| 8899 | 357  | 
(* cert_typ *) (*exception TYPE*)  | 
| 256 | 358  | 
|
| 9504 | 359  | 
fun cert_typ_no_norm tsig T =  | 
| 2964 | 360  | 
(case typ_errors tsig (T, []) of  | 
| 9504 | 361  | 
[] => T  | 
| 3790 | 362  | 
| errs => raise TYPE (cat_lines errs, [T], []));  | 
| 256 | 363  | 
|
| 9504 | 364  | 
fun cert_typ tsig T = norm_typ tsig (cert_typ_no_norm tsig T);  | 
365  | 
||
| 256 | 366  | 
|
367  | 
||
| 422 | 368  | 
(** merge type signatures **)  | 
| 256 | 369  | 
|
| 7641 | 370  | 
(* merge classrel *)  | 
| 0 | 371  | 
|
| 422 | 372  | 
fun assoc_union (as1, []) = as1  | 
373  | 
| assoc_union (as1, (key, l2) :: as2) =  | 
|
| 
2182
 
29e56f003599
Removal of polymorphic equality via mem, subset, eq_set, etc
 
paulson 
parents: 
2148 
diff
changeset
 | 
374  | 
(case assoc_string (as1, key) of  | 
| 7641 | 375  | 
Some l1 => assoc_union (overwrite (as1, (key, l1 union_string l2)), as2)  | 
| 422 | 376  | 
| None => assoc_union ((key, l2) :: as1, as2));  | 
| 0 | 377  | 
|
| 2964 | 378  | 
fun merge_classrel (classrel1, classrel2) =  | 
| 7069 | 379  | 
let  | 
380  | 
val classrel = transitive_closure (assoc_union (Symtab.dest classrel1, Symtab.dest classrel2))  | 
|
| 
2182
 
29e56f003599
Removal of polymorphic equality via mem, subset, eq_set, etc
 
paulson 
parents: 
2148 
diff
changeset
 | 
381  | 
in  | 
| 2964 | 382  | 
if exists (op mem_string) classrel then  | 
| 422 | 383  | 
      error ("Cyclic class structure!")   (* FIXME improve msg, raise TERM *)
 | 
| 7069 | 384  | 
else Symtab.make classrel  | 
| 416 | 385  | 
end;  | 
386  | 
||
387  | 
||
| 422 | 388  | 
(* coregularity *)  | 
| 0 | 389  | 
|
| 7641 | 390  | 
local  | 
391  | 
||
| 0 | 392  | 
(* 'is_unique_decl' checks if there exists just one declaration t:(Ss)C *)  | 
393  | 
||
| 963 | 394  | 
fun is_unique_decl ars (t,(C,w)) = case assoc (ars, C) of  | 
| 0 | 395  | 
Some(w1) => if w = w1 then () else  | 
| 256 | 396  | 
        error("There are two declarations\n" ^
 | 
| 2964 | 397  | 
Sorts.str_of_arity(t, w, [C]) ^ " and\n" ^  | 
398  | 
Sorts.str_of_arity(t, w1, [C]) ^ "\n" ^  | 
|
| 0 | 399  | 
"with the same result class.")  | 
400  | 
| None => ();  | 
|
401  | 
||
| 963 | 402  | 
(* 'coreg' checks if there are two declarations t:(Ss1)C1 and t:(Ss2)C2  | 
| 0 | 403  | 
such that C1 >= C2 then Ss1 >= Ss2 (elementwise) *)  | 
404  | 
||
| 963 | 405  | 
fun coreg_err(t, (C1,w1), (C2,w2)) =  | 
| 2964 | 406  | 
    error("Declarations " ^ Sorts.str_of_arity(t, w1, [C1]) ^ " and "
 | 
407  | 
^ Sorts.str_of_arity(t, w2, [C2]) ^ " are in conflict");  | 
|
| 0 | 408  | 
|
| 2964 | 409  | 
fun coreg classrel (t, Cw1) =  | 
410  | 
let  | 
|
411  | 
fun check1(Cw1 as (C1,w1), Cw2 as (C2,w2)) =  | 
|
412  | 
if leq classrel (C1,C2) then  | 
|
413  | 
if Sorts.sorts_le classrel (w1,w2) then ()  | 
|
414  | 
else coreg_err(t, Cw1, Cw2)  | 
|
415  | 
else ()  | 
|
416  | 
fun check(Cw2) = (check1(Cw1,Cw2); check1(Cw2,Cw1))  | 
|
| 963 | 417  | 
in seq check end;  | 
| 0 | 418  | 
|
| 7641 | 419  | 
in  | 
420  | 
||
| 2964 | 421  | 
fun add_arity classrel ars (tCw as (_,Cw)) =  | 
422  | 
(is_unique_decl ars tCw; coreg classrel tCw ars; Cw ins ars);  | 
|
| 0 | 423  | 
|
| 7641 | 424  | 
end;  | 
| 0 | 425  | 
|
426  | 
||
| 963 | 427  | 
(* 'merge_arities' builds the union of two 'arities' lists;  | 
| 422 | 428  | 
it only checks the two restriction conditions and inserts afterwards  | 
429  | 
all elements of the second list into the first one *)  | 
|
430  | 
||
| 7641 | 431  | 
local  | 
432  | 
||
| 7069 | 433  | 
fun merge_arities_aux classrel =  | 
| 2964 | 434  | 
let fun test_ar t (ars1, sw) = add_arity classrel ars1 (t,sw);  | 
| 422 | 435  | 
|
| 963 | 436  | 
fun merge_c (arities1, (c as (t, ars2))) = case assoc (arities1, t) of  | 
437  | 
Some(ars1) =>  | 
|
438  | 
let val ars = foldl (test_ar t) (ars1, ars2)  | 
|
439  | 
in overwrite (arities1, (t,ars)) end  | 
|
440  | 
| None => c::arities1  | 
|
| 422 | 441  | 
in foldl merge_c end;  | 
442  | 
||
| 7641 | 443  | 
in  | 
444  | 
||
| 7069 | 445  | 
fun merge_arities classrel (a1, a2) =  | 
446  | 
Symtab.make (merge_arities_aux classrel (Symtab.dest a1, Symtab.dest a2));  | 
|
| 422 | 447  | 
|
| 7641 | 448  | 
end;  | 
449  | 
||
450  | 
||
451  | 
(* tycons *)  | 
|
452  | 
||
453  | 
fun varying_decls t =  | 
|
454  | 
  error ("Type constructor " ^ quote t ^ " has varying number of arguments");
 | 
|
455  | 
||
| 7069 | 456  | 
fun add_tycons (tycons, tn as (t,n)) =  | 
457  | 
(case Symtab.lookup (tycons, t) of  | 
|
458  | 
Some m => if m = n then tycons else varying_decls t  | 
|
459  | 
| None => Symtab.update (tn, tycons));  | 
|
460  | 
||
| 7641 | 461  | 
|
462  | 
(* merge_abbrs *)  | 
|
463  | 
||
| 7069 | 464  | 
fun merge_abbrs abbrs =  | 
465  | 
Symtab.merge (op =) abbrs handle Symtab.DUPS dups => raise TERM (dup_tyabbrs dups, []);  | 
|
| 422 | 466  | 
|
467  | 
||
| 7641 | 468  | 
(* merge_tsigs *)  | 
| 422 | 469  | 
|
| 7641 | 470  | 
fun merge_tsigs  | 
471  | 
 (TySg {classes = classes1, default = default1, classrel = classrel1, tycons = tycons1,
 | 
|
472  | 
log_types = _, univ_witness = _, arities = arities1, abbrs = abbrs1},  | 
|
473  | 
  TySg {classes = classes2, default = default2, classrel = classrel2, tycons = tycons2,
 | 
|
474  | 
log_types = _, univ_witness = _, arities = arities2, abbrs = abbrs2}) =  | 
|
475  | 
let  | 
|
476  | 
val classes' = classes1 union_string classes2;  | 
|
477  | 
val classrel' = merge_classrel (classrel1, classrel2);  | 
|
478  | 
val arities' = merge_arities classrel' (arities1, arities2);  | 
|
479  | 
val tycons' = foldl add_tycons (tycons1, Symtab.dest tycons2);  | 
|
480  | 
val default' = Sorts.norm_sort classrel' (default1 @ default2);  | 
|
481  | 
val abbrs' = merge_abbrs (abbrs1, abbrs2);  | 
|
482  | 
in  | 
|
483  | 
make_tsig (classes', classrel', default', tycons', [], None, abbrs', arities')  | 
|
484  | 
|> rebuild_tsig  | 
|
485  | 
end;  | 
|
| 422 | 486  | 
|
487  | 
||
488  | 
||
489  | 
(*** extend type signatures ***)  | 
|
490  | 
||
| 2964 | 491  | 
(** add classes and classrel relations **)  | 
| 422 | 492  | 
|
493  | 
fun add_classes classes cs =  | 
|
| 
2182
 
29e56f003599
Removal of polymorphic equality via mem, subset, eq_set, etc
 
paulson 
parents: 
2148 
diff
changeset
 | 
494  | 
(case cs inter_string classes of  | 
| 422 | 495  | 
[] => cs @ classes  | 
496  | 
| dups => err_dup_classes cs);  | 
|
497  | 
||
498  | 
||
| 2964 | 499  | 
(*'add_classrel' adds a tuple consisting of a new class (the new class has  | 
| 422 | 500  | 
already been inserted into the 'classes' list) and its superclasses (they  | 
| 2964 | 501  | 
must be declared in 'classes' too) to the 'classrel' list of the given type  | 
| 422 | 502  | 
signature; furthermore all inherited superclasses according to the  | 
503  | 
superclasses brought with are inserted and there is a check that there are  | 
|
504  | 
no cycles (i.e. C <= D <= C, with C <> D);*)  | 
|
505  | 
||
| 2964 | 506  | 
fun add_classrel classes (classrel, (s, ges)) =  | 
| 621 | 507  | 
let  | 
| 2964 | 508  | 
fun upd (classrel, s') =  | 
| 
2182
 
29e56f003599
Removal of polymorphic equality via mem, subset, eq_set, etc
 
paulson 
parents: 
2148 
diff
changeset
 | 
509  | 
if s' mem_string classes then  | 
| 7069 | 510  | 
let val ges' = the (Symtab.lookup (classrel, s))  | 
511  | 
in case Symtab.lookup (classrel, s') of  | 
|
| 
2182
 
29e56f003599
Removal of polymorphic equality via mem, subset, eq_set, etc
 
paulson 
parents: 
2148 
diff
changeset
 | 
512  | 
Some sups => if s mem_string sups  | 
| 422 | 513  | 
                           then error(" Cycle :" ^ s^" <= "^ s'^" <= "^ s )
 | 
| 7069 | 514  | 
else Symtab.update ((s, sups union_string ges'), classrel)  | 
| 2964 | 515  | 
| None => classrel  | 
| 621 | 516  | 
end  | 
| 8899 | 517  | 
else err_undeclared_class s'  | 
| 7069 | 518  | 
in foldl upd (Symtab.update ((s, ges), classrel), ges) end;  | 
| 422 | 519  | 
|
520  | 
||
521  | 
(* 'extend_classes' inserts all new classes into the corresponding  | 
|
| 2964 | 522  | 
   lists ('classes', 'classrel') if possible *)
 | 
| 422 | 523  | 
|
| 2964 | 524  | 
fun extend_classes (classes, classrel, new_classes) =  | 
| 621 | 525  | 
let  | 
526  | 
val classes' = add_classes classes (map fst new_classes);  | 
|
| 2964 | 527  | 
val classrel' = foldl (add_classrel classes') (classrel, new_classes);  | 
528  | 
in (classes', classrel') end;  | 
|
| 422 | 529  | 
|
530  | 
||
| 621 | 531  | 
(* ext_tsig_classes *)  | 
532  | 
||
533  | 
fun ext_tsig_classes tsig new_classes =  | 
|
534  | 
let  | 
|
| 7641 | 535  | 
    val TySg {classes, classrel, default, tycons, log_types, univ_witness, abbrs, arities} = tsig;
 | 
| 11022 | 536  | 
val (classes', classrel') = extend_classes (classes,classrel, new_classes);  | 
| 12222 | 537  | 
in  | 
538  | 
make_tsig (classes', classrel', default, tycons, log_types, univ_witness, abbrs, arities)  | 
|
539  | 
|> rebuild_tsig  | 
|
540  | 
end;  | 
|
| 621 | 541  | 
|
542  | 
||
| 2964 | 543  | 
(* ext_tsig_classrel *)  | 
| 422 | 544  | 
|
| 2964 | 545  | 
fun ext_tsig_classrel tsig pairs =  | 
| 422 | 546  | 
let  | 
| 7641 | 547  | 
    val TySg {classes, classrel, default, tycons, log_types, univ_witness, abbrs, arities} = tsig;
 | 
| 11022 | 548  | 
val cert = cert_class tsig;  | 
| 422 | 549  | 
|
550  | 
(* FIXME clean! *)  | 
|
| 2964 | 551  | 
val classrel' =  | 
| 11022 | 552  | 
merge_classrel (classrel, Symtab.make (map (fn (c1, c2) => (cert c1, [cert c2])) pairs));  | 
| 422 | 553  | 
in  | 
| 7641 | 554  | 
make_tsig (classes, classrel', default, tycons, log_types, univ_witness, abbrs, arities)  | 
555  | 
|> rebuild_tsig  | 
|
| 422 | 556  | 
end;  | 
557  | 
||
558  | 
||
559  | 
(* ext_tsig_defsort *)  | 
|
560  | 
||
| 7641 | 561  | 
fun ext_tsig_defsort  | 
562  | 
    (TySg {classes, classrel, default = _, tycons, log_types, univ_witness, abbrs, arities, ...}) default =
 | 
|
563  | 
make_tsig (classes, classrel, default, tycons, log_types, univ_witness, abbrs, arities);  | 
|
| 422 | 564  | 
|
565  | 
||
566  | 
||
| 621 | 567  | 
(** add types **)  | 
| 582 | 568  | 
|
| 7641 | 569  | 
fun ext_tsig_types (TySg {classes, classrel, default, tycons, log_types, univ_witness, abbrs, arities}) ts =
 | 
| 582 | 570  | 
let  | 
571  | 
fun check_type (c, n) =  | 
|
572  | 
if n < 0 then err_neg_args c  | 
|
| 7069 | 573  | 
else if is_some (Symtab.lookup (tycons, c)) then err_dup_tycon c  | 
574  | 
else if is_some (Symtab.lookup (abbrs, c)) then error (ty_confl c)  | 
|
| 582 | 575  | 
else ();  | 
| 7641 | 576  | 
val _ = seq check_type ts;  | 
577  | 
val tycons' = Symtab.extend (tycons, ts);  | 
|
578  | 
val arities' = Symtab.extend (arities, map (rpair [] o #1) ts);  | 
|
579  | 
in make_tsig (classes, classrel, default, tycons', log_types, univ_witness, abbrs, arities') end;  | 
|
| 582 | 580  | 
|
581  | 
||
582  | 
||
583  | 
(** add type abbreviations **)  | 
|
584  | 
||
585  | 
fun abbr_errors tsig (a, (lhs_vs, rhs)) =  | 
|
586  | 
let  | 
|
| 963 | 587  | 
    val TySg {tycons, abbrs, ...} = tsig;
 | 
| 621 | 588  | 
val rhs_vs = map (#1 o #1) (typ_tvars rhs);  | 
| 582 | 589  | 
|
590  | 
val dup_lhs_vars =  | 
|
591  | 
(case duplicates lhs_vs of  | 
|
592  | 
[] => []  | 
|
| 621 | 593  | 
| vs => ["Duplicate variables on lhs: " ^ commas_quote vs]);  | 
| 582 | 594  | 
|
595  | 
val extra_rhs_vars =  | 
|
596  | 
(case gen_rems (op =) (rhs_vs, lhs_vs) of  | 
|
597  | 
[] => []  | 
|
| 621 | 598  | 
| vs => ["Extra variables on rhs: " ^ commas_quote vs]);  | 
| 582 | 599  | 
|
600  | 
val tycon_confl =  | 
|
| 7069 | 601  | 
if is_none (Symtab.lookup (tycons, a)) then []  | 
| 582 | 602  | 
else [ty_confl a];  | 
603  | 
||
604  | 
val dup_abbr =  | 
|
| 7069 | 605  | 
if is_none (Symtab.lookup (abbrs, a)) then []  | 
| 582 | 606  | 
else ["Duplicate declaration of abbreviation"];  | 
607  | 
in  | 
|
608  | 
dup_lhs_vars @ extra_rhs_vars @ tycon_confl @ dup_abbr @  | 
|
609  | 
typ_errors tsig (rhs, [])  | 
|
610  | 
end;  | 
|
611  | 
||
| 621 | 612  | 
fun prep_abbr tsig (a, vs, raw_rhs) =  | 
613  | 
let  | 
|
| 4974 | 614  | 
fun err msgs = (seq error_msg msgs;  | 
| 621 | 615  | 
      error ("The error(s) above occurred in type abbreviation " ^ quote a));
 | 
616  | 
||
617  | 
val rhs = rem_sorts (varifyT (no_tvars raw_rhs))  | 
|
618  | 
handle TYPE (msg, _, _) => err [msg];  | 
|
619  | 
val abbr = (a, (vs, rhs));  | 
|
620  | 
in  | 
|
| 582 | 621  | 
(case abbr_errors tsig abbr of  | 
| 621 | 622  | 
[] => abbr  | 
623  | 
| msgs => err msgs)  | 
|
| 582 | 624  | 
end;  | 
625  | 
||
| 7641 | 626  | 
fun add_abbr  | 
627  | 
    (tsig as TySg {classes, classrel, default, tycons, log_types, univ_witness, arities, abbrs}, abbr) =
 | 
|
628  | 
make_tsig (classes, classrel, default, tycons, log_types, univ_witness,  | 
|
629  | 
Symtab.update (prep_abbr tsig abbr, abbrs), arities);  | 
|
| 621 | 630  | 
|
631  | 
fun ext_tsig_abbrs tsig raw_abbrs = foldl add_abbr (tsig, raw_abbrs);  | 
|
| 582 | 632  | 
|
633  | 
||
634  | 
||
| 422 | 635  | 
(** add arities **)  | 
636  | 
||
| 0 | 637  | 
(* 'coregular' checks  | 
| 963 | 638  | 
- the two restrictions 'is_unique_decl' and 'coreg'  | 
| 256 | 639  | 
- if the classes in the new type declarations are known in the  | 
| 0 | 640  | 
given type signature  | 
641  | 
- if one type constructor has always the same number of arguments;  | 
|
| 256 | 642  | 
if one type declaration has passed all checks it is inserted into  | 
| 963 | 643  | 
the 'arities' association list of the given type signatrure *)  | 
| 0 | 644  | 
|
| 2964 | 645  | 
fun coregular (classes, classrel, tycons) =  | 
| 8899 | 646  | 
let fun ex C = if C mem_string classes then () else err_undeclared_class(C);  | 
| 0 | 647  | 
|
| 7069 | 648  | 
fun addar(arities, (t, (w, C))) = case Symtab.lookup (tycons, t) of  | 
| 0 | 649  | 
Some(n) => if n <> length w then varying_decls(t) else  | 
| 963 | 650  | 
((seq o seq) ex w; ex C;  | 
| 7069 | 651  | 
let val ars = the (Symtab.lookup (arities, t))  | 
| 2964 | 652  | 
val ars' = add_arity classrel ars (t,(C,w))  | 
| 7069 | 653  | 
in Symtab.update ((t,ars'), arities) end)  | 
| 8899 | 654  | 
| None => error (undeclared_type t);  | 
| 0 | 655  | 
|
| 963 | 656  | 
in addar end;  | 
| 0 | 657  | 
|
658  | 
||
| 963 | 659  | 
(* 'close' extends the 'arities' association list after all new type  | 
| 0 | 660  | 
declarations have been inserted successfully:  | 
661  | 
for every declaration t:(Ss)C , for all classses D with C <= D:  | 
|
662  | 
if there is no declaration t:(Ss')C' with C < C' and C' <= D  | 
|
| 963 | 663  | 
then insert the declaration t:(Ss)D into 'arities'  | 
| 0 | 664  | 
this means, if there exists a declaration t:(Ss)C and there is  | 
665  | 
no declaration t:(Ss')D with C <=D then the declaration holds  | 
|
| 256 | 666  | 
for all range classes more general than C *)  | 
667  | 
||
| 2964 | 668  | 
fun close classrel arities =  | 
| 7069 | 669  | 
let fun check sl (l, (s, dom)) = case Symtab.lookup (classrel, s) of  | 
| 621 | 670  | 
Some sups =>  | 
| 256 | 671  | 
let fun close_sup (l, sup) =  | 
| 2964 | 672  | 
if exists (fn s'' => less classrel (s, s'') andalso  | 
673  | 
leq classrel (s'', sup)) sl  | 
|
| 0 | 674  | 
then l  | 
| 256 | 675  | 
else (sup, dom)::l  | 
676  | 
in foldl close_sup (l, sups) end  | 
|
| 0 | 677  | 
| None => l;  | 
| 256 | 678  | 
fun ext (s, l) = (s, foldl (check (map #1 l)) (l, l));  | 
| 963 | 679  | 
in map ext arities end;  | 
| 0 | 680  | 
|
| 422 | 681  | 
|
| 621 | 682  | 
(* ext_tsig_arities *)  | 
| 256 | 683  | 
|
| 2964 | 684  | 
fun norm_domain classrel =  | 
685  | 
let fun one_min (f, (doms, ran)) = (f, (map (Sorts.norm_sort classrel) doms, ran))  | 
|
686  | 
in map one_min end;  | 
|
687  | 
||
| 621 | 688  | 
fun ext_tsig_arities tsig sarities =  | 
| 416 | 689  | 
let  | 
| 7641 | 690  | 
    val TySg {classes, classrel, default, tycons, log_types, univ_witness, arities, abbrs} = tsig;
 | 
| 963 | 691  | 
val arities1 =  | 
| 7641 | 692  | 
flat (map (fn (t, ss, cs) => map (fn c => (t, (ss, c))) cs) sarities);  | 
693  | 
val arities2 =  | 
|
694  | 
foldl (coregular (classes, classrel, tycons)) (arities, norm_domain classrel arities1)  | 
|
| 7069 | 695  | 
|> Symtab.dest |> close classrel |> Symtab.make;  | 
| 416 | 696  | 
in  | 
| 7641 | 697  | 
make_tsig (classes, classrel, default, tycons, log_types, univ_witness, abbrs, arities2)  | 
698  | 
|> rebuild_tsig  | 
|
| 416 | 699  | 
end;  | 
| 0 | 700  | 
|
701  | 
||
| 416 | 702  | 
|
| 2964 | 703  | 
(*** type unification and friends ***)  | 
| 0 | 704  | 
|
| 2964 | 705  | 
(** matching **)  | 
| 0 | 706  | 
|
| 2964 | 707  | 
exception TYPE_MATCH;  | 
| 0 | 708  | 
|
| 2964 | 709  | 
fun typ_match tsig =  | 
710  | 
let  | 
|
711  | 
fun match (subs, (TVar (v, S), T)) =  | 
|
| 
8406
 
a217b0cd304d
Type.unify and Type.typ_match now use Vartab instead of association lists.
 
berghofe 
parents: 
7641 
diff
changeset
 | 
712  | 
(case Vartab.lookup (subs, v) of  | 
| 
 
a217b0cd304d
Type.unify and Type.typ_match now use Vartab instead of association lists.
 
berghofe 
parents: 
7641 
diff
changeset
 | 
713  | 
None => (Vartab.update_new ((v, (check_has_sort (tsig, T, S); T)), subs)  | 
| 2964 | 714  | 
handle TYPE _ => raise TYPE_MATCH)  | 
715  | 
| Some U => if U = T then subs else raise TYPE_MATCH)  | 
|
716  | 
| match (subs, (Type (a, Ts), Type (b, Us))) =  | 
|
717  | 
if a <> b then raise TYPE_MATCH  | 
|
718  | 
else foldl match (subs, Ts ~~ Us)  | 
|
719  | 
| match (subs, (TFree x, TFree y)) =  | 
|
720  | 
if x = y then subs else raise TYPE_MATCH  | 
|
721  | 
| match _ = raise TYPE_MATCH;  | 
|
722  | 
in match end;  | 
|
| 0 | 723  | 
|
| 2964 | 724  | 
fun typ_instance (tsig, T, U) =  | 
| 
8406
 
a217b0cd304d
Type.unify and Type.typ_match now use Vartab instead of association lists.
 
berghofe 
parents: 
7641 
diff
changeset
 | 
725  | 
(typ_match tsig (Vartab.empty, (U, T)); true) handle TYPE_MATCH => false;  | 
| 2964 | 726  | 
|
| 0 | 727  | 
|
| 2964 | 728  | 
|
729  | 
(** unification **)  | 
|
730  | 
||
| 0 | 731  | 
exception TUNIFY;  | 
732  | 
||
733  | 
||
| 2964 | 734  | 
(* occurs check *)  | 
| 0 | 735  | 
|
| 2964 | 736  | 
fun occurs v tye =  | 
737  | 
let  | 
|
738  | 
fun occ (Type (_, Ts)) = exists occ Ts  | 
|
739  | 
| occ (TFree _) = false  | 
|
740  | 
| occ (TVar (w, _)) =  | 
|
741  | 
eq_ix (v, w) orelse  | 
|
| 
8406
 
a217b0cd304d
Type.unify and Type.typ_match now use Vartab instead of association lists.
 
berghofe 
parents: 
7641 
diff
changeset
 | 
742  | 
(case Vartab.lookup (tye, w) of  | 
| 2964 | 743  | 
None => false  | 
744  | 
| Some U => occ U);  | 
|
| 0 | 745  | 
in occ end;  | 
746  | 
||
| 2964 | 747  | 
|
748  | 
(* chase variable assignments *)  | 
|
749  | 
||
750  | 
(*if devar returns a type var then it must be unassigned*)  | 
|
751  | 
fun devar (T as TVar (v, _), tye) =  | 
|
| 
8406
 
a217b0cd304d
Type.unify and Type.typ_match now use Vartab instead of association lists.
 
berghofe 
parents: 
7641 
diff
changeset
 | 
752  | 
(case Vartab.lookup (tye, v) of  | 
| 2964 | 753  | 
Some U => devar (U, tye)  | 
754  | 
| None => T)  | 
|
| 256 | 755  | 
| devar (T, tye) = T;  | 
| 0 | 756  | 
|
| 1627 | 757  | 
|
| 2964 | 758  | 
(* unify *)  | 
759  | 
||
| 12528 | 760  | 
fun unify (tsig as TySg {classrel, arities, ...}) (tyenv, maxidx) TU =
 | 
| 2964 | 761  | 
let  | 
762  | 
val tyvar_count = ref maxidx;  | 
|
763  | 
    fun gen_tyvar S = TVar (("'a", inc tyvar_count), S);
 | 
|
764  | 
||
765  | 
fun mg_domain a S =  | 
|
| 7641 | 766  | 
Sorts.mg_domain (classrel, arities) a S handle Sorts.DOMAIN _ => raise TUNIFY;  | 
| 2964 | 767  | 
|
768  | 
fun meet ((_, []), tye) = tye  | 
|
769  | 
| meet ((TVar (xi, S'), S), tye) =  | 
|
770  | 
if Sorts.sort_le classrel (S', S) then tye  | 
|
| 
13666
 
a2730043029b
Removed add_env because Vartab.map was too slow for large environments.
 
berghofe 
parents: 
12726 
diff
changeset
 | 
771  | 
else Vartab.update_new ((xi,  | 
| 
 
a2730043029b
Removed add_env because Vartab.map was too slow for large environments.
 
berghofe 
parents: 
12726 
diff
changeset
 | 
772  | 
gen_tyvar (Sorts.inter_sort classrel (S', S))), tye)  | 
| 2964 | 773  | 
| meet ((TFree (_, S'), S), tye) =  | 
774  | 
if Sorts.sort_le classrel (S', S) then tye  | 
|
775  | 
else raise TUNIFY  | 
|
776  | 
| meet ((Type (a, Ts), S), tye) = meets ((Ts, mg_domain a S), tye)  | 
|
777  | 
and meets (([], []), tye) = tye  | 
|
778  | 
| meets ((T :: Ts, S :: Ss), tye) =  | 
|
779  | 
meets ((Ts, Ss), meet ((devar (T, tye), S), tye))  | 
|
780  | 
| meets _ = sys_error "meets";  | 
|
781  | 
||
782  | 
fun unif ((ty1, ty2), tye) =  | 
|
783  | 
(case (devar (ty1, tye), devar (ty2, tye)) of  | 
|
784  | 
(T as TVar (v, S1), U as TVar (w, S2)) =>  | 
|
785  | 
if eq_ix (v, w) then tye  | 
|
| 
13666
 
a2730043029b
Removed add_env because Vartab.map was too slow for large environments.
 
berghofe 
parents: 
12726 
diff
changeset
 | 
786  | 
else if Sorts.sort_le classrel (S1, S2) then  | 
| 
 
a2730043029b
Removed add_env because Vartab.map was too slow for large environments.
 
berghofe 
parents: 
12726 
diff
changeset
 | 
787  | 
Vartab.update_new ((w, T), tye)  | 
| 
 
a2730043029b
Removed add_env because Vartab.map was too slow for large environments.
 
berghofe 
parents: 
12726 
diff
changeset
 | 
788  | 
else if Sorts.sort_le classrel (S2, S1) then  | 
| 
 
a2730043029b
Removed add_env because Vartab.map was too slow for large environments.
 
berghofe 
parents: 
12726 
diff
changeset
 | 
789  | 
Vartab.update_new ((v, U), tye)  | 
| 2964 | 790  | 
else  | 
791  | 
let val S = gen_tyvar (Sorts.inter_sort classrel (S1, S2)) in  | 
|
| 
13666
 
a2730043029b
Removed add_env because Vartab.map was too slow for large environments.
 
berghofe 
parents: 
12726 
diff
changeset
 | 
792  | 
Vartab.update_new ((v, S), Vartab.update_new ((w, S), tye))  | 
| 2964 | 793  | 
end  | 
794  | 
| (TVar (v, S), T) =>  | 
|
795  | 
if occurs v tye T then raise TUNIFY  | 
|
| 
13666
 
a2730043029b
Removed add_env because Vartab.map was too slow for large environments.
 
berghofe 
parents: 
12726 
diff
changeset
 | 
796  | 
else meet ((T, S), Vartab.update_new ((v, T), tye))  | 
| 2964 | 797  | 
| (T, TVar (v, S)) =>  | 
798  | 
if occurs v tye T then raise TUNIFY  | 
|
| 
13666
 
a2730043029b
Removed add_env because Vartab.map was too slow for large environments.
 
berghofe 
parents: 
12726 
diff
changeset
 | 
799  | 
else meet ((T, S), Vartab.update_new ((v, T), tye))  | 
| 2964 | 800  | 
| (Type (a, Ts), Type (b, Us)) =>  | 
801  | 
if a <> b then raise TUNIFY  | 
|
802  | 
else foldr unif (Ts ~~ Us, tye)  | 
|
803  | 
| (T, U) => if T = U then tye else raise TUNIFY);  | 
|
| 12528 | 804  | 
in (unif (TU, tyenv), ! tyvar_count) end;  | 
| 0 | 805  | 
|
806  | 
||
| 2964 | 807  | 
(* raw_unify *)  | 
| 0 | 808  | 
|
| 2964 | 809  | 
(*purely structural unification -- ignores sorts*)  | 
| 450 | 810  | 
fun raw_unify (ty1, ty2) =  | 
| 12528 | 811  | 
(unify tsig0 (Vartab.empty, 0) (rem_sorts ty1, rem_sorts ty2); true)  | 
| 450 | 812  | 
handle TUNIFY => false;  | 
813  | 
||
814  | 
||
| 0 | 815  | 
|
| 2964 | 816  | 
(** type inference **)  | 
| 
1435
 
aefcd255ed4a
Removed bug in type unification. Negative indexes are not used any longer.
 
nipkow 
parents: 
1392 
diff
changeset
 | 
817  | 
|
| 3790 | 818  | 
(* sort constraints *)  | 
819  | 
||
820  | 
fun get_sort tsig def_sort map_sort raw_env =  | 
|
821  | 
let  | 
|
822  | 
fun eq ((xi, S), (xi', S')) =  | 
|
823  | 
xi = xi' andalso eq_sort tsig (S, S');  | 
|
824  | 
||
825  | 
val env = gen_distinct eq (map (apsnd map_sort) raw_env);  | 
|
826  | 
val _ =  | 
|
827  | 
(case gen_duplicates eq_fst env of  | 
|
828  | 
[] => ()  | 
|
829  | 
      | dups => error ("Inconsistent sort constraints for type variable(s) " ^
 | 
|
830  | 
commas (map (quote o Syntax.string_of_vname' o fst) dups)));  | 
|
| 2587 | 831  | 
|
| 3790 | 832  | 
fun get xi =  | 
833  | 
(case (assoc (env, xi), def_sort xi) of  | 
|
834  | 
(None, None) => defaultS tsig  | 
|
835  | 
| (None, Some S) => S  | 
|
836  | 
| (Some S, None) => S  | 
|
837  | 
| (Some S, Some S') =>  | 
|
| 3804 | 838  | 
if eq_sort tsig (S, S') then S'  | 
| 3790 | 839  | 
          else error ("Sort constraint inconsistent with default for type variable " ^
 | 
840  | 
quote (Syntax.string_of_vname' xi)));  | 
|
841  | 
in get end;  | 
|
842  | 
||
843  | 
||
844  | 
(* type constraints *)  | 
|
| 2587 | 845  | 
|
| 2964 | 846  | 
fun constrain t T =  | 
847  | 
if T = dummyT then t  | 
|
848  | 
  else Const ("_type_constraint_", T) $ t;
 | 
|
| 0 | 849  | 
|
850  | 
||
| 4603 | 851  | 
(* user parameters *)  | 
852  | 
||
853  | 
fun is_param (x, _) = size x > 0 andalso ord x = ord "?";  | 
|
| 12528 | 854  | 
fun param i (x, S) = TVar (("?" ^ x, i), S);
 | 
855  | 
||
| 12726 | 856  | 
fun paramify_dummies (maxidx, TFree ("'_dummy_", S)) =
 | 
857  | 
      (maxidx + 1, param (maxidx + 1) ("'dummy", S))
 | 
|
858  | 
| paramify_dummies (maxidx, Type (a, Ts)) =  | 
|
859  | 
let val (maxidx', Ts') = foldl_map paramify_dummies (maxidx, Ts)  | 
|
860  | 
in (maxidx', Type (a, Ts')) end  | 
|
| 12528 | 861  | 
| paramify_dummies arg = arg;  | 
| 4603 | 862  | 
|
863  | 
||
| 2964 | 864  | 
(* decode_types *)  | 
| 0 | 865  | 
|
| 3804 | 866  | 
(*transform parse tree into raw term*)  | 
| 3790 | 867  | 
fun decode_types tsig is_const def_type def_sort map_const map_type map_sort tm =  | 
| 2964 | 868  | 
let  | 
869  | 
fun get_type xi = if_none (def_type xi) dummyT;  | 
|
| 
5080
 
ce093ff0880e
defaults for free variables hide consts of same name;
 
wenzelm 
parents: 
4974 
diff
changeset
 | 
870  | 
fun is_free x = is_some (def_type (x, ~1));  | 
| 3790 | 871  | 
val raw_env = Syntax.raw_term_sorts tm;  | 
872  | 
val sort_of = get_sort tsig def_sort map_sort raw_env;  | 
|
| 2964 | 873  | 
|
| 3804 | 874  | 
val certT = cert_typ tsig o map_type;  | 
| 12314 | 875  | 
fun decodeT t = certT (Syntax.typ_of_term sort_of map_sort t);  | 
| 0 | 876  | 
|
| 2964 | 877  | 
    fun decode (Const ("_constrain", _) $ t $ typ) =
 | 
878  | 
constrain (decode t) (decodeT typ)  | 
|
| 3804 | 879  | 
      | decode (Const ("_constrainAbs", _) $ (Abs (x, T, t)) $ typ) =
 | 
| 2964 | 880  | 
if T = dummyT then Abs (x, decodeT typ, decode t)  | 
| 3804 | 881  | 
else constrain (Abs (x, certT T, decode t)) (decodeT typ --> dummyT)  | 
882  | 
| decode (Abs (x, T, t)) = Abs (x, certT T, decode t)  | 
|
| 2964 | 883  | 
| decode (t $ u) = decode t $ decode u  | 
| 3804 | 884  | 
| decode (Free (x, T)) =  | 
| 3790 | 885  | 
let val c = map_const x in  | 
| 8721 | 886  | 
if not (is_free x) andalso (is_const c orelse NameSpace.is_qualified c) then  | 
| 
5080
 
ce093ff0880e
defaults for free variables hide consts of same name;
 
wenzelm 
parents: 
4974 
diff
changeset
 | 
887  | 
Const (c, certT T)  | 
| 3790 | 888  | 
else if T = dummyT then Free (x, get_type (x, ~1))  | 
| 3804 | 889  | 
else constrain (Free (x, certT T)) (get_type (x, ~1))  | 
| 3790 | 890  | 
end  | 
| 3804 | 891  | 
| decode (Var (xi, T)) =  | 
| 2964 | 892  | 
if T = dummyT then Var (xi, get_type xi)  | 
| 3804 | 893  | 
else constrain (Var (xi, certT T)) (get_type xi)  | 
| 2964 | 894  | 
| decode (t as Bound _) = t  | 
| 3804 | 895  | 
| decode (Const (c, T)) = Const (map_const c, certT T);  | 
| 8610 | 896  | 
in decode tm end;  | 
| 0 | 897  | 
|
898  | 
||
| 2964 | 899  | 
(* infer_types *)  | 
| 
949
 
83c588d6fee9
Changed treatment of during type inference internally generated type
 
nipkow 
parents: 
621 
diff
changeset
 | 
900  | 
|
| 
 
83c588d6fee9
Changed treatment of during type inference internally generated type
 
nipkow 
parents: 
621 
diff
changeset
 | 
901  | 
(*  | 
| 2964 | 902  | 
Given [T1,...,Tn] and [t1,...,tn], ensure that the type of ti  | 
903  | 
unifies with Ti (for i=1,...,n).  | 
|
904  | 
||
905  | 
tsig: type signature  | 
|
| 3790 | 906  | 
const_type: name mapping and signature lookup  | 
| 2964 | 907  | 
def_type: partial map from indexnames to types (constrains Frees, Vars)  | 
908  | 
def_sort: partial map from indexnames to sorts (constrains TFrees, TVars)  | 
|
909  | 
used: list of already used type variables  | 
|
910  | 
freeze: if true then generated parameters are turned into TFrees, else TVars  | 
|
| 
949
 
83c588d6fee9
Changed treatment of during type inference internally generated type
 
nipkow 
parents: 
621 
diff
changeset
 | 
911  | 
*)  | 
| 0 | 912  | 
|
| 3790 | 913  | 
fun infer_types prt prT tsig const_type def_type def_sort  | 
914  | 
map_const map_type map_sort used freeze pat_Ts raw_ts =  | 
|
| 565 | 915  | 
let  | 
| 2964 | 916  | 
    val TySg {classrel, arities, ...} = tsig;
 | 
917  | 
val pat_Ts' = map (cert_typ tsig) pat_Ts;  | 
|
| 3790 | 918  | 
val is_const = is_some o const_type;  | 
| 2964 | 919  | 
val raw_ts' =  | 
| 3790 | 920  | 
map (decode_types tsig is_const def_type def_sort map_const map_type map_sort) raw_ts;  | 
| 2964 | 921  | 
val (ts, Ts, unifier) =  | 
| 2979 | 922  | 
TypeInfer.infer_types prt prT const_type classrel arities used freeze  | 
| 4603 | 923  | 
is_param raw_ts' pat_Ts';  | 
| 7641 | 924  | 
in (ts, unifier) end;  | 
| 0 | 925  | 
|
| 3790 | 926  | 
|
| 0 | 927  | 
end;  |