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