| author | bulwahn | 
| Fri, 08 Apr 2011 16:31:14 +0200 | |
| changeset 42316 | 12635bb655fd | 
| parent 41353 | 684003dbda54 | 
| child 42284 | 326f57825e1a | 
| permissions | -rw-r--r-- | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
1  | 
(* Title: Tools/subtyping.ML  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
2  | 
Author: Dmitriy Traytel, TU Muenchen  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
3  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
4  | 
Coercive subtyping via subtype constraints.  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
5  | 
*)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
6  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
7  | 
signature SUBTYPING =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
8  | 
sig  | 
| 
41353
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
9  | 
datatype variance = COVARIANT | CONTRAVARIANT | INVARIANT | INVARIANT_TO of typ;  | 
| 
40939
 
2c150063cd4d
setup subtyping/coercions once in HOL.thy, but enable it only later via configuration option;
 
wenzelm 
parents: 
40938 
diff
changeset
 | 
10  | 
val coercion_enabled: bool Config.T  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
11  | 
val infer_types: Proof.context -> (string -> typ option) -> (indexname -> typ option) ->  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
12  | 
term list -> term list  | 
| 
40284
 
c9acf88447e6
export declarations by default, to allow other ML packages by-pass concrete syntax;
 
wenzelm 
parents: 
40283 
diff
changeset
 | 
13  | 
val add_type_map: term -> Context.generic -> Context.generic  | 
| 
 
c9acf88447e6
export declarations by default, to allow other ML packages by-pass concrete syntax;
 
wenzelm 
parents: 
40283 
diff
changeset
 | 
14  | 
val add_coercion: term -> Context.generic -> Context.generic  | 
| 40836 | 15  | 
val gen_coercion: Proof.context -> typ Vartab.table -> (typ * typ) -> term  | 
| 40283 | 16  | 
val setup: theory -> theory  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
17  | 
end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
18  | 
|
| 40283 | 19  | 
structure Subtyping: SUBTYPING =  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
20  | 
struct  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
21  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
22  | 
(** coercions data **)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
23  | 
|
| 
41353
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
24  | 
datatype variance = COVARIANT | CONTRAVARIANT | INVARIANT | INVARIANT_TO of typ;  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
25  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
26  | 
datatype data = Data of  | 
| 40282 | 27  | 
  {coes: term Symreltab.table,  (*coercions table*)
 | 
28  | 
coes_graph: unit Graph.T, (*coercions graph*)  | 
|
29  | 
tmaps: (term * variance list) Symtab.table}; (*map functions*)  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
30  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
31  | 
fun make_data (coes, coes_graph, tmaps) =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
32  | 
  Data {coes = coes, coes_graph = coes_graph, tmaps = tmaps};
 | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
33  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
34  | 
structure Data = Generic_Data  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
35  | 
(  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
36  | 
type T = data;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
37  | 
val empty = make_data (Symreltab.empty, Graph.empty, Symtab.empty);  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
38  | 
val extend = I;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
39  | 
fun merge  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
40  | 
    (Data {coes = coes1, coes_graph = coes_graph1, tmaps = tmaps1},
 | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
41  | 
      Data {coes = coes2, coes_graph = coes_graph2, tmaps = tmaps2}) =
 | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
42  | 
make_data (Symreltab.merge (op aconv) (coes1, coes2),  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
43  | 
Graph.merge (op =) (coes_graph1, coes_graph2),  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
44  | 
Symtab.merge (eq_pair (op aconv) (op =)) (tmaps1, tmaps2));  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
45  | 
);  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
46  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
47  | 
fun map_data f =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
48  | 
  Data.map (fn Data {coes, coes_graph, tmaps} =>
 | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
49  | 
make_data (f (coes, coes_graph, tmaps)));  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
50  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
51  | 
fun map_coes f =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
52  | 
map_data (fn (coes, coes_graph, tmaps) =>  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
53  | 
(f coes, coes_graph, tmaps));  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
54  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
55  | 
fun map_coes_graph f =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
56  | 
map_data (fn (coes, coes_graph, tmaps) =>  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
57  | 
(coes, f coes_graph, tmaps));  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
58  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
59  | 
fun map_coes_and_graph f =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
60  | 
map_data (fn (coes, coes_graph, tmaps) =>  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
61  | 
let val (coes', coes_graph') = f (coes, coes_graph);  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
62  | 
in (coes', coes_graph', tmaps) end);  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
63  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
64  | 
fun map_tmaps f =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
65  | 
map_data (fn (coes, coes_graph, tmaps) =>  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
66  | 
(coes, coes_graph, f tmaps));  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
67  | 
|
| 40285 | 68  | 
val rep_data = (fn Data args => args) o Data.get o Context.Proof;  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
69  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
70  | 
val coes_of = #coes o rep_data;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
71  | 
val coes_graph_of = #coes_graph o rep_data;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
72  | 
val tmaps_of = #tmaps o rep_data;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
73  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
74  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
75  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
76  | 
(** utils **)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
77  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
78  | 
fun nameT (Type (s, [])) = s;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
79  | 
fun t_of s = Type (s, []);  | 
| 40286 | 80  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
81  | 
fun sort_of (TFree (_, S)) = SOME S  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
82  | 
| sort_of (TVar (_, S)) = SOME S  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
83  | 
| sort_of _ = NONE;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
84  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
85  | 
val is_typeT = fn (Type _) => true | _ => false;  | 
| 
41353
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
86  | 
val is_stypeT = fn (Type (_, [])) => true | _ => false;  | 
| 40282 | 87  | 
val is_compT = fn (Type (_, _ :: _)) => true | _ => false;  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
88  | 
val is_freeT = fn (TFree _) => true | _ => false;  | 
| 40286 | 89  | 
val is_fixedvarT = fn (TVar (xi, _)) => not (Type_Infer.is_param xi) | _ => false;  | 
| 
41353
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
90  | 
val is_funtype = fn (Type ("fun", [_, _])) => true | _ => false;
 | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
