| author | wenzelm | 
| Wed, 23 Mar 2022 13:05:54 +0100 | |
| changeset 75312 | e641ac92b489 | 
| parent 74561 | 8e6c973003c8 | 
| child 80634 | a90ab1ea6458 | 
| 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  | 
| 
40939
 
2c150063cd4d
setup subtyping/coercions once in HOL.thy, but enable it only later via configuration option;
 
wenzelm 
parents: 
40938 
diff
changeset
 | 
9  | 
val coercion_enabled: bool Config.T  | 
| 
40284
 
c9acf88447e6
export declarations by default, to allow other ML packages by-pass concrete syntax;
 
wenzelm 
parents: 
40283 
diff
changeset
 | 
10  | 
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
 | 
11  | 
val add_coercion: term -> Context.generic -> Context.generic  | 
| 45059 | 12  | 
val print_coercions: Proof.context -> unit  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
13  | 
end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
14  | 
|
| 40283 | 15  | 
structure Subtyping: SUBTYPING =  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
16  | 
struct  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
17  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
18  | 
(** coercions data **)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
19  | 
|
| 
41353
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
20  | 
datatype variance = COVARIANT | CONTRAVARIANT | INVARIANT | INVARIANT_TO of typ;  | 
| 51327 | 21  | 
datatype coerce_arg = PERMIT | FORBID | LEAVE  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
22  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
23  | 
datatype data = Data of  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
24  | 
  {coes: (term * ((typ list * typ list) * term list)) Symreltab.table,  (*coercions table*)
 | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
25  | 
(*full coercions graph - only used at coercion declaration/deletion*)  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
26  | 
full_graph: int Graph.T,  | 
| 52432 | 27  | 
(*coercions graph restricted to base types - for efficiency reasons stored in the context*)  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
28  | 
coes_graph: int Graph.T,  | 
| 
51319
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
29  | 
tmaps: (term * variance list) Symtab.table, (*map functions*)  | 
| 51327 | 30  | 
coerce_args: coerce_arg list Symtab.table (*special constants with non-coercible arguments*)};  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
31  | 
|
| 
51319
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
32  | 
fun make_data (coes, full_graph, coes_graph, tmaps, coerce_args) =  | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
33  | 
  Data {coes = coes, full_graph = full_graph, coes_graph = coes_graph,
 | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
34  | 
tmaps = tmaps, coerce_args = coerce_args};  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
35  | 
|
| 
45935
 
32f769f94ea4
meaningful error message on failing merges of coercion tables
 
traytel 
parents: 
45429 
diff
changeset
 | 
36  | 
fun merge_error_coes (a, b) =  | 
| 
 
32f769f94ea4
meaningful error message on failing merges of coercion tables
 
traytel 
parents: 
45429 
diff
changeset
 | 
37  | 
  error ("Cannot merge coercion tables.\nConflicting declarations for coercions from " ^
 | 
| 
 
32f769f94ea4
meaningful error message on failing merges of coercion tables
 
traytel 
parents: 
45429 
diff
changeset
 | 
38  | 
quote a ^ " to " ^ quote b ^ ".");  | 
| 
 
32f769f94ea4
meaningful error message on failing merges of coercion tables
 
traytel 
parents: 
45429 
diff
changeset
 | 
39  | 
|
| 
 
32f769f94ea4
meaningful error message on failing merges of coercion tables
 
traytel 
parents: 
45429 
diff
changeset
 | 
40  | 
fun merge_error_tmaps C =  | 
| 
 
32f769f94ea4
meaningful error message on failing merges of coercion tables
 
traytel 
parents: 
45429 
diff
changeset
 | 
41  | 
  error ("Cannot merge coercion map tables.\nConflicting declarations for the constructor " ^
 | 
| 
 
32f769f94ea4
meaningful error message on failing merges of coercion tables
 
traytel 
parents: 
45429 
diff
changeset
 | 
42  | 
quote C ^ ".");  | 
| 
 
32f769f94ea4
meaningful error message on failing merges of coercion tables
 
traytel 
parents: 
45429 
diff
changeset
 | 