91  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
92  | 
|
| 40836 | 93  | 
(* unification *)  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
94  | 
|
| 40836 | 95  | 
exception TYPE_INFERENCE_ERROR of unit -> string;  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
96  | 
exception NO_UNIFIER of string * typ Vartab.table;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
97  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
98  | 
fun unify weak ctxt =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
99  | 
let  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
100  | 
val thy = ProofContext.theory_of ctxt;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
101  | 
val pp = Syntax.pp ctxt;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
102  | 
val arity_sorts = Type.arity_sorts pp (Sign.tsig_of thy);  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
103  | 
|
| 40282 | 104  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
105  | 
(* adjust sorts of parameters *)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
106  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
107  | 
fun not_of_sort x S' S =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
108  | 
"Variable " ^ x ^ "::" ^ Syntax.string_of_sort ctxt S' ^ " not of sort " ^  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
109  | 
Syntax.string_of_sort ctxt S;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
110  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
111  | 
fun meet (_, []) tye_idx = tye_idx  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
112  | 
| meet (Type (a, Ts), S) (tye_idx as (tye, _)) =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
113  | 
meets (Ts, arity_sorts a S handle ERROR msg => raise NO_UNIFIER (msg, tye)) tye_idx  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
114  | 
| meet (TFree (x, S'), S) (tye_idx as (tye, _)) =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
115  | 
if Sign.subsort thy (S', S) then tye_idx  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
116  | 
else raise NO_UNIFIER (not_of_sort x S' S, tye)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
117  | 
| meet (TVar (xi, S'), S) (tye_idx as (tye, idx)) =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
118  | 
if Sign.subsort thy (S', S) then tye_idx  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
119  | 
else if Type_Infer.is_param xi then  | 
| 40286 | 120  | 
(Vartab.update_new  | 
121  | 
(xi, Type_Infer.mk_param idx (Sign.inter_sort thy (S', S))) tye, idx + 1)  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
122  | 
else raise NO_UNIFIER (not_of_sort (Term.string_of_vname xi) S' S, tye)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
123  | 
and meets (T :: Ts, S :: Ss) (tye_idx as (tye, _)) =  | 
| 40286 | 124  | 
meets (Ts, Ss) (meet (Type_Infer.deref tye T, S) tye_idx)  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
125  | 
| meets _ tye_idx = tye_idx;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
126  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
127  | 
val weak_meet = if weak then fn _ => I else meet  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
128  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
129  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
130  | 
(* occurs check and assignment *)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
131  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
132  | 
fun occurs_check tye xi (TVar (xi', _)) =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
133  | 
          if xi = xi' then raise NO_UNIFIER ("Occurs check!", tye)
 | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
134  | 
else  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
135  | 
(case Vartab.lookup tye xi' of  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
136  | 
NONE => ()  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
137  | 
| SOME T => occurs_check tye xi T)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
138  | 
| occurs_check tye xi (Type (_, Ts)) = List.app (occurs_check tye xi) Ts  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
139  | 
| occurs_check _ _ _ = ();  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
140  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
141  | 
fun assign xi (T as TVar (xi', _)) S env =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
142  | 
if xi = xi' then env  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
143  | 
else env |> weak_meet (T, S) |>> Vartab.update_new (xi, T)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
144  | 
| assign xi T S (env as (tye, _)) =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
145  | 
(occurs_check tye xi T; env |> weak_meet (T, S) |>> Vartab.update_new (xi, T));  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
146  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
147  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
148  | 
(* unification *)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
149  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
150  | 
fun show_tycon (a, Ts) =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
151  | 
quote (Syntax.string_of_typ ctxt (Type (a, replicate (length Ts) dummyT)));  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
152  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
153  | 
fun unif (T1, T2) (env as (tye, _)) =  | 
| 40286 | 154  | 
(case pairself (`Type_Infer.is_paramT o Type_Infer.deref tye) (T1, T2) of  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
155  | 
((true, TVar (xi, S)), (_, T)) => assign xi T S env  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
156  | 
| ((_, T), (true, TVar (xi, S))) => assign xi T S env  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
157  | 
| ((_, Type (a, Ts)), (_, Type (b, Us))) =>  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
158  | 
if weak andalso null Ts andalso null Us then env  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
159  | 
else if a <> b then  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
160  | 
raise NO_UNIFIER  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
161  | 
              ("Clash of types " ^ show_tycon (a, Ts) ^ " and " ^ show_tycon (b, Us), tye)
 | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
162  | 
else fold unif (Ts ~~ Us) env  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
163  | 
      | ((_, T), (_, U)) => if T = U then env else raise NO_UNIFIER ("", tye));
 | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
164  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
165  | 
in unif end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
166  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
167  | 
val weak_unify = unify true;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
168  | 
val strong_unify = unify false;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
169  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
170  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
171  | 
(* Typ_Graph shortcuts *)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
172  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
173  | 
val add_edge = Typ_Graph.add_edge_acyclic;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
174  | 
fun get_preds G T = Typ_Graph.all_preds G [T];  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
175  | 
fun get_succs G T = Typ_Graph.all_succs G [T];  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
176  | 
fun maybe_new_typnode T G = perhaps (try (Typ_Graph.new_node (T, ()))) G;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
177  | 
fun maybe_new_typnodes Ts G = fold maybe_new_typnode Ts G;  | 
| 40282 | 178  | 
fun new_imm_preds G Ts =  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
179  | 
subtract (op =) Ts (distinct (op =) (maps (Typ_Graph.imm_preds G) Ts));  | 
| 40282 | 180  | 
fun new_imm_succs G Ts =  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
181  | 
subtract op= Ts (distinct (op =) (maps (Typ_Graph.imm_succs G) Ts));  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
182  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
183  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
184  | 
(* Graph shortcuts *)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
185  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
186  | 
fun maybe_new_node s G = perhaps (try (Graph.new_node (s, ()))) G  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
187  | 
fun maybe_new_nodes ss G = fold maybe_new_node ss G  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
188  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
189  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
190  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
191  | 
(** error messages **)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
192  | 
|
| 40836 | 193  | 
fun gen_msg err msg =  | 
194  | 
err () ^ "\nNow trying to infer coercions:\n\nCoercion inference failed" ^  | 
|
195  | 
(if msg = "" then "" else ": " ^ msg) ^ "\n";  | 
|
196  | 
||
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
197  | 
fun prep_output ctxt tye bs ts Ts =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
198  | 
let  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
199  | 
val (Ts_bTs', ts') = Type_Infer.finish ctxt tye (Ts @ map snd bs, ts);  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
200  | 
val (Ts', Ts'') = chop (length Ts) Ts_bTs';  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
201  | 
fun prep t =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
202  | 
let val xs = rev (Term.variant_frees t (rev (map fst bs ~~ Ts'')))  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
203  | 
in Term.subst_bounds (map Syntax.mark_boundT xs, t) end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
204  | 
in (map prep ts', Ts') end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
205  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
206  | 
fun err_loose i = error ("Loose bound variable: B." ^ string_of_int i);
 | 
| 40836 | 207  | 
|
208  | 
fun unif_failed msg =  | 
|
209  | 
"Type unification failed" ^ (if msg = "" then "" else ": " ^ msg) ^ "\n\n";  | 
|
210  | 
||
211  | 
fun err_appl_msg ctxt msg tye bs t T u U () =  | 
|
212  | 
let val ([t', u'], [T', U']) = prep_output ctxt tye bs [t, u] [T, U]  | 
|
213  | 
in unif_failed msg ^ Type.appl_error (Syntax.pp ctxt) t' T' u' U' ^ "\n" end;  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
214  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
215  | 
fun err_list ctxt msg tye Ts =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
216  | 
let  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
217  | 
val (_, Ts') = prep_output ctxt tye [] [] Ts;  | 
| 40836 | 218  | 
val text = cat_lines ([msg,  | 
219  | 
"Cannot unify a list of types that should be the same:",  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
220  | 
(Pretty.string_of (Pretty.list "[" "]" (map (Pretty.typ (Syntax.pp ctxt)) Ts')))]);  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
221  | 
in  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
222  | 
error text  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
223  | 
end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
224  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
225  | 
fun err_bound ctxt msg tye packs =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
226  | 
let  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
227  | 
val pp = Syntax.pp ctxt;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
228  | 
val (ts, Ts) = fold  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
229  | 
(fn (bs, t $ u, U, _, U') => fn (ts, Ts) =>  | 
| 40836 | 230  | 
let val (t', T') = prep_output ctxt tye bs [t, u] [U', U]  | 
| 40282 | 231  | 
in (t' :: ts, T' :: Ts) end)  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
232  | 
packs ([], []);  | 
| 40836 | 233  | 
val text = cat_lines ([msg, "Cannot fulfil subtype constraints:"] @  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
234  | 
(map2 (fn [t, u] => fn [T, U] => Pretty.string_of (  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
235  | 
Pretty.block [  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
236  | 
Pretty.typ pp T, Pretty.brk 2, Pretty.str "<:", Pretty.brk 2, Pretty.typ pp U,  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
237  | 
Pretty.brk 3, Pretty.str "from function application", Pretty.brk 2,  | 
| 40836 | 238  | 
Pretty.block [Pretty.term pp (t $ u)]]))  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
239  | 
ts Ts))  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
240  | 
in  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
241  | 
error text  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
242  | 
end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
243  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
244  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
245  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
246  | 
(** constraint generation **)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
247  | 
|
| 40836 | 248  | 
fun generate_constraints ctxt err =  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
249  | 
let  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
250  | 
fun gen cs _ (Const (_, T)) tye_idx = (T, tye_idx, cs)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
251  | 
| gen cs _ (Free (_, T)) tye_idx = (T, tye_idx, cs)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
252  | 
| gen cs _ (Var (_, T)) tye_idx = (T, tye_idx, cs)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
253  | 
| gen cs bs (Bound i) tye_idx =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
254  | 
(snd (nth bs i handle Subscript => err_loose i), tye_idx, cs)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
255  | 
| gen cs bs (Abs (x, T, t)) tye_idx =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
256  | 
let val (U, tye_idx', cs') = gen cs ((x, T) :: bs) t tye_idx  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
257  | 
in (T --> U, tye_idx', cs') end  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
258  | 
| gen cs bs (t $ u) tye_idx =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
259  | 
let  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
260  | 
val (T, tye_idx', cs') = gen cs bs t tye_idx;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
261  | 
val (U', (tye, idx), cs'') = gen cs' bs u tye_idx';  | 
| 40286 | 262  | 
val U = Type_Infer.mk_param idx [];  | 
263  | 
val V = Type_Infer.mk_param (idx + 1) [];  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
264  | 
val tye_idx''= strong_unify ctxt (U --> V, T) (tye, idx + 2)  | 
| 
41353
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
265  | 
handle NO_UNIFIER (msg, _) => error (gen_msg err msg);  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
266  | 
val error_pack = (bs, t $ u, U, V, U');  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
267  | 
in (V, tye_idx'', ((U', U), error_pack) :: cs'') end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
268  | 
in  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
269  | 
gen [] []  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
270  | 
end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
271  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
272  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
273  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
274  | 
(** constraint resolution **)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
275  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
276  | 
exception BOUND_ERROR of string;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
277  | 
|
| 40836 | 278  | 
fun process_constraints ctxt err cs tye_idx =  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
279  | 
let  | 
| 40285 | 280  | 
val coes_graph = coes_graph_of ctxt;  | 
281  | 
val tmaps = tmaps_of ctxt;  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
282  | 
val tsig = Sign.tsig_of (ProofContext.theory_of ctxt);  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
283  | 
val pp = Syntax.pp ctxt;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
284  | 
val arity_sorts = Type.arity_sorts pp tsig;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
285  | 
val subsort = Type.subsort tsig;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
286  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
287  | 
fun split_cs _ [] = ([], [])  | 
| 40282 | 288  | 
| split_cs f (c :: cs) =  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
289  | 
(case pairself f (fst c) of  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
290  | 
(false, false) => apsnd (cons c) (split_cs f cs)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
291  | 
| _ => apfst (cons c) (split_cs f cs));  | 
| 
41353
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
292  | 
|
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
293  | 
fun unify_list (T :: Ts) tye_idx =  | 
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
294  | 
fold (fn U => fn tye_idx' => strong_unify ctxt (T, U) tye_idx') Ts tye_idx;  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
295  | 
|
| 40282 | 296  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
297  | 
(* check whether constraint simplification will terminate using weak unification *)  | 
| 40282 | 298  | 
|
| 
41353
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
299  | 
val _ = fold (fn (TU, _) => fn tye_idx =>  | 
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
300  | 
weak_unify ctxt TU tye_idx handle NO_UNIFIER (msg, _) =>  | 
| 40836 | 301  | 
        error (gen_msg err ("weak unification of subtype constraints fails\n" ^ msg))) cs tye_idx;
 | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
302  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
303  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
304  | 
(* simplify constraints *)  | 
| 40282 | 305  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
306  | 
fun simplify_constraints cs tye_idx =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
307  | 
let  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
308  | 
fun contract a Ts Us error_pack done todo tye idx =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
309  | 
let  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
310  | 
val arg_var =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
311  | 
(case Symtab.lookup tmaps a of  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
312  | 
(*everything is invariant for unknown constructors*)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
313  | 
NONE => replicate (length Ts) INVARIANT  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
314  | 
| SOME av => snd av);  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
315  | 
fun new_constraints (variance, constraint) (cs, tye_idx) =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
316  | 
(case variance of  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
317  | 
COVARIANT => (constraint :: cs, tye_idx)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
318  | 
| CONTRAVARIANT => (swap constraint :: cs, tye_idx)  | 
| 
41353
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
319  | 
| INVARIANT_TO T => (cs, unify_list [T, fst constraint, snd constraint] tye_idx  | 
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
320  | 
handle NO_UNIFIER (msg, _) =>  | 
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
321  | 
err_list ctxt (gen_msg err  | 
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
322  | 
"failed to unify invariant arguments w.r.t. to the known map function")  | 
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
323  | 
(fst tye_idx) Ts)  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
324  | 
| INVARIANT => (cs, strong_unify ctxt constraint tye_idx  | 
| 
41353
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
325  | 
handle NO_UNIFIER (msg, _) =>  | 
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
326  | 
                    error (gen_msg err ("failed to unify invariant arguments" ^ msg))));
 | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
327  | 
val (new, (tye', idx')) = apfst (fn cs => (cs ~~ replicate (length cs) error_pack))  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
328  | 
(fold new_constraints (arg_var ~~ (Ts ~~ Us)) ([], (tye, idx)));  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
329  | 
val test_update = is_compT orf is_freeT orf is_fixedvarT;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
330  | 
val (ch, done') =  | 
| 40286 | 331  | 
if not (null new) then ([], done)  | 
332  | 
else split_cs (test_update o Type_Infer.deref tye') done;  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
333  | 
val todo' = ch @ todo;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
334  | 
in  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
335  | 
simplify done' (new @ todo') (tye', idx')  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
336  | 
end  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
337  | 
(*xi is definitely a parameter*)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
338  | 
and expand varleq xi S a Ts error_pack done todo tye idx =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
339  | 
let  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
340  | 
val n = length Ts;  | 
| 40286 | 341  | 
val args = map2 Type_Infer.mk_param (idx upto idx + n - 1) (arity_sorts a S);  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
342  | 
val tye' = Vartab.update_new (xi, Type(a, args)) tye;  | 
| 40286 | 343  | 
val (ch, done') = split_cs (is_compT o Type_Infer.deref tye') done;  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
344  | 
val todo' = ch @ todo;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
345  | 
val new =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
346  | 
if varleq then (Type(a, args), Type (a, Ts))  | 
| 40286 | 347  | 
else (Type (a, Ts), Type (a, args));  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
348  | 
in  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
349  | 
simplify done' ((new, error_pack) :: todo') (tye', idx + n)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
350  | 
end  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
351  | 
(*TU is a pair of a parameter and a free/fixed variable*)  | 
| 
41353
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
352  | 
and eliminate TU done todo tye idx =  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
353  | 
let  | 
| 40286 | 354  | 
val [TVar (xi, S)] = filter Type_Infer.is_paramT TU;  | 
355  | 
val [T] = filter_out Type_Infer.is_paramT TU;  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
356  | 
val SOME S' = sort_of T;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
357  | 
val test_update = if is_freeT T then is_freeT else is_fixedvarT;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
358  | 
val tye' = Vartab.update_new (xi, T) tye;  | 
| 40286 | 359  | 
val (ch, done') = split_cs (test_update o Type_Infer.deref tye') done;  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
360  | 
val todo' = ch @ todo;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
361  | 
in  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
362  | 
if subsort (S', S) (*TODO check this*)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
363  | 
then simplify done' todo' (tye', idx)  | 
| 40836 | 364  | 
else error (gen_msg err "sort mismatch")  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
365  | 
end  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
366  | 
and simplify done [] tye_idx = (done, tye_idx)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
367  | 
| simplify done (((T, U), error_pack) :: todo) (tye_idx as (tye, idx)) =  | 
| 40286 | 368  | 
(case (Type_Infer.deref tye T, Type_Infer.deref tye U) of  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
369  | 
(Type (a, []), Type (b, [])) =>  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
370  | 
if a = b then simplify done todo tye_idx  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
371  | 
else if Graph.is_edge coes_graph (a, b) then simplify done todo tye_idx  | 
| 40836 | 372  | 
else error (gen_msg err (a ^ " is not a subtype of " ^ b))  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
373  | 
| (Type (a, Ts), Type (b, Us)) =>  | 
| 40836 | 374  | 
if a <> b then error (gen_msg err "different constructors")  | 
375  | 
(fst tye_idx) error_pack  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
376  | 
else contract a Ts Us error_pack done todo tye idx  | 
| 40282 | 377  | 
| (TVar (xi, S), Type (a, Ts as (_ :: _))) =>  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
378  | 
expand true xi S a Ts error_pack done todo tye idx  | 
| 40282 | 379  | 
| (Type (a, Ts as (_ :: _)), TVar (xi, S)) =>  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
380  | 
expand false xi S a Ts error_pack done todo tye idx  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
381  | 
| (T, U) =>  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
382  | 
if T = U then simplify done todo tye_idx  | 
| 40282 | 383  | 
else if exists (is_freeT orf is_fixedvarT) [T, U] andalso  | 
| 40286 | 384  | 
exists Type_Infer.is_paramT [T, U]  | 
| 
41353
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
385  | 
then eliminate [T, U] done todo tye idx  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
386  | 
else if exists (is_freeT orf is_fixedvarT) [T, U]  | 
| 40836 | 387  | 
then error (gen_msg err "not eliminated free/fixed variables")  | 
| 40282 | 388  | 
else simplify (((T, U), error_pack) :: done) todo tye_idx);  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
389  | 
in  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
390  | 
simplify [] cs tye_idx  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
391  | 
end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
392  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
393  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
394  | 
(* do simplification *)  | 
| 40282 | 395  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
396  | 
val (cs', tye_idx') = simplify_constraints cs tye_idx;  | 
| 40836 | 397  | 
|
398  | 
fun find_error_pack lower T' = map_filter  | 
|
399  | 
(fn ((T, U), pack) => if if lower then T' = U else T' = T then SOME pack else NONE) cs';  | 
|
400  | 
||
401  | 
fun find_cycle_packs nodes =  | 
|
402  | 
let  | 
|
403  | 
val (but_last, last) = split_last nodes  | 
|
404  | 
val pairs = (last, hd nodes) :: (but_last ~~ tl nodes);  | 
|
405  | 
in  | 
|
406  | 
map_filter  | 
|
| 40838 | 407  | 
(fn (TU, pack) => if member (op =) pairs TU then SOME pack else NONE)  | 
| 40836 | 408  | 
cs'  | 
409  | 
end;  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
410  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
411  | 
(*styps stands either for supertypes or for subtypes of a type T  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
412  | 
in terms of the subtype-relation (excluding T itself)*)  | 
| 40282 | 413  | 
fun styps super T =  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
414  | 
(if super then Graph.imm_succs else Graph.imm_preds) coes_graph T  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
415  | 
handle Graph.UNDEF _ => [];  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
416  | 
|
| 40282 | 417  | 
fun minmax sup (T :: Ts) =  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
418  | 
let  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
419  | 
fun adjust T U = if sup then (T, U) else (U, T);  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
420  | 
fun extract T [] = T  | 
| 40282 | 421  | 
| extract T (U :: Us) =  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
422  | 
if Graph.is_edge coes_graph (adjust T U) then extract T Us  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
423  | 
else if Graph.is_edge coes_graph (adjust U T) then extract U Us  | 
| 40836 | 424  | 
else raise BOUND_ERROR "uncomparable types in type list";  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
425  | 
in  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
426  | 
t_of (extract T Ts)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
427  | 
end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
428  | 
|
| 40282 | 429  | 
fun ex_styp_of_sort super T styps_and_sorts =  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
430  | 
let  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
431  | 
fun adjust T U = if super then (T, U) else (U, T);  | 
| 40282 | 432  | 
fun styp_test U Ts = forall  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
433  | 
(fn T => T = U orelse Graph.is_edge coes_graph (adjust U T)) Ts;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
434  | 
fun fitting Ts S U = Type.of_sort tsig (t_of U, S) andalso styp_test U Ts  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
435  | 
in  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
436  | 
forall (fn (Ts, S) => exists (fitting Ts S) (T :: styps super T)) styps_and_sorts  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
437  | 
end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
438  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
439  | 
(* computes the tightest possible, correct assignment for 'a::S  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
440  | 
e.g. in the supremum case (sup = true):  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
441  | 
------- 'a::S---  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
442  | 
/ / \ \  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
443  | 
/ / \ \  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
444  | 
'b::C1 'c::C2 ... T1 T2 ...  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
445  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
446  | 
sorts - list of sorts [C1, C2, ...]  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
447  | 
T::Ts - non-empty list of base types [T1, T2, ...]  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
448  | 
*)  | 
| 40282 | 449  | 
fun tightest sup S styps_and_sorts (T :: Ts) =  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
450  | 
let  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
451  | 
fun restriction T = Type.of_sort tsig (t_of T, S)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
452  | 
andalso ex_styp_of_sort (not sup) T styps_and_sorts;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
453  | 
fun candidates T = inter (op =) (filter restriction (T :: styps sup T));  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
454  | 
in  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
455  | 
(case fold candidates Ts (filter restriction (T :: styps sup T)) of  | 
| 40836 | 456  | 
          [] => raise BOUND_ERROR ("no " ^ (if sup then "supremum" else "infimum"))
 | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
457  | 
| [T] => t_of T  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
458  | 
| Ts => minmax sup Ts)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
459  | 
end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
460  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
461  | 
fun build_graph G [] tye_idx = (G, tye_idx)  | 
| 40282 | 462  | 
| build_graph G ((T, U) :: cs) tye_idx =  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
463  | 
if T = U then build_graph G cs tye_idx  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
464  | 
else  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
465  | 
let  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
466  | 
val G' = maybe_new_typnodes [T, U] G;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
467  | 
val (G'', tye_idx') = (add_edge (T, U) G', tye_idx)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
468  | 
handle Typ_Graph.CYCLES cycles =>  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
469  | 
let  | 
| 40836 | 470  | 
val (tye, idx) =  | 
471  | 
fold  | 
|
472  | 
(fn cycle => fn tye_idx' => (unify_list cycle tye_idx'  | 
|
| 
41353
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
473  | 
handle NO_UNIFIER (msg, _) =>  | 
| 40836 | 474  | 
err_bound ctxt  | 
475  | 
                            (gen_msg err ("constraint cycle not unifiable" ^ msg)) (fst tye_idx)
 | 
|
476  | 
(find_cycle_packs cycle)))  | 
|
477  | 
cycles tye_idx  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
478  | 
in  | 
| 40836 | 479  | 
collapse (tye, idx) cycles G  | 
480  | 
end  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
481  | 
in  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
482  | 
build_graph G'' cs tye_idx'  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
483  | 
end  | 
| 40836 | 484  | 
and collapse (tye, idx) cycles G = (*nodes non-empty list*)  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
485  | 
let  | 
| 40836 | 486  | 
(*all cycles collapse to one node,  | 
487  | 
because all of them share at least the nodes x and y*)  | 
|
488  | 
val nodes = (distinct (op =) (flat cycles));  | 
|
489  | 
val T = Type_Infer.deref tye (hd nodes);  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
490  | 
val P = new_imm_preds G nodes;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
491  | 
val S = new_imm_succs G nodes;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
492  | 
val G' = Typ_Graph.del_nodes (tl nodes) G;  | 
| 40836 | 493  | 
fun check_and_gen super T' =  | 
494  | 
let val U = Type_Infer.deref tye T';  | 
|
495  | 
in  | 
|
496  | 
if not (is_typeT T) orelse not (is_typeT U) orelse T = U  | 
|
497  | 
then if super then (hd nodes, T') else (T', hd nodes)  | 
|
498  | 
else  | 
|
499  | 
if super andalso  | 
|
500  | 
Graph.is_edge coes_graph (nameT T, nameT U) then (hd nodes, T')  | 
|
501  | 
else if not super andalso  | 
|
502  | 
Graph.is_edge coes_graph (nameT U, nameT T) then (T', hd nodes)  | 
|
503  | 
else err_bound ctxt (gen_msg err "cycle elimination produces inconsistent graph")  | 
|
504  | 
(fst tye_idx)  | 
|
505  | 
(maps find_cycle_packs cycles @ find_error_pack super T')  | 
|
506  | 
end;  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
507  | 
in  | 
| 40836 | 508  | 
build_graph G' (map (check_and_gen false) P @ map (check_and_gen true) S) (tye, idx)  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
509  | 
end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
510  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
511  | 
fun assign_bound lower G key (tye_idx as (tye, _)) =  | 
| 40286 | 512  | 
if Type_Infer.is_paramT (Type_Infer.deref tye key) then  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
513  | 
let  | 
| 40286 | 514  | 
val TVar (xi, S) = Type_Infer.deref tye key;  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
515  | 
val get_bound = if lower then get_preds else get_succs;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
516  | 
val raw_bound = get_bound G key;  | 
| 40286 | 517  | 
val bound = map (Type_Infer.deref tye) raw_bound;  | 
518  | 
val not_params = filter_out Type_Infer.is_paramT bound;  | 
|
| 40282 | 519  | 
fun to_fulfil T =  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
520  | 
(case sort_of T of  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
521  | 
NONE => NONE  | 
| 40282 | 522  | 
| SOME S =>  | 
| 40286 | 523  | 
SOME  | 
524  | 
(map nameT  | 
|
525  | 
(filter_out Type_Infer.is_paramT (map (Type_Infer.deref tye) (get_bound G T))),  | 
|
526  | 
S));  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
527  | 
val styps_and_sorts = distinct (op =) (map_filter to_fulfil raw_bound);  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
528  | 
val assignment =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
529  | 
if null bound orelse null not_params then NONE  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
530  | 
else SOME (tightest lower S styps_and_sorts (map nameT not_params)  | 
| 40836 | 531  | 
handle BOUND_ERROR msg =>  | 
532  | 
err_bound ctxt (gen_msg err msg) tye (find_error_pack lower key))  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
533  | 
in  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
534  | 
(case assignment of  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
535  | 
NONE => tye_idx  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
536  | 
| SOME T =>  | 
| 40286 | 537  | 
if Type_Infer.is_paramT T then tye_idx  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
538  | 
else if lower then (*upper bound check*)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
539  | 
let  | 
| 40286 | 540  | 
val other_bound = map (Type_Infer.deref tye) (get_succs G key);  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
541  | 
val s = nameT T;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
542  | 
in  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
543  | 
if subset (op = o apfst nameT) (filter is_typeT other_bound, s :: styps true s)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
544  | 
then apfst (Vartab.update (xi, T)) tye_idx  | 
| 40836 | 545  | 
                  else err_bound ctxt (gen_msg err ("assigned simple type " ^ s ^
 | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
546  | 
" clashes with the upper bound of variable " ^  | 
| 40836 | 547  | 
Syntax.string_of_typ ctxt (TVar(xi, S)))) tye (find_error_pack (not lower) key)  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
548  | 
end  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
549  | 
else apfst (Vartab.update (xi, T)) tye_idx)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
550  | 
end  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
551  | 
else tye_idx;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
552  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
553  | 
val assign_lb = assign_bound true;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
554  | 
val assign_ub = assign_bound false;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
555  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
556  | 
fun assign_alternating ts' ts G tye_idx =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
557  | 
if ts' = ts then tye_idx  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
558  | 
else  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
559  | 
let  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
560  | 
val (tye_idx' as (tye, _)) = fold (assign_lb G) ts tye_idx  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
561  | 
|> fold (assign_ub G) ts;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
562  | 
in  | 
| 40836 | 563  | 
assign_alternating ts  | 
564  | 
(filter (Type_Infer.is_paramT o Type_Infer.deref tye) ts) G tye_idx'  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
565  | 
end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
566  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
567  | 
(*Unify all weakly connected components of the constraint forest,  | 
| 40282 | 568  | 
that contain only params. These are the only WCCs that contain  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
569  | 
params anyway.*)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
570  | 
fun unify_params G (tye_idx as (tye, _)) =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
571  | 
let  | 
| 40286 | 572  | 
val max_params =  | 
573  | 
filter (Type_Infer.is_paramT o Type_Infer.deref tye) (Typ_Graph.maximals G);  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
574  | 
val to_unify = map (fn T => T :: get_preds G T) max_params;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
575  | 
in  | 
| 40836 | 576  | 
fold  | 
577  | 
(fn Ts => fn tye_idx' => unify_list Ts tye_idx'  | 
|
| 
41353
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
578  | 
handle NO_UNIFIER (msg, _) => err_list ctxt (gen_msg err msg) (fst tye_idx) Ts)  | 
| 40836 | 579  | 
to_unify tye_idx  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
580  | 
end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
581  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
582  | 
fun solve_constraints G tye_idx = tye_idx  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
583  | 
|> assign_alternating [] (Typ_Graph.keys G) G  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
584  | 
|> unify_params G;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
585  | 
in  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
586  | 
build_graph Typ_Graph.empty (map fst cs') tye_idx'  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
587  | 
|-> solve_constraints  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
588  | 
end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
589  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
590  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
591  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
592  | 
(** coercion insertion **)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
593  | 
|
| 40836 | 594  | 
fun gen_coercion ctxt tye (T1, T2) =  | 
595  | 
(case pairself (Type_Infer.deref tye) (T1, T2) of  | 
|
596  | 
((Type (a, [])), (Type (b, []))) =>  | 
|
597  | 
if a = b  | 
|
598  | 
then Abs (Name.uu, Type (a, []), Bound 0)  | 
|
599  | 
else  | 
|
600  | 
(case Symreltab.lookup (coes_of ctxt) (a, b) of  | 
|
601  | 
NONE => raise Fail (a ^ " is not a subtype of " ^ b)  | 
|
602  | 
| SOME co => co)  | 
|
603  | 
| ((Type (a, Ts)), (Type (b, Us))) =>  | 
|
604  | 
if a <> b  | 
|
605  | 
        then raise Fail ("Different constructors: " ^ a ^ " and " ^ b)
 | 
|
606  | 
else  | 
|
607  | 
let  | 
|
608  | 
fun inst t Ts =  | 
|
609  | 
Term.subst_vars  | 
|
610  | 
(((Term.add_tvar_namesT (fastype_of t) []) ~~ rev Ts), []) t;  | 
|
| 
41353
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
611  | 
fun sub_co (COVARIANT, TU) = SOME (gen_coercion ctxt tye TU)  | 
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
612  | 
| sub_co (CONTRAVARIANT, TU) = SOME (gen_coercion ctxt tye (swap TU))  | 
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
613  | 
| sub_co (INVARIANT_TO T, _) = NONE;  | 
| 40836 | 614  | 
fun ts_of [] = []  | 
615  | 
              | ts_of (Type ("fun", [x1, x2]) :: xs) = x1 :: x2 :: (ts_of xs);
 | 
|
616  | 
in  | 
|
617  | 
(case Symtab.lookup (tmaps_of ctxt) a of  | 
|
618  | 
              NONE => raise Fail ("No map function for " ^ a ^ " known")
 | 
|
619  | 
| SOME tmap =>  | 
|
620  | 
let  | 
|
| 
41353
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
621  | 
val used_coes = map_filter sub_co ((snd tmap) ~~ (Ts ~~ Us));  | 
| 40836 | 622  | 
in  | 
623  | 
Term.list_comb  | 
|
624  | 
(inst (fst tmap) (ts_of (map fastype_of used_coes)), used_coes)  | 
|
625  | 
end)  | 
|
626  | 
end  | 
|
627  | 
| (T, U) =>  | 
|
628  | 
if Type.could_unify (T, U)  | 
|
629  | 
then Abs (Name.uu, T, Bound 0)  | 
|
630  | 
        else raise Fail ("Cannot generate coercion from "
 | 
|
631  | 
^ Syntax.string_of_typ ctxt T ^ " to " ^ Syntax.string_of_typ ctxt U));  | 
|
632  | 
||
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
633  | 
fun insert_coercions ctxt tye ts =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
634  | 
let  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
635  | 
fun insert _ (Const (c, T)) =  | 
| 40836 | 636  | 
let val T' = T;  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
637  | 
in (Const (c, T'), T') end  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
638  | 
| insert _ (Free (x, T)) =  | 
| 40836 | 639  | 
let val T' = T;  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
640  | 
in (Free (x, T'), T') end  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
641  | 
| insert _ (Var (xi, T)) =  | 
| 40836 | 642  | 
let val T' = T;  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
643  | 
in (Var (xi, T'), T') end  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
644  | 
| insert bs (Bound i) =  | 
| 40836 | 645  | 
let val T = nth bs i handle Subscript => err_loose i;  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
646  | 
in (Bound i, T) end  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
647  | 
| insert bs (Abs (x, T, t)) =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
648  | 
let  | 
| 40836 | 649  | 
val T' = T;  | 
| 40282 | 650  | 
val (t', T'') = insert (T' :: bs) t;  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
651  | 
in  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
652  | 
(Abs (x, T', t'), T' --> T'')  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
653  | 
end  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
654  | 
| insert bs (t $ u) =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
655  | 
let  | 
| 40836 | 656  | 
            val (t', Type ("fun", [U, T])) = apsnd (Type_Infer.deref tye) (insert bs t);
 | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
657  | 
val (u', U') = insert bs u;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
658  | 
in  | 
| 40836 | 659  | 
if can (fn TU => strong_unify ctxt TU (tye, 0)) (U, U')  | 
660  | 
then (t' $ u', T)  | 
|
661  | 
else (t' $ (gen_coercion ctxt tye (U', U) $ u'), T)  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
662  | 
end  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
663  | 
in  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
664  | 
map (fst o insert []) ts  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
665  | 
end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
666  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
667  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
668  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
669  | 
(** assembling the pipeline **)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
670  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
671  | 
fun infer_types ctxt const_type var_type raw_ts =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
672  | 
let  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
673  | 
val (idx, ts) = Type_Infer.prepare ctxt const_type var_type raw_ts;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
674  | 
|
| 40836 | 675  | 
fun inf _ (t as (Const (_, T))) tye_idx = (t, T, tye_idx)  | 
676  | 
| inf _ (t as (Free (_, T))) tye_idx = (t, T, tye_idx)  | 
|
677  | 
| inf _ (t as (Var (_, T))) tye_idx = (t, T, tye_idx)  | 
|
678  | 
| inf bs (t as (Bound i)) tye_idx =  | 
|
679  | 
(t, snd (nth bs i handle Subscript => err_loose i), tye_idx)  | 
|
680  | 
| inf bs (Abs (x, T, t)) tye_idx =  | 
|
681  | 
let val (t', U, tye_idx') = inf ((x, T) :: bs) t tye_idx  | 
|
682  | 
in (Abs (x, T, t'), T --> U, tye_idx') end  | 
|
683  | 
| inf bs (t $ u) tye_idx =  | 
|
684  | 
let  | 
|
685  | 
val (t', T, tye_idx') = inf bs t tye_idx;  | 
|
686  | 
val (u', U, (tye, idx)) = inf bs u tye_idx';  | 
|
687  | 
val V = Type_Infer.mk_param idx [];  | 
|
688  | 
val (tu, tye_idx'') = (t' $ u', strong_unify ctxt (U --> V, T) (tye, idx + 1))  | 
|
689  | 
handle NO_UNIFIER (msg, tye') =>  | 
|
690  | 
raise TYPE_INFERENCE_ERROR (err_appl_msg ctxt msg tye' bs t T u U);  | 
|
691  | 
in (tu, V, tye_idx'') end;  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
692  | 
|
| 
40938
 
e258f6817add
use "fold_map" instead of "fold (fn .. => .. (ts @ [t], ..)) .."
 
traytel 
parents: 
40840 
diff
changeset
 | 
693  | 
fun infer_single t tye_idx =  | 
| 40836 | 694  | 
let val (t, _, tye_idx') = inf [] t tye_idx;  | 
| 
40938
 
e258f6817add
use "fold_map" instead of "fold (fn .. => .. (ts @ [t], ..)) .."
 
traytel 
parents: 
40840 
diff
changeset
 | 
695  | 
in (t, tye_idx') end;  | 
| 40836 | 696  | 
|
| 
40938
 
e258f6817add
use "fold_map" instead of "fold (fn .. => .. (ts @ [t], ..)) .."
 
traytel 
parents: 
40840 
diff
changeset
 | 
697  | 
val (ts', (tye, _)) = (fold_map infer_single ts (Vartab.empty, idx)  | 
| 40836 | 698  | 
handle TYPE_INFERENCE_ERROR err =>  | 
699  | 
let  | 
|
700  | 
fun gen_single t (tye_idx, constraints) =  | 
|
701  | 
let val (_, tye_idx', constraints') = generate_constraints ctxt err t tye_idx  | 
|
702  | 
in (tye_idx', constraints' @ constraints) end;  | 
|
703  | 
||
704  | 
val (tye_idx, constraints) = fold gen_single ts ((Vartab.empty, idx), []);  | 
|
705  | 
val (tye, idx) = process_constraints ctxt err constraints tye_idx;  | 
|
706  | 
in  | 
|
707  | 
(insert_coercions ctxt tye ts, (tye, idx))  | 
|
708  | 
end);  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
709  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
710  | 
val (_, ts'') = Type_Infer.finish ctxt tye ([], ts');  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
711  | 
in ts'' end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
712  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
713  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
714  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
715  | 
(** installation **)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
716  | 
|
| 40283 | 717  | 
(* term check *)  | 
718  | 
||
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
719  | 
fun coercion_infer_types ctxt =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
720  | 
infer_types ctxt  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
721  | 
(try (Consts.the_constraint (ProofContext.consts_of ctxt)))  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
722  | 
(ProofContext.def_type ctxt);  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
723  | 
|
| 
40939
 
2c150063cd4d
setup subtyping/coercions once in HOL.thy, but enable it only later via configuration option;
 
wenzelm 
parents: 
40938 
diff
changeset
 | 
724  | 
val (coercion_enabled, coercion_enabled_setup) = Attrib.config_bool "coercion_enabled" (K false);  | 
| 
 
2c150063cd4d
setup subtyping/coercions once in HOL.thy, but enable it only later via configuration option;
 
wenzelm 
parents: 
40938 
diff
changeset
 | 
725  | 
|
| 40283 | 726  | 
val add_term_check =  | 
727  | 
Syntax.add_term_check ~100 "coercions"  | 
|
728  | 
(fn xs => fn ctxt =>  | 
|
| 
40939
 
2c150063cd4d
setup subtyping/coercions once in HOL.thy, but enable it only later via configuration option;
 
wenzelm 
parents: 
40938 
diff
changeset
 | 
729  | 
if Config.get ctxt coercion_enabled then  | 
| 
 
2c150063cd4d
setup subtyping/coercions once in HOL.thy, but enable it only later via configuration option;
 
wenzelm 
parents: 
40938 
diff
changeset
 | 
730  | 
let val xs' = coercion_infer_types ctxt xs  | 
| 
 
2c150063cd4d
setup subtyping/coercions once in HOL.thy, but enable it only later via configuration option;
 
wenzelm 
parents: 
40938 
diff
changeset
 | 
731  | 
in if eq_list (op aconv) (xs, xs') then NONE else SOME (xs', ctxt) end  | 
| 
 
2c150063cd4d
setup subtyping/coercions once in HOL.thy, but enable it only later via configuration option;
 
wenzelm 
parents: 
40938 
diff
changeset
 | 
732  | 
else NONE);  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
733  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
734  | 
|
| 40283 | 735  | 
(* declarations *)  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
736  | 
|
| 
40284
 
c9acf88447e6
export declarations by default, to allow other ML packages by-pass concrete syntax;
 
wenzelm 
parents: 
40283 
diff
changeset
 | 
737  | 
fun add_type_map raw_t context =  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
738  | 
let  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
739  | 
val ctxt = Context.proof_of context;  | 
| 
40284
 
c9acf88447e6
export declarations by default, to allow other ML packages by-pass concrete syntax;
 
wenzelm 
parents: 
40283 
diff
changeset
 | 
740  | 
val t = singleton (Variable.polymorphic ctxt) raw_t;  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
741  | 
|
| 
41353
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
742  | 
fun err_str t = "\n\nThe provided function has the type\n" ^  | 
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
743  | 
Syntax.string_of_typ ctxt (fastype_of t) ^  | 
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
744  | 
"\n\nThe general type signature of a map function is" ^  | 
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
745  | 
"\nf1 => f2 => ... => fn => C [x1, ..., xn] => C [y1, ..., yn]" ^  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
746  | 
"\nwhere C is a constructor and fi is of type (xi => yi) or (yi => xi)";  | 
| 
41353
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
747  | 
|
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
748  | 
val ((fis, T1), T2) = apfst split_last (strip_type (fastype_of t))  | 
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
749  | 
      handle Empty => error ("Not a proper map function:" ^ err_str t);
 | 
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
750  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
751  | 
fun gen_arg_var ([], []) = []  | 
| 40282 | 752  | 
| gen_arg_var ((T, T') :: Ts, (U, U') :: Us) =  | 
| 
41353
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
753  | 
if U = U' then  | 
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
754  | 
if is_stypeT U then INVARIANT_TO U :: gen_arg_var ((T, T') :: Ts, Us)  | 
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
755  | 
            else error ("Invariant xi and yi should be base types:" ^ err_str t)
 | 
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
756  | 
else if T = U andalso T' = U' then COVARIANT :: gen_arg_var (Ts, Us)  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
757  | 
else if T = U' andalso T' = U then CONTRAVARIANT :: gen_arg_var (Ts, Us)  | 
| 
41353
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
758  | 
          else error ("Functions do not apply to arguments correctly:" ^ err_str t)
 | 
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
759  | 
| gen_arg_var (_, Ts) =  | 
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
760  | 
if forall (op = andf is_stypeT o fst) Ts  | 
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
761  | 
then map (INVARIANT_TO o fst) Ts  | 
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
762  | 
          else error ("Different numbers of functions and variant arguments\n" ^ err_str t);
 | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
763  | 
|
| 
41353
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
764  | 
(*retry flag needed to adjust the type lists, when given a map over type constructor fun*)  | 
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
765  | 
fun check_map_fun fis (Type (C1, Ts)) (Type (C2, Us)) retry =  | 
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
766  | 
if C1 = C2 andalso not (null fis) andalso forall is_funtype fis  | 
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
767  | 
then ((map dest_funT fis, Ts ~~ Us), C1)  | 
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
768  | 
          else error ("Not a proper map function:" ^ err_str t)
 | 
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
769  | 
| check_map_fun fis T1 T2 true =  | 
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
770  | 
let val (fis', T') = split_last fis  | 
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
771  | 
in check_map_fun fis' T' (T1 --> T2) false end  | 
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
772  | 
      | check_map_fun _ _ _ _ = error ("Not a proper map function:" ^ err_str t);
 | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
773  | 
|
| 
41353
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
774  | 
val res = check_map_fun fis T1 T2 true;  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
775  | 
val res_av = gen_arg_var (fst res);  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
776  | 
in  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
777  | 
map_tmaps (Symtab.update (snd res, (t, res_av))) context  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
778  | 
end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
779  | 
|
| 
40284
 
c9acf88447e6
export declarations by default, to allow other ML packages by-pass concrete syntax;
 
wenzelm 
parents: 
40283 
diff
changeset
 | 
780  | 
fun add_coercion raw_t context =  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
781  | 
let  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
782  | 
val ctxt = Context.proof_of context;  | 
| 
40284
 
c9acf88447e6
export declarations by default, to allow other ML packages by-pass concrete syntax;
 
wenzelm 
parents: 
40283 
diff
changeset
 | 
783  | 
val t = singleton (Variable.polymorphic ctxt) raw_t;  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
784  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
785  | 
    fun err_coercion () = error ("Bad type for coercion " ^
 | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
786  | 
Syntax.string_of_term ctxt t ^ ":\n" ^  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
787  | 
Syntax.string_of_typ ctxt (fastype_of t));  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
788  | 
|
| 40840 | 789  | 
val (T1, T2) = Term.dest_funT (fastype_of t)  | 
790  | 
handle TYPE _ => err_coercion ();  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
791  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
792  | 
val a =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
793  | 
(case T1 of  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
794  | 
Type (x, []) => x  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
795  | 
| _ => err_coercion ());  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
796  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
797  | 
val b =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
798  | 
(case T2 of  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
799  | 
Type (x, []) => x  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
800  | 
| _ => err_coercion ());  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
801  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
802  | 
fun coercion_data_update (tab, G) =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
803  | 
let  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
804  | 
val G' = maybe_new_nodes [a, b] G  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
805  | 
val G'' = Graph.add_edge_trans_acyclic (a, b) G'  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
806  | 
handle Graph.CYCLES _ => error (a ^ " is already a subtype of " ^ b ^  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
807  | 
"!\n\nCannot add coercion of type: " ^ a ^ " => " ^ b);  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
808  | 
val new_edges =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
809  | 
flat (Graph.dest G'' |> map (fn (x, ys) => ys |> map_filter (fn y =>  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
810  | 
if Graph.is_edge G' (x, y) then NONE else SOME (x, y))));  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
811  | 
val G_and_new = Graph.add_edge (a, b) G';  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
812  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
813  | 
fun complex_coercion tab G (a, b) =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
814  | 
let  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
815  | 
val path = hd (Graph.irreducible_paths G (a, b))  | 
| 40836 | 816  | 
val path' = fst (split_last path) ~~ tl path  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
817  | 
in Abs (Name.uu, Type (a, []),  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
818  | 
fold (fn t => fn u => t $ u) (map (the o Symreltab.lookup tab) path') (Bound 0))  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
819  | 
end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
820  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
821  | 
val tab' = fold  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
822  | 
(fn pair => fn tab => Symreltab.update (pair, complex_coercion tab G_and_new pair) tab)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
823  | 
(filter (fn pair => pair <> (a, b)) new_edges)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
824  | 
(Symreltab.update ((a, b), t) tab);  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
825  | 
in  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
826  | 
(tab', G'')  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
827  | 
end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
828  | 
in  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
829  | 
map_coes_and_graph coercion_data_update context  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
830  | 
end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
831  | 
|
| 40283 | 832  | 
|
833  | 
(* theory setup *)  | 
|
834  | 
||
835  | 
val setup =  | 
|
| 
40939
 
2c150063cd4d
setup subtyping/coercions once in HOL.thy, but enable it only later via configuration option;
 
wenzelm 
parents: 
40938 
diff
changeset
 | 
836  | 
coercion_enabled_setup #>  | 
| 40283 | 837  | 
Context.theory_map add_term_check #>  | 
| 
40284
 
c9acf88447e6
export declarations by default, to allow other ML packages by-pass concrete syntax;
 
wenzelm 
parents: 
40283 
diff
changeset
 | 
838  | 
  Attrib.setup @{binding coercion}
 | 
| 
 
c9acf88447e6
export declarations by default, to allow other ML packages by-pass concrete syntax;
 
wenzelm 
parents: 
40283 
diff
changeset
 | 
839  | 
(Args.term >> (fn t => Thm.declaration_attribute (K (add_coercion t))))  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
840  | 
"declaration of new coercions" #>  | 
| 40297 | 841  | 
  Attrib.setup @{binding coercion_map}
 | 
| 
40284
 
c9acf88447e6
export declarations by default, to allow other ML packages by-pass concrete syntax;
 
wenzelm 
parents: 
40283 
diff
changeset
 | 
842  | 
(Args.term >> (fn t => Thm.declaration_attribute (K (add_type_map t))))  | 
| 40283 | 843  | 
"declaration of new map functions";  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
844  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
845  | 
end;  |