43  | 
|
| 
51319
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
44  | 
fun merge_error_coerce_args C =  | 
| 
55303
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
45  | 
  error ("Cannot merge tables for constants with coercion-invariant arguments.\n" ^
 | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
46  | 
"Conflicting declarations for the constant " ^ quote C ^ ".");  | 
| 
51319
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
47  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
48  | 
structure Data = Generic_Data  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
49  | 
(  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
50  | 
type T = data;  | 
| 
51319
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
51  | 
val empty = make_data (Symreltab.empty, Graph.empty, Graph.empty, Symtab.empty, Symtab.empty);  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
52  | 
fun merge  | 
| 
51319
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
53  | 
    (Data {coes = coes1, full_graph = full_graph1, coes_graph = coes_graph1,
 | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
54  | 
tmaps = tmaps1, coerce_args = coerce_args1},  | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
55  | 
      Data {coes = coes2, full_graph = full_graph2, coes_graph = coes_graph2,
 | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
56  | 
tmaps = tmaps2, coerce_args = coerce_args2}) =  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
57  | 
make_data (Symreltab.merge (eq_pair (op aconv)  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
58  | 
(eq_pair (eq_pair (eq_list (op =)) (eq_list (op =))) (eq_list (op aconv))))  | 
| 
45935
 
32f769f94ea4
meaningful error message on failing merges of coercion tables
 
traytel 
parents: 
45429 
diff
changeset
 | 
59  | 
(coes1, coes2) handle Symreltab.DUP key => merge_error_coes key,  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
60  | 
Graph.merge (op =) (full_graph1, full_graph2),  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
61  | 
Graph.merge (op =) (coes_graph1, coes_graph2),  | 
| 
45935
 
32f769f94ea4
meaningful error message on failing merges of coercion tables
 
traytel 
parents: 
45429 
diff
changeset
 | 
62  | 
Symtab.merge (eq_pair (op aconv) (op =)) (tmaps1, tmaps2)  | 
| 
51319
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
63  | 
handle Symtab.DUP key => merge_error_tmaps key,  | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
64  | 
Symtab.merge (eq_list (op =)) (coerce_args1, coerce_args2)  | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
65  | 
handle Symtab.DUP key => merge_error_coerce_args key);  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
66  | 
);  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
67  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
68  | 
fun map_data f =  | 
| 
51319
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
69  | 
  Data.map (fn Data {coes, full_graph, coes_graph, tmaps, coerce_args} =>
 | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
70  | 
make_data (f (coes, full_graph, coes_graph, tmaps, coerce_args)));  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
71  | 
|
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
72  | 
fun map_coes_and_graphs f =  | 
| 
51319
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
73  | 
map_data (fn (coes, full_graph, coes_graph, tmaps, coerce_args) =>  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
74  | 
let val (coes', full_graph', coes_graph') = f (coes, full_graph, coes_graph);  | 
| 
51319
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
75  | 
in (coes', full_graph', coes_graph', tmaps, coerce_args) end);  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
76  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
77  | 
fun map_tmaps f =  | 
| 
51319
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
78  | 
map_data (fn (coes, full_graph, coes_graph, tmaps, coerce_args) =>  | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
79  | 
(coes, full_graph, coes_graph, f tmaps, coerce_args));  | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
80  | 
|
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
81  | 
fun map_coerce_args f =  | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
82  | 
map_data (fn (coes, full_graph, coes_graph, tmaps, coerce_args) =>  | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
83  | 
(coes, full_graph, coes_graph, tmaps, f coerce_args));  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
84  | 
|
| 40285 | 85  | 
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
 | 
86  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
87  | 
val coes_of = #coes o rep_data;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
88  | 
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
 | 
89  | 
val tmaps_of = #tmaps o rep_data;  | 
| 
51319
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
90  | 
val coerce_args_of = #coerce_args o rep_data;  | 
| 
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  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
93  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
94  | 
(** utils **)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
95  | 
|
| 
46614
 
165886a4fe64
clarified Graph.restrict (formerly Graph.subgraph) based on public graph operations;
 
wenzelm 
parents: 
45935 
diff
changeset
 | 
96  | 
fun restrict_graph G = Graph.restrict (fn x => Graph.get_node G x = 0) G;  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
97  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
98  | 
fun nameT (Type (s, [])) = s;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
99  | 
fun t_of s = Type (s, []);  | 
| 40286 | 100  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
101  | 
fun sort_of (TFree (_, S)) = SOME S  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
102  | 
| sort_of (TVar (_, S)) = SOME S  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
103  | 
| sort_of _ = NONE;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
104  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
105  | 
val is_typeT = fn (Type _) => true | _ => false;  | 
| 40282 | 106  | 
val is_compT = fn (Type (_, _ :: _)) => true | _ => false;  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
107  | 
val is_freeT = fn (TFree _) => true | _ => false;  | 
| 40286 | 108  | 
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
 | 
109  | 
val is_funtype = fn (Type ("fun", [_, _])) => true | _ => false;
 | 
| 51335 | 110  | 
|
111  | 
fun mk_identity T = Abs (Name.uu, T, Bound 0);  | 
|
| 
43591
 
d4cbd6feffdf
collapse map functions with identity subcoercions to identities;
 
traytel 
parents: 
43278 
diff
changeset
 | 
112  | 
val is_identity = fn (Abs (_, _, Bound 0)) => true | _ => false;  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
113  | 
|
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
114  | 
fun instantiate t Ts = Term.subst_TVars  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
115  | 
((Term.add_tvar_namesT (fastype_of t) []) ~~ rev Ts) t;  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
116  | 
|
| 
55303
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
117  | 
exception COERCION_GEN_ERROR of unit -> string * Buffer.T;  | 
| 
54584
 
2bbcbf8cf47e
possibility to fold coercion inference error messages; tuned;
 
traytel 
parents: 
54470 
diff
changeset
 | 
118  | 
|
| 
55303
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
119  | 
infixr ++> (*composition with deferred error message*)  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
120  | 
fun (err : unit -> string * Buffer.T) ++> s =  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
121  | 
err #> apsnd (Buffer.add s);  | 
| 
54584
 
2bbcbf8cf47e
possibility to fold coercion inference error messages; tuned;
 
traytel 
parents: 
54470 
diff
changeset
 | 
122  | 
|
| 
55303
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
123  | 
fun eval_err err =  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
124  | 
let val (s, buf) = err ()  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
125  | 
in s ^ Markup.markup Markup.text_fold (Buffer.content buf) end;  | 
| 
54584
 
2bbcbf8cf47e
possibility to fold coercion inference error messages; tuned;
 
traytel 
parents: 
54470 
diff
changeset
 | 
126  | 
|
| 
55303
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
127  | 
fun eval_error err = error (eval_err err);  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
128  | 
|
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
129  | 
fun inst_collect tye err T U =  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
130  | 
(case (T, Type_Infer.deref tye U) of  | 
| 
54584
 
2bbcbf8cf47e
possibility to fold coercion inference error messages; tuned;
 
traytel 
parents: 
54470 
diff
changeset
 | 
131  | 
(TVar (xi, _), U) => [(xi, U)]  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
132  | 
| (Type (a, Ts), Type (b, Us)) =>  | 
| 
55303
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
133  | 
if a <> b then eval_error err else inst_collects tye err Ts Us  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
134  | 
| (_, U') => if T <> U' then eval_error err else [])  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
135  | 
and inst_collects tye err Ts Us =  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
136  | 
fold2 (fn T => fn U => fn is => inst_collect tye err T U @ is) Ts Us [];  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
137  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
138  | 
|
| 40836 | 139  | 
(* unification *)  | 
| 
40281
 
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  | 
exception NO_UNIFIER of string * typ Vartab.table;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
142  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
143  | 
fun unify weak ctxt =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
144  | 
let  | 
| 42361 | 145  | 
val thy = Proof_Context.theory_of ctxt;  | 
| 59840 | 146  | 
val arity_sorts = Proof_Context.arity_sorts ctxt;  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
147  | 
|
| 40282 | 148  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
149  | 
(* adjust sorts of parameters *)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
150  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
151  | 
fun not_of_sort x S' S =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
152  | 
"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
 | 
153  | 
Syntax.string_of_sort ctxt S;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
154  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
155  | 
fun meet (_, []) tye_idx = tye_idx  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
156  | 
| 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
 | 
157  | 
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
 | 
158  | 
| 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
 | 
159  | 
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
 | 
160  | 
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
 | 
161  | 
| 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
 | 
162  | 
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
 | 
163  | 
else if Type_Infer.is_param xi then  | 
| 40286 | 164  | 
(Vartab.update_new  | 
165  | 
(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
 | 
166  | 
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
 | 
167  | 
and meets (T :: Ts, S :: Ss) (tye_idx as (tye, _)) =  | 
| 40286 | 168  | 
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
 | 
169  | 
| meets _ tye_idx = tye_idx;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
170  | 
|
| 55301 | 171  | 
val weak_meet = if weak then fn _ => I else meet;  | 
| 
40281
 
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  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
174  | 
(* occurs check and assignment *)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
175  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
176  | 
fun occurs_check tye xi (TVar (xi', _)) =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
177  | 
          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
 | 
178  | 
else  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
179  | 
(case Vartab.lookup tye xi' of  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
180  | 
NONE => ()  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
181  | 
| SOME T => occurs_check tye xi T)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
182  | 
| 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
 | 
183  | 
| occurs_check _ _ _ = ();  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
184  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
185  | 
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
 | 
186  | 
if xi = xi' then env  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
187  | 
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
 | 
188  | 
| assign xi T S (env as (tye, _)) =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
189  | 
(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
 | 
190  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
191  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
192  | 
(* unification *)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
193  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
194  | 
fun show_tycon (a, Ts) =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
195  | 
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
 | 
196  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
197  | 
fun unif (T1, T2) (env as (tye, _)) =  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58893 
diff
changeset
 | 
198  | 
(case apply2 (`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
 | 
199  | 
((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
 | 
200  | 
| ((_, 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
 | 
201  | 
| ((_, Type (a, Ts)), (_, Type (b, Us))) =>  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
202  | 
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
 | 
203  | 
else if a <> b then  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
204  | 
raise NO_UNIFIER  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
205  | 
              ("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
 | 
206  | 
else fold unif (Ts ~~ Us) env  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
207  | 
      | ((_, 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
 | 
208  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
209  | 
in unif end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
210  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
211  | 
val weak_unify = unify true;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
212  | 
val strong_unify = unify false;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
213  | 
|
| 
 
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  | 
(* Typ_Graph shortcuts *)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
216  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
217  | 
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
 | 
218  | 
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
 | 
219  | 
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
 | 
220  | 
fun maybe_new_typnodes Ts G = fold maybe_new_typnode Ts G;  | 
| 
44338
 
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
 
wenzelm 
parents: 
43591 
diff
changeset
 | 
221  | 
fun new_imm_preds G Ts = (* FIXME inefficient *)  | 
| 
 
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
 
wenzelm 
parents: 
43591 
diff
changeset
 | 
222  | 
subtract (op =) Ts (distinct (op =) (maps (Typ_Graph.immediate_preds G) Ts));  | 
| 
 
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
 
wenzelm 
parents: 
43591 
diff
changeset
 | 
223  | 
fun new_imm_succs G Ts = (* FIXME inefficient *)  | 
| 
 
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
 
wenzelm 
parents: 
43591 
diff
changeset
 | 
224  | 
subtract (op =) Ts (distinct (op =) (maps (Typ_Graph.immediate_succs G) Ts));  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
225  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
226  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
227  | 
(* Graph shortcuts *)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
228  | 
|
| 55301 | 229  | 
fun maybe_new_node s G = perhaps (try (Graph.new_node s)) G;  | 
230  | 
fun maybe_new_nodes ss G = fold maybe_new_node ss G;  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
231  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
232  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
233  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
234  | 
(** error messages **)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
235  | 
|
| 
54584
 
2bbcbf8cf47e
possibility to fold coercion inference error messages; tuned;
 
traytel 
parents: 
54470 
diff
changeset
 | 
236  | 
fun gen_err err msg =  | 
| 
55303
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
237  | 
  err ++> ("\nNow trying to infer coercions globally.\n\nCoercion inference failed" ^
 | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
238  | 
(if msg = "" then "" else ":\n" ^ msg) ^ "\n");  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
239  | 
|
| 
54584
 
2bbcbf8cf47e
possibility to fold coercion inference error messages; tuned;
 
traytel 
parents: 
54470 
diff
changeset
 | 
240  | 
val gen_msg = eval_err oo gen_err  | 
| 40836 | 241  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
242  | 
fun prep_output ctxt tye bs ts Ts =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
243  | 
let  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
244  | 
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
 | 
245  | 
val (Ts', Ts'') = chop (length Ts) Ts_bTs';  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
246  | 
fun prep t =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
247  | 
let val xs = rev (Term.variant_frees t (rev (map fst bs ~~ Ts'')))  | 
| 
49660
 
de49d9b4d7bc
more explicit Syntax_Trans.mark_bound_abs/mark_bound_body: preserve type information for show_markup;
 
wenzelm 
parents: 
49564 
diff
changeset
 | 
248  | 
in Term.subst_bounds (map Syntax_Trans.mark_bound_abs xs, t) end;  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
249  | 
in (map prep ts', Ts') end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
250  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
251  | 
fun err_loose i = error ("Loose bound variable: B." ^ string_of_int i);
 | 
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
252  | 
|
| 40836 | 253  | 
fun unif_failed msg =  | 
254  | 
"Type unification failed" ^ (if msg = "" then "" else ": " ^ msg) ^ "\n\n";  | 
|
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
255  | 
|
| 40836 | 256  | 
fun err_appl_msg ctxt msg tye bs t T u U () =  | 
| 55301 | 257  | 
let val ([t', u'], [T', U']) = prep_output ctxt tye bs [t, u] [T, U] in  | 
258  | 
(unif_failed msg ^ Type.appl_error ctxt t' T' u' U' ^ "\n\n",  | 
|
| 
55303
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
259  | 
Buffer.empty |> Buffer.add "Coercion Inference:\n\n")  | 
| 55301 | 260  | 
end;  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
261  | 
|
| 
54584
 
2bbcbf8cf47e
possibility to fold coercion inference error messages; tuned;
 
traytel 
parents: 
54470 
diff
changeset
 | 
262  | 
fun err_list ctxt err tye Ts =  | 
| 
55303
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
263  | 
let val (_, Ts') = prep_output ctxt tye [] [] Ts in  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
264  | 
eval_error (err ++>  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
265  | 
      ("\nCannot unify a list of types that should be the same:\n" ^
 | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
266  | 
Pretty.string_of (Pretty.list "[" "]" (map (Syntax.pretty_typ ctxt) Ts'))))  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
267  | 
end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
268  | 
|
| 
54584
 
2bbcbf8cf47e
possibility to fold coercion inference error messages; tuned;
 
traytel 
parents: 
54470 
diff
changeset
 | 
269  | 
fun err_bound ctxt err tye packs =  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
270  | 
let  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
271  | 
val (ts, Ts) = fold  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
272  | 
(fn (bs, t $ u, U, _, U') => fn (ts, Ts) =>  | 
| 40836 | 273  | 
let val (t', T') = prep_output ctxt tye bs [t, u] [U', U]  | 
| 40282 | 274  | 
in (t' :: ts, T' :: Ts) end)  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
275  | 
packs ([], []);  | 
| 
55303
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
276  | 
val msg =  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
277  | 
Pretty.string_of (Pretty.big_list "Cannot fulfil subtype constraints:"  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
278  | 
(map2 (fn [t, u] => fn [T, U] =>  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
279  | 
Pretty.block [  | 
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
280  | 
Syntax.pretty_typ ctxt T, Pretty.brk 2, Pretty.str "<:", Pretty.brk 2,  | 
| 
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
281  | 
Syntax.pretty_typ ctxt U, Pretty.brk 3,  | 
| 
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
282  | 
Pretty.str "from function application", Pretty.brk 2,  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
283  | 
Pretty.block [Syntax.pretty_term ctxt (t $ u)]])  | 
| 
55303
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
284  | 
ts Ts));  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
285  | 
  in eval_error (err ++> ("\n" ^ msg)) end;
 | 
| 
40281
 
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  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
288  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
289  | 
(** constraint generation **)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
290  | 
|
| 
51319
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
291  | 
fun update_coerce_arg ctxt old t =  | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
292  | 
let  | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
293  | 
val mk_coerce_args = the_default [] o Symtab.lookup (coerce_args_of ctxt);  | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
294  | 
fun update _ [] = old  | 
| 51327 | 295  | 
| update 0 (coerce :: _) = (case coerce of LEAVE => old | PERMIT => true | FORBID => false)  | 
| 
51319
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
296  | 
| update n (_ :: cs) = update (n - 1) cs;  | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
297  | 
val (f, n) = Term.strip_comb (Type.strip_constraints t) ||> length;  | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
298  | 
in  | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
299  | 
update n (case f of Const (name, _) => mk_coerce_args name | _ => [])  | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
300  | 
end;  | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
301  | 
|
| 40836 | 302  | 
fun generate_constraints ctxt err =  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
303  | 
let  | 
| 
51319
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
304  | 
fun gen _ cs _ (Const (_, T)) tye_idx = (T, tye_idx, cs)  | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
305  | 
| gen _ cs _ (Free (_, T)) tye_idx = (T, tye_idx, cs)  | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
306  | 
| gen _ cs _ (Var (_, T)) tye_idx = (T, tye_idx, cs)  | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
307  | 
| gen _ cs bs (Bound i) tye_idx =  | 
| 43278 | 308  | 
(snd (nth bs i handle General.Subscript => err_loose i), tye_idx, cs)  | 
| 
51319
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
309  | 
| gen coerce cs bs (Abs (x, T, t)) tye_idx =  | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
310  | 
let val (U, tye_idx', cs') = gen coerce cs ((x, T) :: bs) t tye_idx  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
311  | 
in (T --> U, tye_idx', cs') end  | 
| 
51319
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
312  | 
| gen coerce cs bs (t $ u) tye_idx =  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
313  | 
let  | 
| 
51319
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
314  | 
val (T, tye_idx', cs') = gen coerce cs bs t tye_idx;  | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
315  | 
val coerce' = update_coerce_arg ctxt coerce t;  | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
316  | 
val (U', (tye, idx), cs'') = gen coerce' cs' bs u tye_idx';  | 
| 40286 | 317  | 
val U = Type_Infer.mk_param idx [];  | 
318  | 
val V = Type_Infer.mk_param (idx + 1) [];  | 
|
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
319  | 
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
 | 
320  | 
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
 | 
321  | 
val error_pack = (bs, t $ u, U, V, U');  | 
| 52432 | 322  | 
in  | 
| 
51319
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
323  | 
if coerce'  | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
324  | 
then (V, tye_idx'', ((U', U), error_pack) :: cs'')  | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
325  | 
else (V,  | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
326  | 
strong_unify ctxt (U, U') tye_idx''  | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
327  | 
handle NO_UNIFIER (msg, _) => error (gen_msg err msg),  | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
328  | 
cs'')  | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
329  | 
end;  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
330  | 
in  | 
| 
51319
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
331  | 
gen true [] []  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
332  | 
end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
333  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
334  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
335  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
336  | 
(** constraint resolution **)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
337  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
338  | 
exception BOUND_ERROR of string;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
339  | 
|
| 40836 | 340  | 
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
 | 
341  | 
let  | 
| 42388 | 342  | 
val thy = Proof_Context.theory_of ctxt;  | 
343  | 
||
| 40285 | 344  | 
val coes_graph = coes_graph_of ctxt;  | 
345  | 
val tmaps = tmaps_of ctxt;  | 
|
| 59840 | 346  | 
val arity_sorts = Proof_Context.arity_sorts ctxt;  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
347  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
348  | 
fun split_cs _ [] = ([], [])  | 
| 40282 | 349  | 
| split_cs f (c :: cs) =  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58893 
diff
changeset
 | 
350  | 
(case apply2 f (fst c) of  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
351  | 
(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
 | 
352  | 
| _ => apfst (cons c) (split_cs f cs));  | 
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
353  | 
|
| 
41353
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
354  | 
fun unify_list (T :: Ts) tye_idx =  | 
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
355  | 
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
 | 
356  | 
|
| 40282 | 357  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
358  | 
(* check whether constraint simplification will terminate using weak unification *)  | 
| 40282 | 359  | 
|
| 
41353
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
360  | 
val _ = fold (fn (TU, _) => fn tye_idx =>  | 
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
361  | 
weak_unify ctxt TU tye_idx handle NO_UNIFIER (msg, _) =>  | 
| 40836 | 362  | 
        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
 | 
363  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
364  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
365  | 
(* simplify constraints *)  | 
| 40282 | 366  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
367  | 
fun simplify_constraints cs tye_idx =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
368  | 
let  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
369  | 
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
 | 
370  | 
let  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
371  | 
val arg_var =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
372  | 
(case Symtab.lookup tmaps a of  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
373  | 
(*everything is invariant for unknown constructors*)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
374  | 
NONE => replicate (length Ts) INVARIANT  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
375  | 
| SOME av => snd av);  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
376  | 
fun new_constraints (variance, constraint) (cs, tye_idx) =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
377  | 
(case variance of  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
378  | 
COVARIANT => (constraint :: cs, tye_idx)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
379  | 
| CONTRAVARIANT => (swap constraint :: cs, tye_idx)  | 
| 
41353
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
380  | 
| INVARIANT_TO T => (cs, unify_list [T, fst constraint, snd constraint] tye_idx  | 
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
381  | 
handle NO_UNIFIER (msg, _) =>  | 
| 
54584
 
2bbcbf8cf47e
possibility to fold coercion inference error messages; tuned;
 
traytel 
parents: 
54470 
diff
changeset
 | 
382  | 
err_list ctxt (gen_err err  | 
| 
 
2bbcbf8cf47e
possibility to fold coercion inference error messages; tuned;
 
traytel 
parents: 
54470 
diff
changeset
 | 
383  | 
                      ("failed to unify invariant arguments w.r.t. to the known map function\n" ^
 | 
| 
 
2bbcbf8cf47e
possibility to fold coercion inference error messages; tuned;
 
traytel 
parents: 
54470 
diff
changeset
 | 
384  | 
msg))  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
385  | 
(fst tye_idx) (T :: Ts))  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
386  | 
| INVARIANT => (cs, strong_unify ctxt constraint tye_idx  | 
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
387  | 
handle NO_UNIFIER (msg, _) =>  | 
| 51248 | 388  | 
                    error (gen_msg err ("failed to unify invariant arguments\n" ^ msg))));
 | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
389  | 
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
 | 
390  | 
(fold new_constraints (arg_var ~~ (Ts ~~ Us)) ([], (tye, idx)));  | 
| 
49142
 
0f81eca1e473
more conservative rechecking of processed constraints in subtyping constraint simplification
 
traytel 
parents: 
47060 
diff
changeset
 | 
391  | 
val test_update = is_typeT orf is_freeT orf is_fixedvarT;  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
392  | 
val (ch, done') =  | 
| 
51246
 
755562fd2d9d
apply unifying substitution before building the constraint graph
 
traytel 
parents: 
49660 
diff
changeset
 | 
393  | 
done  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58893 
diff
changeset
 | 
394  | 
|> map (apfst (apply2 (Type_Infer.deref tye')))  | 
| 
51246
 
755562fd2d9d
apply unifying substitution before building the constraint graph
 
traytel 
parents: 
49660 
diff
changeset
 | 
395  | 
|> (if not (null new) then rpair [] else split_cs test_update);  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
396  | 
val todo' = ch @ todo;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
397  | 
in  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
398  | 
simplify done' (new @ todo') (tye', idx')  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
399  | 
end  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
400  | 
(*xi is definitely a parameter*)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
401  | 
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
 | 
402  | 
let  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
403  | 
val n = length Ts;  | 
| 40286 | 404  | 
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
 | 
405  | 
val tye' = Vartab.update_new (xi, Type(a, args)) tye;  | 
| 40286 | 406  | 
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
 | 
407  | 
val todo' = ch @ todo;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
408  | 
val new =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
409  | 
if varleq then (Type(a, args), Type (a, Ts))  | 
| 40286 | 410  | 
else (Type (a, Ts), Type (a, args));  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
411  | 
in  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
412  | 
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
 | 
413  | 
end  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
414  | 
(*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
 | 
415  | 
and eliminate TU done todo tye idx =  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
416  | 
let  | 
| 40286 | 417  | 
val [TVar (xi, S)] = filter Type_Infer.is_paramT TU;  | 
418  | 
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
 | 
419  | 
val SOME S' = sort_of T;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
420  | 
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
 | 
421  | 
val tye' = Vartab.update_new (xi, T) tye;  | 
| 40286 | 422  | 
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
 | 
423  | 
val todo' = ch @ todo;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
424  | 
in  | 
| 42388 | 425  | 
if Sign.subsort thy (S', S) (*TODO check this*)  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
426  | 
then simplify done' todo' (tye', idx)  | 
| 40836 | 427  | 
else error (gen_msg err "sort mismatch")  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
428  | 
end  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
429  | 
and simplify done [] tye_idx = (done, tye_idx)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
430  | 
| simplify done (((T, U), error_pack) :: todo) (tye_idx as (tye, idx)) =  | 
| 40286 | 431  | 
(case (Type_Infer.deref tye T, Type_Infer.deref tye U) of  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
432  | 
(T1 as Type (a, []), T2 as Type (b, [])) =>  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
433  | 
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
 | 
434  | 
else if Graph.is_edge coes_graph (a, b) then simplify done todo tye_idx  | 
| 
55303
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
435  | 
else  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
436  | 
error (gen_msg err (quote (Syntax.string_of_typ ctxt T1) ^  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
437  | 
" is not a subtype of " ^ quote (Syntax.string_of_typ ctxt T2)))  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
438  | 
| (Type (a, Ts), Type (b, Us)) =>  | 
| 
55303
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
439  | 
if a <> b then  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
440  | 
error (gen_msg err "different constructors") (fst tye_idx) error_pack  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
441  | 
else contract a Ts Us error_pack done todo tye idx  | 
| 40282 | 442  | 
| (TVar (xi, S), Type (a, Ts as (_ :: _))) =>  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
443  | 
expand true xi S a Ts error_pack done todo tye idx  | 
| 40282 | 444  | 
| (Type (a, Ts as (_ :: _)), TVar (xi, S)) =>  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
445  | 
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
 | 
446  | 
| (T, U) =>  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
447  | 
if T = U then simplify done todo tye_idx  | 
| 40282 | 448  | 
else if exists (is_freeT orf is_fixedvarT) [T, U] andalso  | 
| 40286 | 449  | 
exists Type_Infer.is_paramT [T, U]  | 
| 
41353
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
450  | 
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
 | 
451  | 
else if exists (is_freeT orf is_fixedvarT) [T, U]  | 
| 40836 | 452  | 
then error (gen_msg err "not eliminated free/fixed variables")  | 
| 40282 | 453  | 
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
 | 
454  | 
in  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
455  | 
simplify [] cs tye_idx  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
456  | 
end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
457  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
458  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
459  | 
(* do simplification *)  | 
| 40282 | 460  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
461  | 
val (cs', tye_idx') = simplify_constraints cs tye_idx;  | 
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
462  | 
|
| 
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
463  | 
fun find_error_pack lower T' = map_filter  | 
| 40836 | 464  | 
(fn ((T, U), pack) => if if lower then T' = U else T' = T then SOME pack else NONE) cs';  | 
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
465  | 
|
| 
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
466  | 
fun find_cycle_packs nodes =  | 
| 40836 | 467  | 
let  | 
468  | 
val (but_last, last) = split_last nodes  | 
|
469  | 
val pairs = (last, hd nodes) :: (but_last ~~ tl nodes);  | 
|
470  | 
in  | 
|
471  | 
map_filter  | 
|
| 40838 | 472  | 
(fn (TU, pack) => if member (op =) pairs TU then SOME pack else NONE)  | 
| 40836 | 473  | 
cs'  | 
474  | 
end;  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
475  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
476  | 
(*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
 | 
477  | 
in terms of the subtype-relation (excluding T itself)*)  | 
| 40282 | 478  | 
fun styps super T =  | 
| 
44338
 
700008399ee5
refined Graph implementation: more abstract/scalable Graph.Keys instead of plain lists -- order of adjacency is now standardized wrt. Key.ord;
 
wenzelm 
parents: 
43591 
diff
changeset
 | 
479  | 
(if super then Graph.immediate_succs else Graph.immediate_preds) coes_graph T  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
480  | 
handle Graph.UNDEF _ => [];  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
481  | 
|
| 40282 | 482  | 
fun minmax sup (T :: Ts) =  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
483  | 
let  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
484  | 
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
 | 
485  | 
fun extract T [] = T  | 
| 40282 | 486  | 
| extract T (U :: Us) =  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
487  | 
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
 | 
488  | 
else if Graph.is_edge coes_graph (adjust U T) then extract U Us  | 
| 40836 | 489  | 
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
 | 
490  | 
in  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
491  | 
t_of (extract T Ts)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
492  | 
end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
493  | 
|
| 40282 | 494  | 
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
 | 
495  | 
let  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
496  | 
fun adjust T U = if super then (T, U) else (U, T);  | 
| 40282 | 497  | 
fun styp_test U Ts = forall  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
498  | 
(fn T => T = U orelse Graph.is_edge coes_graph (adjust U T)) Ts;  | 
| 55301 | 499  | 
fun fitting Ts S U = Sign.of_sort thy (t_of U, S) andalso styp_test U Ts;  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
500  | 
in  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
501  | 
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
 | 
502  | 
end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
503  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
504  | 
(* 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
 | 
505  | 
e.g. in the supremum case (sup = true):  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
506  | 
------- 'a::S---  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
507  | 
/ / \ \  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
508  | 
/ / \ \  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
509  | 
'b::C1 'c::C2 ... T1 T2 ...  | 
| 
 
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  | 
sorts - list of sorts [C1, C2, ...]  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
512  | 
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
 | 
513  | 
*)  | 
| 40282 | 514  | 
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
 | 
515  | 
let  | 
| 42388 | 516  | 
fun restriction T = Sign.of_sort thy (t_of T, S)  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
517  | 
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
 | 
518  | 
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
 | 
519  | 
in  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
520  | 
(case fold candidates Ts (filter restriction (T :: styps sup T)) of  | 
| 40836 | 521  | 
          [] => 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
 | 
522  | 
| [T] => t_of T  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
523  | 
| Ts => minmax sup Ts)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
524  | 
end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
525  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
526  | 
fun build_graph G [] tye_idx = (G, tye_idx)  | 
| 40282 | 527  | 
| 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
 | 
528  | 
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
 | 
529  | 
else  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
530  | 
let  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
531  | 
val G' = maybe_new_typnodes [T, U] G;  | 
| 45059 | 532  | 
val (G'', tye_idx') = (Typ_Graph.add_edge_acyclic (T, U) G', tye_idx)  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
533  | 
handle Typ_Graph.CYCLES cycles =>  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
534  | 
let  | 
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
535  | 
val (tye, idx) =  | 
| 
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
536  | 
fold  | 
| 40836 | 537  | 
(fn cycle => fn tye_idx' => (unify_list cycle tye_idx'  | 
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
538  | 
handle NO_UNIFIER (msg, _) =>  | 
| 
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
539  | 
err_bound ctxt  | 
| 
54584
 
2bbcbf8cf47e
possibility to fold coercion inference error messages; tuned;
 
traytel 
parents: 
54470 
diff
changeset
 | 
540  | 
                            (gen_err err ("constraint cycle not unifiable\n" ^ msg)) (fst tye_idx)
 | 
| 40836 | 541  | 
(find_cycle_packs cycle)))  | 
542  | 
cycles tye_idx  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
543  | 
in  | 
| 40836 | 544  | 
collapse (tye, idx) cycles G  | 
545  | 
end  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
546  | 
in  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
547  | 
build_graph G'' cs tye_idx'  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
548  | 
end  | 
| 40836 | 549  | 
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
 | 
550  | 
let  | 
| 40836 | 551  | 
(*all cycles collapse to one node,  | 
552  | 
because all of them share at least the nodes x and y*)  | 
|
553  | 
val nodes = (distinct (op =) (flat cycles));  | 
|
554  | 
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
 | 
555  | 
val P = new_imm_preds G nodes;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
556  | 
val S = new_imm_succs G nodes;  | 
| 
46665
 
919dfcdf6d8a
discontinued slightly odd Graph.del_nodes (inefficient due to full Table.map);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
557  | 
val G' = fold Typ_Graph.del_node (tl nodes) G;  | 
| 40836 | 558  | 
fun check_and_gen super T' =  | 
559  | 
let val U = Type_Infer.deref tye T';  | 
|
560  | 
in  | 
|
561  | 
if not (is_typeT T) orelse not (is_typeT U) orelse T = U  | 
|
562  | 
then if super then (hd nodes, T') else (T', hd nodes)  | 
|
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
563  | 
else  | 
| 
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
564  | 
if super andalso  | 
| 40836 | 565  | 
Graph.is_edge coes_graph (nameT T, nameT U) then (hd nodes, T')  | 
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
566  | 
else if not super andalso  | 
| 40836 | 567  | 
Graph.is_edge coes_graph (nameT U, nameT T) then (T', hd nodes)  | 
| 
55303
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
568  | 
else  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
569  | 
err_bound ctxt (gen_err err "cycle elimination produces inconsistent graph")  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
570  | 
(fst tye_idx)  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
571  | 
(maps find_cycle_packs cycles @ find_error_pack super T')  | 
| 40836 | 572  | 
end;  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
573  | 
in  | 
| 40836 | 574  | 
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
 | 
575  | 
end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
576  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
577  | 
fun assign_bound lower G key (tye_idx as (tye, _)) =  | 
| 40286 | 578  | 
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
 | 
579  | 
let  | 
| 40286 | 580  | 
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
 | 
581  | 
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
 | 
582  | 
val raw_bound = get_bound G key;  | 
| 40286 | 583  | 
val bound = map (Type_Infer.deref tye) raw_bound;  | 
584  | 
val not_params = filter_out Type_Infer.is_paramT bound;  | 
|
| 40282 | 585  | 
fun to_fulfil T =  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
586  | 
(case sort_of T of  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
587  | 
NONE => NONE  | 
| 40282 | 588  | 
| SOME S =>  | 
| 40286 | 589  | 
SOME  | 
590  | 
(map nameT  | 
|
| 
42405
 
13ecdb3057d8
split Type_Infer into early and late part, after Proof_Context;
 
wenzelm 
parents: 
42402 
diff
changeset
 | 
591  | 
(filter_out Type_Infer.is_paramT  | 
| 
 
13ecdb3057d8
split Type_Infer into early and late part, after Proof_Context;
 
wenzelm 
parents: 
42402 
diff
changeset
 | 
592  | 
(map (Type_Infer.deref tye) (get_bound G T))), S));  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
593  | 
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
 | 
594  | 
val assignment =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
595  | 
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
 | 
596  | 
else SOME (tightest lower S styps_and_sorts (map nameT not_params)  | 
| 
54584
 
2bbcbf8cf47e
possibility to fold coercion inference error messages; tuned;
 
traytel 
parents: 
54470 
diff
changeset
 | 
597  | 
handle BOUND_ERROR msg => err_bound ctxt (gen_err err msg) tye  | 
| 55301 | 598  | 
(maps (find_error_pack (not lower)) raw_bound));  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
599  | 
in  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
600  | 
(case assignment of  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
601  | 
NONE => tye_idx  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
602  | 
| SOME T =>  | 
| 40286 | 603  | 
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
 | 
604  | 
else if lower then (*upper bound check*)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
605  | 
let  | 
| 40286 | 606  | 
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
 | 
607  | 
val s = nameT T;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
608  | 
in  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
609  | 
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
 | 
610  | 
then apfst (Vartab.update (xi, T)) tye_idx  | 
| 
55303
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
611  | 
else  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
612  | 
err_bound ctxt  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
613  | 
(gen_err err  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
614  | 
(Pretty.string_of (Pretty.block  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
615  | 
[Pretty.str "assigned base type", Pretty.brk 1,  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
616  | 
Pretty.quote (Syntax.pretty_typ ctxt T), Pretty.brk 1,  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
617  | 
Pretty.str "clashes with the upper bound of variable", Pretty.brk 1,  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
618  | 
Syntax.pretty_typ ctxt (TVar (xi, S))])))  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
619  | 
tye  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
620  | 
(maps (find_error_pack lower) other_bound)  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
621  | 
end  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
622  | 
else apfst (Vartab.update (xi, T)) tye_idx)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
623  | 
end  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
624  | 
else tye_idx;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
625  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
626  | 
val assign_lb = assign_bound true;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
627  | 
val assign_ub = assign_bound false;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
628  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
629  | 
fun assign_alternating ts' ts G tye_idx =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
630  | 
if ts' = ts then tye_idx  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
631  | 
else  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
632  | 
let  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
633  | 
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
 | 
634  | 
|> fold (assign_ub G) ts;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
635  | 
in  | 
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
636  | 
assign_alternating ts  | 
| 40836 | 637  | 
(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
 | 
638  | 
end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
639  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
640  | 
(*Unify all weakly connected components of the constraint forest,  | 
| 40282 | 641  | 
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
 | 
642  | 
params anyway.*)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
643  | 
fun unify_params G (tye_idx as (tye, _)) =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
644  | 
let  | 
| 40286 | 645  | 
val max_params =  | 
646  | 
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
 | 
647  | 
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
 | 
648  | 
in  | 
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
649  | 
fold  | 
| 40836 | 650  | 
(fn Ts => fn tye_idx' => unify_list Ts tye_idx'  | 
| 
54584
 
2bbcbf8cf47e
possibility to fold coercion inference error messages; tuned;
 
traytel 
parents: 
54470 
diff
changeset
 | 
651  | 
handle NO_UNIFIER (msg, _) => err_list ctxt (gen_err err msg) (fst tye_idx) Ts)  | 
| 40836 | 652  | 
to_unify tye_idx  | 
| 
40281
 
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  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
655  | 
fun solve_constraints G tye_idx = tye_idx  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
656  | 
|> assign_alternating [] (Typ_Graph.keys G) G  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
657  | 
|> unify_params G;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
658  | 
in  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
659  | 
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
 | 
660  | 
|-> solve_constraints  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
661  | 
end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
662  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
663  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
664  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
665  | 
(** coercion insertion **)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
666  | 
|
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
667  | 
fun gen_coercion ctxt err tye TU =  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
668  | 
let  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58893 
diff
changeset
 | 
669  | 
fun gen (T1, T2) =  | 
| 
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58893 
diff
changeset
 | 
670  | 
(case apply2 (Type_Infer.deref tye) (T1, T2) of  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
671  | 
(T1 as (Type (a, [])), T2 as (Type (b, []))) =>  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
672  | 
if a = b  | 
| 51335 | 673  | 
then mk_identity T1  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
674  | 
else  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
675  | 
(case Symreltab.lookup (coes_of ctxt) (a, b) of  | 
| 
55303
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
676  | 
NONE =>  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
677  | 
raise COERCION_GEN_ERROR (err ++>  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
678  | 
(Pretty.string_of o Pretty.block)  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
679  | 
[Pretty.quote (Syntax.pretty_typ ctxt T1), Pretty.brk 1,  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
680  | 
Pretty.str "is not a subtype of", Pretty.brk 1,  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
681  | 
Pretty.quote (Syntax.pretty_typ ctxt T2)])  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
682  | 
| SOME (co, _) => co)  | 
| 
45102
 
7bb89635eb51
correct coercion generation in case of unknown map functions
 
traytel 
parents: 
45060 
diff
changeset
 | 
683  | 
| (T1 as Type (a, Ts), T2 as Type (b, Us)) =>  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
684  | 
if a <> b  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
685  | 
then  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
686  | 
(case Symreltab.lookup (coes_of ctxt) (a, b) of  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
687  | 
(*immediate error - cannot fix complex coercion with the global algorithm*)  | 
| 
55303
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
688  | 
NONE =>  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
689  | 
eval_error (err ++>  | 
| 55304 | 690  | 
                    ("No coercion known for type constructors: " ^
 | 
691  | 
quote (Proof_Context.markup_type ctxt a) ^ " and " ^  | 
|
692  | 
quote (Proof_Context.markup_type ctxt b)))  | 
|
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
693  | 
| SOME (co, ((Ts', Us'), _)) =>  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
694  | 
let  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
695  | 
val co_before = gen (T1, Type (a, Ts'));  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
696  | 
val coT = range_type (fastype_of co_before);  | 
| 
55303
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
697  | 
val insts =  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
698  | 
inst_collect tye (err ++> "Could not insert complex coercion")  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
699  | 
(domain_type (fastype_of co)) coT;  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
700  | 
val co' = Term.subst_TVars insts co;  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
701  | 
val co_after = gen (Type (b, (map (typ_subst_TVars insts) Us')), T2);  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
702  | 
in  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
703  | 
Abs (Name.uu, T1, Library.foldr (op $)  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
704  | 
(filter (not o is_identity) [co_after, co', co_before], Bound 0))  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
705  | 
end)  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
706  | 
else  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
707  | 
let  | 
| 51335 | 708  | 
fun sub_co (COVARIANT, TU) = (SOME (gen TU), NONE)  | 
709  | 
| sub_co (CONTRAVARIANT, TU) = (SOME (gen (swap TU)), NONE)  | 
|
710  | 
| sub_co (INVARIANT, (T, _)) = (NONE, SOME T)  | 
|
711  | 
| sub_co (INVARIANT_TO T, _) = (NONE, NONE);  | 
|
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
712  | 
fun ts_of [] = []  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
713  | 
                  | ts_of (Type ("fun", [x1, x2]) :: xs) = x1 :: x2 :: (ts_of xs);
 | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
714  | 
in  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
715  | 
(case Symtab.lookup (tmaps_of ctxt) a of  | 
| 
45102
 
7bb89635eb51
correct coercion generation in case of unknown map functions
 
traytel 
parents: 
45060 
diff
changeset
 | 
716  | 
NONE =>  | 
| 
 
7bb89635eb51
correct coercion generation in case of unknown map functions
 
traytel 
parents: 
45060 
diff
changeset
 | 
717  | 
if Type.could_unify (T1, T2)  | 
| 51335 | 718  | 
then mk_identity T1  | 
| 
55303
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
719  | 
else  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
720  | 
raise COERCION_GEN_ERROR  | 
| 55304 | 721  | 
(err ++>  | 
722  | 
                          ("No map function for " ^ quote (Proof_Context.markup_type ctxt a)
 | 
|
723  | 
^ " known"))  | 
|
| 51335 | 724  | 
| SOME (tmap, variances) =>  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
725  | 
let  | 
| 51335 | 726  | 
val (used_coes, invarTs) =  | 
727  | 
map_split sub_co (variances ~~ (Ts ~~ Us))  | 
|
728  | 
|>> map_filter I  | 
|
729  | 
||> map_filter I;  | 
|
730  | 
val Tinsts = ts_of (map fastype_of used_coes) @ invarTs;  | 
|
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
731  | 
in  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
732  | 
if null (filter (not o is_identity) used_coes)  | 
| 51335 | 733  | 
then mk_identity (Type (a, Ts))  | 
734  | 
else Term.list_comb (instantiate tmap Tinsts, used_coes)  | 
|
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
735  | 
end)  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
736  | 
end  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
737  | 
| (T, U) =>  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
738  | 
if Type.could_unify (T, U)  | 
| 51335 | 739  | 
then mk_identity T  | 
| 
55303
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
740  | 
else raise COERCION_GEN_ERROR (err ++>  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
741  | 
(Pretty.string_of o Pretty.block)  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
742  | 
[Pretty.str "Cannot generate coercion from", Pretty.brk 1,  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
743  | 
Pretty.quote (Syntax.pretty_typ ctxt T), Pretty.brk 1,  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
744  | 
Pretty.str "to", Pretty.brk 1,  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
745  | 
Pretty.quote (Syntax.pretty_typ ctxt U)]));  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
746  | 
in  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
747  | 
gen TU  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
748  | 
end;  | 
| 40836 | 749  | 
|
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
750  | 
fun function_of ctxt err tye T =  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
751  | 
(case Type_Infer.deref tye T of  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
752  | 
Type (C, Ts) =>  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
753  | 
(case Symreltab.lookup (coes_of ctxt) (C, "fun") of  | 
| 55304 | 754  | 
NONE =>  | 
755  | 
          eval_error (err ++> ("No complex coercion from " ^
 | 
|
756  | 
quote (Proof_Context.markup_type ctxt C) ^ " to " ^  | 
|
757  | 
quote (Proof_Context.markup_type ctxt "fun")))  | 
|
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
758  | 
| SOME (co, ((Ts', _), _)) =>  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
759  | 
let  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
760  | 
val co_before = gen_coercion ctxt err tye (Type (C, Ts), Type (C, Ts'));  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
761  | 
val coT = range_type (fastype_of co_before);  | 
| 
55303
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
762  | 
val insts =  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
763  | 
inst_collect tye (err ++> "Could not insert complex coercion")  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
764  | 
(domain_type (fastype_of co)) coT;  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
765  | 
val co' = Term.subst_TVars insts co;  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
766  | 
in  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
767  | 
Abs (Name.uu, Type (C, Ts), Library.foldr (op $)  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
768  | 
(filter (not o is_identity) [co', co_before], Bound 0))  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
769  | 
end)  | 
| 
55303
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
770  | 
| T' =>  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
771  | 
eval_error (err ++>  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
772  | 
(Pretty.string_of o Pretty.block)  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
773  | 
[Pretty.str "No complex coercion from", Pretty.brk 1,  | 
| 55304 | 774  | 
Pretty.quote (Syntax.pretty_typ ctxt T'), Pretty.brk 1,  | 
775  | 
Pretty.str "to", Pretty.brk 1, Proof_Context.pretty_type ctxt "fun"]));  | 
|
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
776  | 
|
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
777  | 
fun insert_coercions ctxt (tye, idx) ts =  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
778  | 
let  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
779  | 
fun insert _ (Const (c, T)) = (Const (c, T), T)  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
780  | 
| insert _ (Free (x, T)) = (Free (x, T), T)  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
781  | 
| insert _ (Var (xi, T)) = (Var (xi, T), T)  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
782  | 
| insert bs (Bound i) =  | 
| 43278 | 783  | 
let val T = nth bs i handle General.Subscript => err_loose i;  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
784  | 
in (Bound i, T) end  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
785  | 
| insert bs (Abs (x, T, t)) =  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
786  | 
let val (t', T') = insert (T :: bs) t;  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
787  | 
in (Abs (x, T, t'), T --> T') end  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
788  | 
| insert bs (t $ u) =  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
789  | 
let  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
790  | 
            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
 | 
791  | 
val (u', U') = insert bs u;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
792  | 
in  | 
| 40836 | 793  | 
if can (fn TU => strong_unify ctxt TU (tye, 0)) (U, U')  | 
794  | 
then (t' $ u', T)  | 
|
| 
55303
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
795  | 
            else (t' $ (gen_coercion ctxt (K ("", Buffer.empty)) tye (U', U) $ u'), T)
 | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
796  | 
end  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
797  | 
in  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
798  | 
map (fst o insert []) ts  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
799  | 
end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
800  | 
|
| 
 
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  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
803  | 
(** assembling the pipeline **)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
804  | 
|
| 42398 | 805  | 
fun coercion_infer_types ctxt raw_ts =  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
806  | 
let  | 
| 
42405
 
13ecdb3057d8
split Type_Infer into early and late part, after Proof_Context;
 
wenzelm 
parents: 
42402 
diff
changeset
 | 
807  | 
val (idx, ts) = Type_Infer_Context.prepare ctxt raw_ts;  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
808  | 
|
| 
51319
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
809  | 
fun inf _ _ (t as (Const (_, T))) tye_idx = (t, T, tye_idx)  | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
810  | 
| inf _ _ (t as (Free (_, T))) tye_idx = (t, T, tye_idx)  | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
811  | 
| inf _ _ (t as (Var (_, T))) tye_idx = (t, T, tye_idx)  | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
812  | 
| inf _ bs (t as (Bound i)) tye_idx =  | 
| 43278 | 813  | 
(t, snd (nth bs i handle General.Subscript => err_loose i), tye_idx)  | 
| 
51319
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
814  | 
| inf coerce bs (Abs (x, T, t)) tye_idx =  | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
815  | 
let val (t', U, tye_idx') = inf coerce ((x, T) :: bs) t tye_idx  | 
| 40836 | 816  | 
in (Abs (x, T, t'), T --> U, tye_idx') end  | 
| 
51319
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
817  | 
| inf coerce bs (t $ u) tye_idx =  | 
| 40836 | 818  | 
let  | 
| 
51319
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
819  | 
val (t', T, tye_idx') = inf coerce bs t tye_idx;  | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
820  | 
val coerce' = update_coerce_arg ctxt coerce t;  | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
821  | 
val (u', U, (tye, idx)) = inf coerce' bs u tye_idx';  | 
| 40836 | 822  | 
val V = Type_Infer.mk_param idx [];  | 
823  | 
val (tu, tye_idx'') = (t' $ u', strong_unify ctxt (U --> V, T) (tye, idx + 1))  | 
|
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
824  | 
handle NO_UNIFIER (msg, tye') =>  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
825  | 
let  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
826  | 
val err = err_appl_msg ctxt msg tye' bs t T u U;  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
827  | 
val W = Type_Infer.mk_param (idx + 1) [];  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
828  | 
val (t'', (tye', idx')) =  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
829  | 
(t', strong_unify ctxt (W --> V, T) (tye, idx + 2))  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
830  | 
handle NO_UNIFIER _ =>  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
831  | 
let  | 
| 
55303
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
832  | 
val err' = err ++> "Local coercion insertion on the operator failed:\n";  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
833  | 
val co = function_of ctxt err' tye T;  | 
| 
51319
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
834  | 
val (t'', T'', tye_idx'') = inf coerce bs (co $ t') (tye, idx + 2);  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
835  | 
in  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
836  | 
(t'', strong_unify ctxt (W --> V, T'') tye_idx''  | 
| 
55303
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
837  | 
handle NO_UNIFIER (msg, _) => eval_error (err' ++> msg))  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
838  | 
end;  | 
| 
55303
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
839  | 
val err' = err ++>  | 
| 
54584
 
2bbcbf8cf47e
possibility to fold coercion inference error messages; tuned;
 
traytel 
parents: 
54470 
diff
changeset
 | 
840  | 
((if t' aconv t'' then ""  | 
| 
55303
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
841  | 
else "Successfully coerced the operator to a function of type:\n" ^  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
842  | 
Syntax.string_of_typ ctxt  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
843  | 
(the_single (snd (prep_output ctxt tye' bs [] [W --> V]))) ^ "\n") ^  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
844  | 
(if coerce' then "Local coercion insertion on the operand failed:\n"  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
845  | 
else "Local coercion insertion on the operand disallowed:\n"));  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
846  | 
val (u'', U', tye_idx') =  | 
| 52432 | 847  | 
if coerce' then  | 
| 
51319
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
848  | 
let val co = gen_coercion ctxt err' tye' (U, W);  | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
849  | 
in inf coerce' bs (if is_identity co then u else co $ u) (tye', idx') end  | 
| 
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
850  | 
else (u, U, (tye', idx'));  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
851  | 
in  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
852  | 
(t'' $ u'', strong_unify ctxt (U', W) tye_idx'  | 
| 
55303
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
853  | 
handle NO_UNIFIER (msg, _) => raise COERCION_GEN_ERROR (err' ++> msg))  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
854  | 
end;  | 
| 40836 | 855  | 
in (tu, V, tye_idx'') end;  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
856  | 
|
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
857  | 
fun infer_single t tye_idx =  | 
| 
51319
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
858  | 
let val (t, _, tye_idx') = inf true [] t tye_idx  | 
| 
40938
 
e258f6817add
use "fold_map" instead of "fold (fn .. => .. (ts @ [t], ..)) .."
 
traytel 
parents: 
40840 
diff
changeset
 | 
859  | 
in (t, tye_idx') end;  | 
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
860  | 
|
| 
40938
 
e258f6817add
use "fold_map" instead of "fold (fn .. => .. (ts @ [t], ..)) .."
 
traytel 
parents: 
40840 
diff
changeset
 | 
861  | 
val (ts', (tye, _)) = (fold_map infer_single ts (Vartab.empty, idx)  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
862  | 
handle COERCION_GEN_ERROR err =>  | 
| 40836 | 863  | 
let  | 
864  | 
fun gen_single t (tye_idx, constraints) =  | 
|
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
865  | 
let val (_, tye_idx', constraints') =  | 
| 
55303
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
866  | 
generate_constraints ctxt (err ++> "\n") t tye_idx  | 
| 40836 | 867  | 
in (tye_idx', constraints' @ constraints) end;  | 
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
868  | 
|
| 40836 | 869  | 
val (tye_idx, constraints) = fold gen_single ts ((Vartab.empty, idx), []);  | 
| 
55303
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
870  | 
val (tye, idx) = process_constraints ctxt (err ++> "\n") constraints tye_idx;  | 
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
871  | 
in  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
872  | 
(insert_coercions ctxt (tye, idx) ts, (tye, idx))  | 
| 40836 | 873  | 
end);  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
874  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
875  | 
val (_, ts'') = Type_Infer.finish ctxt tye ([], ts');  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
876  | 
in ts'' end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
877  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
878  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
879  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
880  | 
(** installation **)  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
881  | 
|
| 40283 | 882  | 
(* term check *)  | 
883  | 
||
| 67149 | 884  | 
val coercion_enabled = Attrib.setup_config_bool \<^binding>\<open>coercion_enabled\<close> (K false);  | 
| 
40939
 
2c150063cd4d
setup subtyping/coercions once in HOL.thy, but enable it only later via configuration option;
 
wenzelm 
parents: 
40938 
diff
changeset
 | 
885  | 
|
| 58826 | 886  | 
val _ =  | 
887  | 
Theory.setup  | 
|
888  | 
(Context.theory_map  | 
|
889  | 
(Syntax_Phases.term_check ~100 "coercions"  | 
|
890  | 
(fn ctxt => Config.get ctxt coercion_enabled ? coercion_infer_types ctxt)));  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
891  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
892  | 
|
| 40283 | 893  | 
(* declarations *)  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
894  | 
|
| 
40284
 
c9acf88447e6
export declarations by default, to allow other ML packages by-pass concrete syntax;
 
wenzelm 
parents: 
40283 
diff
changeset
 | 
895  | 
fun add_type_map raw_t context =  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
896  | 
let  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
897  | 
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
 | 
898  | 
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
 | 
899  | 
|
| 45059 | 900  | 
fun err_str t = "\n\nThe provided function has the type:\n" ^  | 
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
901  | 
Syntax.string_of_typ ctxt (fastype_of t) ^  | 
| 45059 | 902  | 
"\n\nThe general type signature of a map function is:" ^  | 
| 
41353
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
903  | 
"\nf1 => f2 => ... => fn => C [x1, ..., xn] => C [y1, ..., yn]" ^  | 
| 45059 | 904  | 
"\nwhere C is a constructor and fi is of type (xi => yi) or (yi => xi).";  | 
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
905  | 
|
| 
41353
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
906  | 
val ((fis, T1), T2) = apfst split_last (strip_type (fastype_of t))  | 
| 
47060
 
e2741ec9ae36
prefer explicitly qualified exception List.Empty;
 
wenzelm 
parents: 
46961 
diff
changeset
 | 
907  | 
      handle List.Empty => error ("Not a proper map function:" ^ err_str t);
 | 
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42361 
diff
changeset
 | 
908  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
909  | 
fun gen_arg_var ([], []) = []  | 
| 51335 | 910  | 
| gen_arg_var (Ts, (U, U') :: Us) =  | 
| 
41353
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
911  | 
if U = U' then  | 
| 51335 | 912  | 
if null (Term.add_tvarsT U []) then INVARIANT_TO U :: gen_arg_var (Ts, Us)  | 
913  | 
else if Term.is_TVar U then INVARIANT :: gen_arg_var (Ts, Us)  | 
|
914  | 
            else error ("Invariant xi and yi should be variables or variable-free:" ^ err_str t)
 | 
|
915  | 
else  | 
|
916  | 
(case Ts of  | 
|
917  | 
              [] => error ("Different numbers of functions and variant arguments\n" ^ err_str t)
 | 
|
918  | 
| (T, T') :: Ts =>  | 
|
919  | 
if T = U andalso T' = U' then COVARIANT :: gen_arg_var (Ts, Us)  | 
|
920  | 
else if T = U' andalso T' = U then CONTRAVARIANT :: gen_arg_var (Ts, Us)  | 
|
921  | 
              else error ("Functions do not apply to arguments correctly:" ^ err_str t));
 | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
922  | 
|
| 
41353
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
923  | 
(*retry flag needed to adjust the type lists, when given a map over type constructor fun*)  | 
| 55305 | 924  | 
fun check_map_fun fis (Type (C1, Ts)) (Type (C2, Us)) _ =  | 
| 
41353
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
925  | 
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
 | 
926  | 
then ((map dest_funT fis, Ts ~~ Us), C1)  | 
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
927  | 
          else error ("Not a proper map function:" ^ err_str t)
 | 
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
928  | 
| check_map_fun fis T1 T2 true =  | 
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
929  | 
let val (fis', T') = split_last fis  | 
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
930  | 
in check_map_fun fis' T' (T1 --> T2) false end  | 
| 
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
931  | 
      | 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
 | 
932  | 
|
| 
41353
 
684003dbda54
Enabled non fully polymorphic map functions in subtyping
 
traytel 
parents: 
40939 
diff
changeset
 | 
933  | 
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
 | 
934  | 
val res_av = gen_arg_var (fst res);  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
935  | 
in  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
936  | 
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
 | 
937  | 
end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
938  | 
|
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
939  | 
fun transitive_coercion ctxt tab G (a, b) =  | 
| 45059 | 940  | 
let  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
941  | 
fun safe_app t (Abs (x, T', u)) =  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
942  | 
let  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
943  | 
val t' = map_types Type_Infer.paramify_vars t;  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
944  | 
in  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
945  | 
singleton (coercion_infer_types ctxt) (Abs(x, T', (t' $ u)))  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
946  | 
end;  | 
| 45059 | 947  | 
val path = hd (Graph.irreducible_paths G (a, b));  | 
948  | 
val path' = fst (split_last path) ~~ tl path;  | 
|
949  | 
val coercions = map (fst o the o Symreltab.lookup tab) path';  | 
|
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
950  | 
val trans_co = singleton (Variable.polymorphic ctxt)  | 
| 51335 | 951  | 
(fold safe_app coercions (mk_identity dummyT));  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58893 
diff
changeset
 | 
952  | 
val (Ts, Us) = apply2 (snd o Term.dest_Type) (Term.dest_funT (type_of trans_co));  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
953  | 
in  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
954  | 
(trans_co, ((Ts, Us), coercions))  | 
| 45059 | 955  | 
end;  | 
956  | 
||
| 
40284
 
c9acf88447e6
export declarations by default, to allow other ML packages by-pass concrete syntax;
 
wenzelm 
parents: 
40283 
diff
changeset
 | 
957  | 
fun add_coercion raw_t context =  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
958  | 
let  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
959  | 
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
 | 
960  | 
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
 | 
961  | 
|
| 
55303
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
962  | 
fun err_coercion () =  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
963  | 
      error ("Bad type for a coercion:\n" ^
 | 
| 45059 | 964  | 
Syntax.string_of_term ctxt t ^ " :: " ^  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
965  | 
Syntax.string_of_typ ctxt (fastype_of t));  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
966  | 
|
| 40840 | 967  | 
val (T1, T2) = Term.dest_funT (fastype_of t)  | 
968  | 
handle TYPE _ => err_coercion ();  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
969  | 
|
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
970  | 
val (a, Ts) = Term.dest_Type T1  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
971  | 
handle TYPE _ => err_coercion ();  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
972  | 
|
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
973  | 
val (b, Us) = Term.dest_Type T2  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
974  | 
handle TYPE _ => err_coercion ();  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
975  | 
|
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
976  | 
fun coercion_data_update (tab, G, _) =  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
977  | 
let  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
978  | 
val G' = maybe_new_nodes [(a, length Ts), (b, length Us)] G  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
979  | 
val G'' = Graph.add_edge_trans_acyclic (a, b) G'  | 
| 45059 | 980  | 
handle Graph.CYCLES _ => error (  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
981  | 
Syntax.string_of_typ ctxt T2 ^ " is already a subtype of " ^  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
982  | 
Syntax.string_of_typ ctxt T1 ^ "!\n\nCannot add coercion of type: " ^  | 
| 45059 | 983  | 
Syntax.string_of_typ ctxt (T1 --> T2));  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
984  | 
val new_edges =  | 
| 49560 | 985  | 
flat (Graph.dest G'' |> map (fn ((x, _), ys) => ys |> map_filter (fn y =>  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
986  | 
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
 | 
987  | 
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
 | 
988  | 
|
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
989  | 
val tab' = fold  | 
| 45059 | 990  | 
(fn pair => fn tab =>  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
991  | 
Symreltab.update (pair, transitive_coercion ctxt tab G_and_new pair) tab)  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
992  | 
(filter (fn pair => pair <> (a, b)) new_edges)  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
993  | 
(Symreltab.update ((a, b), (t, ((Ts, Us), []))) tab);  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
994  | 
in  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
995  | 
(tab', G'', restrict_graph G'')  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
996  | 
end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
997  | 
in  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
998  | 
map_coes_and_graphs coercion_data_update context  | 
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
999  | 
end;  | 
| 
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
1000  | 
|
| 45059 | 1001  | 
fun delete_coercion raw_t context =  | 
1002  | 
let  | 
|
1003  | 
val ctxt = Context.proof_of context;  | 
|
1004  | 
val t = singleton (Variable.polymorphic ctxt) raw_t;  | 
|
1005  | 
||
| 
55303
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
1006  | 
fun err_coercion the =  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
1007  | 
      error ("Not" ^
 | 
| 45059 | 1008  | 
(if the then " the defined " else " a ") ^ "coercion:\n" ^  | 
1009  | 
Syntax.string_of_term ctxt t ^ " :: " ^  | 
|
1010  | 
Syntax.string_of_typ ctxt (fastype_of t));  | 
|
1011  | 
||
1012  | 
val (T1, T2) = Term.dest_funT (fastype_of t)  | 
|
1013  | 
handle TYPE _ => err_coercion false;  | 
|
1014  | 
||
| 
54584
 
2bbcbf8cf47e
possibility to fold coercion inference error messages; tuned;
 
traytel 
parents: 
54470 
diff
changeset
 | 
1015  | 
val (a, _) = dest_Type T1  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
1016  | 
handle TYPE _ => err_coercion false;  | 
| 45059 | 1017  | 
|
| 
54584
 
2bbcbf8cf47e
possibility to fold coercion inference error messages; tuned;
 
traytel 
parents: 
54470 
diff
changeset
 | 
1018  | 
val (b, _) = dest_Type T2  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
1019  | 
handle TYPE _ => err_coercion false;  | 
| 45059 | 1020  | 
|
1021  | 
fun delete_and_insert tab G =  | 
|
1022  | 
let  | 
|
1023  | 
val pairs =  | 
|
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
1024  | 
Symreltab.fold (fn ((a, b), (_, (_, ts))) => fn pairs =>  | 
| 45059 | 1025  | 
if member (op aconv) ts t then (a, b) :: pairs else pairs) tab [(a, b)];  | 
1026  | 
fun delete pair (G, tab) = (Graph.del_edge pair G, Symreltab.delete_safe pair tab);  | 
|
1027  | 
val (G', tab') = fold delete pairs (G, tab);  | 
|
| 49564 | 1028  | 
fun reinsert pair (G, xs) =  | 
1029  | 
(case Graph.irreducible_paths G pair of  | 
|
1030  | 
[] => (G, xs)  | 
|
1031  | 
| _ => (Graph.add_edge pair G, (pair, transitive_coercion ctxt tab' G' pair) :: xs));  | 
|
| 45059 | 1032  | 
val (G'', ins) = fold reinsert pairs (G', []);  | 
1033  | 
in  | 
|
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
1034  | 
(fold Symreltab.update ins tab', G'', restrict_graph G'')  | 
| 45059 | 1035  | 
end  | 
1036  | 
||
| 
55303
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
1037  | 
fun show_term t =  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
1038  | 
Pretty.block [Syntax.pretty_term ctxt t,  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
1039  | 
Pretty.str " :: ", Syntax.pretty_typ ctxt (fastype_of t)];  | 
| 45059 | 1040  | 
|
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
1041  | 
fun coercion_data_update (tab, G, _) =  | 
| 
55303
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
1042  | 
(case Symreltab.lookup tab (a, b) of  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
1043  | 
NONE => err_coercion false  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
1044  | 
| SOME (t', (_, [])) =>  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
1045  | 
if t aconv t'  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
1046  | 
then delete_and_insert tab G  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
1047  | 
else err_coercion true  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
1048  | 
| SOME (t', (_, ts)) =>  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
1049  | 
if t aconv t' then  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
1050  | 
            error ("Cannot delete the automatically derived coercion:\n" ^
 | 
| 45059 | 1051  | 
Syntax.string_of_term ctxt t ^ " :: " ^  | 
| 
55303
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
1052  | 
Syntax.string_of_typ ctxt (fastype_of t) ^ "\n\n" ^  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
1053  | 
Pretty.string_of  | 
| 
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
1054  | 
(Pretty.big_list "Deleting one of the coercions:" (map show_term ts)) ^  | 
| 45059 | 1055  | 
"\nwill also remove the transitive coercion.")  | 
| 
55303
 
35354009ca25
clarified error messages: Pretty.str is for atomic/unformatted strings only, accumulate buffer of formatted strings instead;
 
wenzelm 
parents: 
55302 
diff
changeset
 | 
1056  | 
else err_coercion true);  | 
| 45059 | 1057  | 
in  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
1058  | 
map_coes_and_graphs coercion_data_update context  | 
| 45059 | 1059  | 
end;  | 
1060  | 
||
1061  | 
fun print_coercions ctxt =  | 
|
1062  | 
let  | 
|
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
1063  | 
fun separate _ [] = ([], [])  | 
| 52432 | 1064  | 
| separate P (x :: xs) = (if P x then apfst else apsnd) (cons x) (separate P xs);  | 
| 
45060
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
1065  | 
val (simple, complex) =  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
1066  | 
separate (fn (_, (_, ((Ts, Us), _))) => null Ts andalso null Us)  | 
| 
 
9c2568c0a504
local coercion insertion algorithm to support complex coercions
 
traytel 
parents: 
45059 
diff
changeset
 | 
1067  | 
(Symreltab.dest (coes_of ctxt));  | 
| 52432 | 1068  | 
fun show_coercion ((a, b), (t, ((Ts, Us), _))) =  | 
1069  | 
Pretty.item [Pretty.block  | 
|
1070  | 
[Syntax.pretty_typ ctxt (Type (a, Ts)), Pretty.brk 1,  | 
|
1071  | 
Pretty.str "<:", Pretty.brk 1,  | 
|
1072  | 
Syntax.pretty_typ ctxt (Type (b, Us)), Pretty.brk 3,  | 
|
1073  | 
Pretty.block  | 
|
| 55763 | 1074  | 
[Pretty.keyword2 "using", Pretty.brk 1,  | 
| 52432 | 1075  | 
Pretty.quote (Syntax.pretty_term ctxt t)]]];  | 
1076  | 
||
1077  | 
val type_space = Proof_Context.type_space ctxt;  | 
|
1078  | 
val tmaps =  | 
|
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58893 
diff
changeset
 | 
1079  | 
sort (Name_Space.extern_ord ctxt type_space o apply2 #1)  | 
| 52432 | 1080  | 
(Symtab.dest (tmaps_of ctxt));  | 
| 53539 | 1081  | 
fun show_map (c, (t, _)) =  | 
| 52432 | 1082  | 
Pretty.block  | 
| 53539 | 1083  | 
[Name_Space.pretty ctxt type_space c, Pretty.str ":",  | 
| 52432 | 1084  | 
Pretty.brk 1, Pretty.quote (Syntax.pretty_term ctxt t)];  | 
| 45059 | 1085  | 
in  | 
| 52432 | 1086  | 
[Pretty.big_list "coercions between base types:" (map show_coercion simple),  | 
1087  | 
Pretty.big_list "other coercions:" (map show_coercion complex),  | 
|
1088  | 
Pretty.big_list "coercion maps:" (map show_map tmaps)]  | 
|
| 
56334
 
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
 
wenzelm 
parents: 
55954 
diff
changeset
 | 
1089  | 
end |> Pretty.writeln_chunks;  | 
| 45059 | 1090  | 
|
1091  | 
||
| 58826 | 1092  | 
(* attribute setup *)  | 
| 40283 | 1093  | 
|
| 
51319
 
4a92178011e7
constants with coercion-invariant arguments (possibility to disable/reenable
 
traytel 
parents: 
51248 
diff
changeset
 | 
1094  | 
val parse_coerce_args =  | 
| 51327 | 1095  | 
Args.$$$ "+" >> K PERMIT || Args.$$$ "-" >> K FORBID || Args.$$$ "0" >> K LEAVE  | 
| 40283 | 1096  | 
|
| 58826 | 1097  | 
val _ =  | 
1098  | 
Theory.setup  | 
|
| 67149 | 1099  | 
(Attrib.setup \<^binding>\<open>coercion\<close>  | 
| 58826 | 1100  | 
(Args.term >> (fn t => Thm.declaration_attribute (K (add_coercion t))))  | 
1101  | 
"declaration of new coercions" #>  | 
|
| 67149 | 1102  | 
Attrib.setup \<^binding>\<open>coercion_delete\<close>  | 
| 58826 | 1103  | 
(Args.term >> (fn t => Thm.declaration_attribute (K (delete_coercion t))))  | 
1104  | 
"deletion of coercions" #>  | 
|
| 67149 | 1105  | 
Attrib.setup \<^binding>\<open>coercion_map\<close>  | 
| 58826 | 1106  | 
(Args.term >> (fn t => Thm.declaration_attribute (K (add_type_map t))))  | 
1107  | 
"declaration of new map functions" #>  | 
|
| 67149 | 1108  | 
Attrib.setup \<^binding>\<open>coercion_args\<close>  | 
| 58826 | 1109  | 
      (Args.const {proper = false, strict = false} -- Scan.lift (Scan.repeat1 parse_coerce_args) >>
 | 
1110  | 
(fn spec => Thm.declaration_attribute (K (map_coerce_args (Symtab.update spec)))))  | 
|
1111  | 
"declaration of new constants with coercion-invariant arguments");  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
1112  | 
|
| 45059 | 1113  | 
|
1114  | 
(* outer syntax commands *)  | 
|
1115  | 
||
1116  | 
val _ =  | 
|
| 67149 | 1117  | 
Outer_Syntax.command \<^command_keyword>\<open>print_coercions\<close>  | 
| 52432 | 1118  | 
"print information about coercions"  | 
1119  | 
(Scan.succeed (Toplevel.keep (print_coercions o Toplevel.context_of)));  | 
|
| 45059 | 1120  | 
|
| 
40281
 
3c6198fd0937
Coercive subtyping via subtype constraints, by Dmitriy Traytel (21-Oct-2010).
 
wenzelm 
parents:  
diff
changeset
 | 
1121  | 
end;  |