| author | kuncar | 
| Tue, 13 Aug 2013 15:59:22 +0200 | |
| changeset 53010 | ec5e6f69bd65 | 
| parent 48272 | db75b4005d9a | 
| child 61262 | 7bd1eb4b056e | 
| permissions | -rw-r--r-- | 
| 2956 | 1  | 
(* Title: Pure/sorts.ML  | 
2  | 
Author: Markus Wenzel and Stefan Berghofer, TU Muenchen  | 
|
3  | 
||
| 19514 | 4  | 
The order-sorted algebra of type classes.  | 
| 19529 | 5  | 
|
6  | 
Classes denote (possibly empty) collections of types that are  | 
|
7  | 
partially ordered by class inclusion. They are represented  | 
|
8  | 
symbolically by strings.  | 
|
9  | 
||
10  | 
Sorts are intersections of finitely many classes. They are represented  | 
|
11  | 
by lists of classes. Normal forms of sorts are sorted lists of  | 
|
12  | 
minimal classes (wrt. current class inclusion).  | 
|
| 2956 | 13  | 
*)  | 
14  | 
||
15  | 
signature SORTS =  | 
|
16  | 
sig  | 
|
| 39687 | 17  | 
val make: sort list -> sort Ord_List.T  | 
18  | 
val subset: sort Ord_List.T * sort Ord_List.T -> bool  | 
|
19  | 
val union: sort Ord_List.T -> sort Ord_List.T -> sort Ord_List.T  | 
|
20  | 
val subtract: sort Ord_List.T -> sort Ord_List.T -> sort Ord_List.T  | 
|
21  | 
val remove_sort: sort -> sort Ord_List.T -> sort Ord_List.T  | 
|
22  | 
val insert_sort: sort -> sort Ord_List.T -> sort Ord_List.T  | 
|
23  | 
val insert_typ: typ -> sort Ord_List.T -> sort Ord_List.T  | 
|
24  | 
val insert_typs: typ list -> sort Ord_List.T -> sort Ord_List.T  | 
|
25  | 
val insert_term: term -> sort Ord_List.T -> sort Ord_List.T  | 
|
26  | 
val insert_terms: term list -> sort Ord_List.T -> sort Ord_List.T  | 
|
| 19645 | 27  | 
type algebra  | 
| 
36328
 
4d9deabf6474
replaced Sorts.rep_algebra by slightly more abstract selectors classes_of/arities_of;
 
wenzelm 
parents: 
36105 
diff
changeset
 | 
28  | 
val classes_of: algebra -> serial Graph.T  | 
| 
37248
 
8e8e5f9d1441
arities: no need to maintain original codomain (cf. f795c1164708) -- completion happens in axclass.ML;
 
wenzelm 
parents: 
36429 
diff
changeset
 | 
29  | 
val arities_of: algebra -> (class * sort list) list Symtab.table  | 
| 
21933
 
819ef284720b
classes: more direct way to achieve topological sorting;
 
wenzelm 
parents: 
21926 
diff
changeset
 | 
30  | 
val all_classes: algebra -> class list  | 
| 19645 | 31  | 
val super_classes: algebra -> class -> class list  | 
32  | 
val class_less: algebra -> class * class -> bool  | 
|
33  | 
val class_le: algebra -> class * class -> bool  | 
|
34  | 
val sort_eq: algebra -> sort * sort -> bool  | 
|
35  | 
val sort_le: algebra -> sort * sort -> bool  | 
|
36  | 
val sorts_le: algebra -> sort list * sort list -> bool  | 
|
37  | 
val inter_sort: algebra -> sort * sort -> sort  | 
|
| 24732 | 38  | 
val minimize_sort: algebra -> sort -> sort  | 
39  | 
val complete_sort: algebra -> sort -> sort  | 
|
| 39687 | 40  | 
val minimal_sorts: algebra -> sort list -> sort Ord_List.T  | 
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
41  | 
val add_class: Context.pretty -> class * class list -> algebra -> algebra  | 
| 
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
42  | 
val add_classrel: Context.pretty -> class * class -> algebra -> algebra  | 
| 
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
43  | 
val add_arities: Context.pretty -> string * (class * sort list) list -> algebra -> algebra  | 
| 19645 | 44  | 
val empty_algebra: algebra  | 
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
45  | 
val merge_algebra: Context.pretty -> algebra * algebra -> algebra  | 
| 
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
46  | 
val subalgebra: Context.pretty -> (class -> bool) -> (class * string -> sort list option)  | 
| 
22181
 
39104d1c43ca
added explicit query function for arities to subalgebra projection
 
haftmann 
parents: 
21933 
diff
changeset
 | 
47  | 
-> algebra -> (sort -> sort) * algebra  | 
| 
19578
 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 
wenzelm 
parents: 
19531 
diff
changeset
 | 
48  | 
type class_error  | 
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
49  | 
val class_error: Context.pretty -> class_error -> string  | 
| 
19578
 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 
wenzelm 
parents: 
19531 
diff
changeset
 | 
50  | 
exception CLASS_ERROR of class_error  | 
| 48272 | 51  | 
val has_instance: algebra -> string -> sort -> bool  | 
| 19645 | 52  | 
val mg_domain: algebra -> string -> sort -> sort list (*exception CLASS_ERROR*)  | 
| 28665 | 53  | 
val meet_sort: algebra -> typ * sort  | 
54  | 
-> sort Vartab.table -> sort Vartab.table (*exception CLASS_ERROR*)  | 
|
55  | 
val meet_sort_typ: algebra -> typ * sort -> typ -> typ (*exception CLASS_ERROR*)  | 
|
| 19645 | 56  | 
val of_sort: algebra -> typ * sort -> bool  | 
| 32791 | 57  | 
val of_sort_derivation: algebra ->  | 
| 
36102
 
a51d1d154c71
of_sort_derivation: pass-through full type information -- following the version by krauss/schropp;
 
wenzelm 
parents: 
35975 
diff
changeset
 | 
58  | 
    {class_relation: typ -> 'a * class -> class -> 'a,
 | 
| 
 
a51d1d154c71
of_sort_derivation: pass-through full type information -- following the version by krauss/schropp;
 
wenzelm 
parents: 
35975 
diff
changeset
 | 
59  | 
     type_constructor: string * typ list -> ('a * class) list list -> class -> 'a,
 | 
| 
22570
 
f166a5416b3f
renamed of_sort_derivation record fields (avoid clash with Alice keywords);
 
wenzelm 
parents: 
22364 
diff
changeset
 | 
60  | 
     type_variable: typ -> ('a * class) list} ->
 | 
| 19584 | 61  | 
typ * sort -> 'a list (*exception CLASS_ERROR*)  | 
| 
35961
 
00e48e1d9afd
Sorts.of_sort_derivation: do not use slow Graph.irreducible_paths here, which not always needed  (SUBTLE CHANGE IN SEMANTICS);
 
wenzelm 
parents: 
35408 
diff
changeset
 | 
62  | 
val classrel_derivation: algebra ->  | 
| 
 
00e48e1d9afd
Sorts.of_sort_derivation: do not use slow Graph.irreducible_paths here, which not always needed  (SUBTLE CHANGE IN SEMANTICS);
 
wenzelm 
parents: 
35408 
diff
changeset
 | 
63  | 
    ('a * class -> class -> 'a) -> 'a * class -> class -> 'a  (*exception CLASS_ERROR*)
 | 
| 
31946
 
99ac0321cd47
witness_sorts: proper type witnesses for hyps, not invented "'hyp" variables;
 
wenzelm 
parents: 
30242 
diff
changeset
 | 
64  | 
val witness_sorts: algebra -> string list -> (typ * sort) list -> sort list -> (typ * sort) list  | 
| 2956 | 65  | 
end;  | 
66  | 
||
| 20573 | 67  | 
structure Sorts: SORTS =  | 
| 2956 | 68  | 
struct  | 
69  | 
||
| 19514 | 70  | 
|
| 19529 | 71  | 
(** ordered lists of sorts **)  | 
| 
14782
 
d6ce35a1c386
incorporate sort ops from term.ML; use Graph.T; misc cleanup;
 
wenzelm 
parents: 
9281 
diff
changeset
 | 
72  | 
|
| 39687 | 73  | 
val make = Ord_List.make Term_Ord.sort_ord;  | 
74  | 
val subset = Ord_List.subset Term_Ord.sort_ord;  | 
|
75  | 
val union = Ord_List.union Term_Ord.sort_ord;  | 
|
76  | 
val subtract = Ord_List.subtract Term_Ord.sort_ord;  | 
|
| 
14782
 
d6ce35a1c386
incorporate sort ops from term.ML; use Graph.T; misc cleanup;
 
wenzelm 
parents: 
9281 
diff
changeset
 | 
77  | 
|
| 39687 | 78  | 
val remove_sort = Ord_List.remove Term_Ord.sort_ord;  | 
79  | 
val insert_sort = Ord_List.insert Term_Ord.sort_ord;  | 
|
| 
14782
 
d6ce35a1c386
incorporate sort ops from term.ML; use Graph.T; misc cleanup;
 
wenzelm 
parents: 
9281 
diff
changeset
 | 
80  | 
|
| 
16598
 
59381032be14
removed obsolete eq_sort, mem_sort, subset_sort, eq_set_sort, ins_sort, union_sort, rems_sort;
 
wenzelm 
parents: 
16400 
diff
changeset
 | 
81  | 
fun insert_typ (TFree (_, S)) Ss = insert_sort S Ss  | 
| 
 
59381032be14
removed obsolete eq_sort, mem_sort, subset_sort, eq_set_sort, ins_sort, union_sort, rems_sort;
 
wenzelm 
parents: 
16400 
diff
changeset
 | 
82  | 
| insert_typ (TVar (_, S)) Ss = insert_sort S Ss  | 
| 
 
59381032be14
removed obsolete eq_sort, mem_sort, subset_sort, eq_set_sort, ins_sort, union_sort, rems_sort;
 
wenzelm 
parents: 
16400 
diff
changeset
 | 
83  | 
| insert_typ (Type (_, Ts)) Ss = insert_typs Ts Ss  | 
| 
 
59381032be14
removed obsolete eq_sort, mem_sort, subset_sort, eq_set_sort, ins_sort, union_sort, rems_sort;
 
wenzelm 
parents: 
16400 
diff
changeset
 | 
84  | 
and insert_typs [] Ss = Ss  | 
| 
 
59381032be14
removed obsolete eq_sort, mem_sort, subset_sort, eq_set_sort, ins_sort, union_sort, rems_sort;
 
wenzelm 
parents: 
16400 
diff
changeset
 | 
85  | 
| insert_typs (T :: Ts) Ss = insert_typs Ts (insert_typ T Ss);  | 
| 
14782
 
d6ce35a1c386
incorporate sort ops from term.ML; use Graph.T; misc cleanup;
 
wenzelm 
parents: 
9281 
diff
changeset
 | 
86  | 
|
| 
16598
 
59381032be14
removed obsolete eq_sort, mem_sort, subset_sort, eq_set_sort, ins_sort, union_sort, rems_sort;
 
wenzelm 
parents: 
16400 
diff
changeset
 | 
87  | 
fun insert_term (Const (_, T)) Ss = insert_typ T Ss  | 
| 
 
59381032be14
removed obsolete eq_sort, mem_sort, subset_sort, eq_set_sort, ins_sort, union_sort, rems_sort;
 
wenzelm 
parents: 
16400 
diff
changeset
 | 
88  | 
| insert_term (Free (_, T)) Ss = insert_typ T Ss  | 
| 
 
59381032be14
removed obsolete eq_sort, mem_sort, subset_sort, eq_set_sort, ins_sort, union_sort, rems_sort;
 
wenzelm 
parents: 
16400 
diff
changeset
 | 
89  | 
| insert_term (Var (_, T)) Ss = insert_typ T Ss  | 
| 
 
59381032be14
removed obsolete eq_sort, mem_sort, subset_sort, eq_set_sort, ins_sort, union_sort, rems_sort;
 
wenzelm 
parents: 
16400 
diff
changeset
 | 
90  | 
| insert_term (Bound _) Ss = Ss  | 
| 
 
59381032be14
removed obsolete eq_sort, mem_sort, subset_sort, eq_set_sort, ins_sort, union_sort, rems_sort;
 
wenzelm 
parents: 
16400 
diff
changeset
 | 
91  | 
| insert_term (Abs (_, T, t)) Ss = insert_term t (insert_typ T Ss)  | 
| 
 
59381032be14
removed obsolete eq_sort, mem_sort, subset_sort, eq_set_sort, ins_sort, union_sort, rems_sort;
 
wenzelm 
parents: 
16400 
diff
changeset
 | 
92  | 
| insert_term (t $ u) Ss = insert_term t (insert_term u Ss);  | 
| 
14782
 
d6ce35a1c386
incorporate sort ops from term.ML; use Graph.T; misc cleanup;
 
wenzelm 
parents: 
9281 
diff
changeset
 | 
93  | 
|
| 
16598
 
59381032be14
removed obsolete eq_sort, mem_sort, subset_sort, eq_set_sort, ins_sort, union_sort, rems_sort;
 
wenzelm 
parents: 
16400 
diff
changeset
 | 
94  | 
fun insert_terms [] Ss = Ss  | 
| 
 
59381032be14
removed obsolete eq_sort, mem_sort, subset_sort, eq_set_sort, ins_sort, union_sort, rems_sort;
 
wenzelm 
parents: 
16400 
diff
changeset
 | 
95  | 
| insert_terms (t :: ts) Ss = insert_terms ts (insert_term t Ss);  | 
| 
14782
 
d6ce35a1c386
incorporate sort ops from term.ML; use Graph.T; misc cleanup;
 
wenzelm 
parents: 
9281 
diff
changeset
 | 
96  | 
|
| 
 
d6ce35a1c386
incorporate sort ops from term.ML; use Graph.T; misc cleanup;
 
wenzelm 
parents: 
9281 
diff
changeset
 | 
97  | 
|
| 19529 | 98  | 
|
99  | 
(** order-sorted algebra **)  | 
|
| 2956 | 100  | 
|
101  | 
(*  | 
|
| 
14782
 
d6ce35a1c386
incorporate sort ops from term.ML; use Graph.T; misc cleanup;
 
wenzelm 
parents: 
9281 
diff
changeset
 | 
102  | 
classes: graph representing class declarations together with proper  | 
| 
 
d6ce35a1c386
incorporate sort ops from term.ML; use Graph.T; misc cleanup;
 
wenzelm 
parents: 
9281 
diff
changeset
 | 
103  | 
subclass relation, which needs to be transitive and acyclic.  | 
| 2956 | 104  | 
|
| 
14782
 
d6ce35a1c386
incorporate sort ops from term.ML; use Graph.T; misc cleanup;
 
wenzelm 
parents: 
9281 
diff
changeset
 | 
105  | 
arities: table of association lists of all type arities; (t, ars)  | 
| 19531 | 106  | 
means that type constructor t has the arities ars; an element  | 
| 
37248
 
8e8e5f9d1441
arities: no need to maintain original codomain (cf. f795c1164708) -- completion happens in axclass.ML;
 
wenzelm 
parents: 
36429 
diff
changeset
 | 
107  | 
(c, Ss) of ars represents the arity t::(Ss)c. "Coregularity" of  | 
| 
 
8e8e5f9d1441
arities: no need to maintain original codomain (cf. f795c1164708) -- completion happens in axclass.ML;
 
wenzelm 
parents: 
36429 
diff
changeset
 | 
108  | 
the arities structure requires that for any two declarations  | 
| 
 
8e8e5f9d1441
arities: no need to maintain original codomain (cf. f795c1164708) -- completion happens in axclass.ML;
 
wenzelm 
parents: 
36429 
diff
changeset
 | 
109  | 
t::(Ss1)c1 and t::(Ss2)c2 such that c1 <= c2 holds Ss1 <= Ss2.  | 
| 2956 | 110  | 
*)  | 
111  | 
||
| 19645 | 112  | 
datatype algebra = Algebra of  | 
| 20573 | 113  | 
 {classes: serial Graph.T,
 | 
| 
37248
 
8e8e5f9d1441
arities: no need to maintain original codomain (cf. f795c1164708) -- completion happens in axclass.ML;
 
wenzelm 
parents: 
36429 
diff
changeset
 | 
114  | 
arities: (class * sort list) list Symtab.table};  | 
| 19645 | 115  | 
|
| 
36328
 
4d9deabf6474
replaced Sorts.rep_algebra by slightly more abstract selectors classes_of/arities_of;
 
wenzelm 
parents: 
36105 
diff
changeset
 | 
116  | 
fun classes_of (Algebra {classes, ...}) = classes;
 | 
| 
 
4d9deabf6474
replaced Sorts.rep_algebra by slightly more abstract selectors classes_of/arities_of;
 
wenzelm 
parents: 
36105 
diff
changeset
 | 
117  | 
fun arities_of (Algebra {arities, ...}) = arities;
 | 
| 19645 | 118  | 
|
119  | 
fun make_algebra (classes, arities) =  | 
|
120  | 
  Algebra {classes = classes, arities = arities};
 | 
|
121  | 
||
122  | 
fun map_classes f (Algebra {classes, arities}) = make_algebra (f classes, arities);
 | 
|
123  | 
fun map_arities f (Algebra {classes, arities}) = make_algebra (classes, f arities);
 | 
|
124  | 
||
125  | 
||
126  | 
(* classes *)  | 
|
127  | 
||
| 
21933
 
819ef284720b
classes: more direct way to achieve topological sorting;
 
wenzelm 
parents: 
21926 
diff
changeset
 | 
128  | 
fun all_classes (Algebra {classes, ...}) = Graph.all_preds classes (Graph.maximals classes);
 | 
| 
 
819ef284720b
classes: more direct way to achieve topological sorting;
 
wenzelm 
parents: 
21926 
diff
changeset
 | 
129  | 
|
| 
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: 
43792 
diff
changeset
 | 
130  | 
val super_classes = Graph.immediate_succs o classes_of;  | 
| 2956 | 131  | 
|
132  | 
||
| 19529 | 133  | 
(* class relations *)  | 
| 2956 | 134  | 
|
| 19645 | 135  | 
val class_less = Graph.is_edge o classes_of;  | 
136  | 
fun class_le algebra (c1, c2) = c1 = c2 orelse class_less algebra (c1, c2);  | 
|
| 2956 | 137  | 
|
138  | 
||
| 19529 | 139  | 
(* sort relations *)  | 
| 2956 | 140  | 
|
| 19645 | 141  | 
fun sort_le algebra (S1, S2) =  | 
| 23585 | 142  | 
S1 = S2 orelse forall (fn c2 => exists (fn c1 => class_le algebra (c1, c2)) S1) S2;  | 
| 2956 | 143  | 
|
| 19645 | 144  | 
fun sorts_le algebra (Ss1, Ss2) =  | 
145  | 
ListPair.all (sort_le algebra) (Ss1, Ss2);  | 
|
| 2956 | 146  | 
|
| 19645 | 147  | 
fun sort_eq algebra (S1, S2) =  | 
148  | 
sort_le algebra (S1, S2) andalso sort_le algebra (S2, S1);  | 
|
| 2956 | 149  | 
|
150  | 
||
| 19529 | 151  | 
(* intersection *)  | 
| 2956 | 152  | 
|
| 19645 | 153  | 
fun inter_class algebra c S =  | 
| 2956 | 154  | 
let  | 
155  | 
fun intr [] = [c]  | 
|
156  | 
| intr (S' as c' :: c's) =  | 
|
| 19645 | 157  | 
if class_le algebra (c', c) then S'  | 
158  | 
else if class_le algebra (c, c') then intr c's  | 
|
| 2956 | 159  | 
else c' :: intr c's  | 
160  | 
in intr S end;  | 
|
161  | 
||
| 19645 | 162  | 
fun inter_sort algebra (S1, S2) =  | 
163  | 
sort_strings (fold (inter_class algebra) S1 S2);  | 
|
| 2956 | 164  | 
|
165  | 
||
| 24732 | 166  | 
(* normal forms *)  | 
| 2956 | 167  | 
|
| 24732 | 168  | 
fun minimize_sort _ [] = []  | 
169  | 
| minimize_sort _ (S as [_]) = S  | 
|
170  | 
| minimize_sort algebra S =  | 
|
| 19645 | 171  | 
filter (fn c => not (exists (fn c' => class_less algebra (c', c)) S)) S  | 
| 19529 | 172  | 
|> sort_distinct string_ord;  | 
| 2990 | 173  | 
|
| 24732 | 174  | 
fun complete_sort algebra =  | 
175  | 
Graph.all_succs (classes_of algebra) o minimize_sort algebra;  | 
|
176  | 
||
| 28623 | 177  | 
fun minimal_sorts algebra raw_sorts =  | 
178  | 
let  | 
|
179  | 
fun le S1 S2 = sort_le algebra (S1, S2);  | 
|
180  | 
val sorts = make (map (minimize_sort algebra) raw_sorts);  | 
|
181  | 
in sorts |> filter_out (fn S => exists (fn S' => le S' S andalso not (le S S')) sorts) end;  | 
|
182  | 
||
| 2990 | 183  | 
|
| 2956 | 184  | 
|
| 19529 | 185  | 
(** build algebras **)  | 
| 19514 | 186  | 
|
187  | 
(* classes *)  | 
|
188  | 
||
| 
23655
 
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
 
wenzelm 
parents: 
23585 
diff
changeset
 | 
189  | 
fun err_dup_class c = error ("Duplicate declaration of class: " ^ quote c);
 | 
| 19514 | 190  | 
|
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
191  | 
fun err_cyclic_classes pp css =  | 
| 19514 | 192  | 
error (cat_lines (map (fn cs =>  | 
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
193  | 
"Cycle in class relation: " ^ Syntax.string_of_classrel (Syntax.init_pretty pp) cs) css));  | 
| 19514 | 194  | 
|
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
195  | 
fun add_class pp (c, cs) = map_classes (fn classes =>  | 
| 19514 | 196  | 
let  | 
| 20573 | 197  | 
val classes' = classes |> Graph.new_node (c, serial ())  | 
| 
23655
 
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
 
wenzelm 
parents: 
23585 
diff
changeset
 | 
198  | 
handle Graph.DUP dup => err_dup_class dup;  | 
| 19514 | 199  | 
val classes'' = classes' |> fold Graph.add_edge_trans_acyclic (map (pair c) cs)  | 
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
200  | 
handle Graph.CYCLES css => err_cyclic_classes pp css;  | 
| 19645 | 201  | 
in classes'' end);  | 
| 19514 | 202  | 
|
203  | 
||
204  | 
(* arities *)  | 
|
205  | 
||
206  | 
local  | 
|
207  | 
||
208  | 
fun for_classes _ NONE = ""  | 
|
| 42387 | 209  | 
| for_classes ctxt (SOME (c1, c2)) = " for classes " ^ Syntax.string_of_classrel ctxt [c1, c2];  | 
| 19514 | 210  | 
|
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
211  | 
fun err_conflict pp t cc (c, Ss) (c', Ss') =  | 
| 
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
212  | 
let val ctxt = Syntax.init_pretty pp in  | 
| 
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
213  | 
    error ("Conflict of type arities" ^ for_classes ctxt cc ^ ":\n  " ^
 | 
| 
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
214  | 
Syntax.string_of_arity ctxt (t, Ss, [c]) ^ " and\n " ^  | 
| 
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
215  | 
Syntax.string_of_arity ctxt (t, Ss', [c']))  | 
| 
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
216  | 
end;  | 
| 19514 | 217  | 
|
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
218  | 
fun coregular pp algebra t (c, Ss) ars =  | 
| 19514 | 219  | 
let  | 
| 
37248
 
8e8e5f9d1441
arities: no need to maintain original codomain (cf. f795c1164708) -- completion happens in axclass.ML;
 
wenzelm 
parents: 
36429 
diff
changeset
 | 
220  | 
fun conflict (c', Ss') =  | 
| 19645 | 221  | 
if class_le algebra (c, c') andalso not (sorts_le algebra (Ss, Ss')) then  | 
| 19514 | 222  | 
SOME ((c, c'), (c', Ss'))  | 
| 19645 | 223  | 
else if class_le algebra (c', c) andalso not (sorts_le algebra (Ss', Ss)) then  | 
| 19514 | 224  | 
SOME ((c', c), (c', Ss'))  | 
225  | 
else NONE;  | 
|
226  | 
in  | 
|
227  | 
(case get_first conflict ars of  | 
|
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
228  | 
SOME ((c1, c2), (c', Ss')) => err_conflict pp t (SOME (c1, c2)) (c, Ss) (c', Ss')  | 
| 
37248
 
8e8e5f9d1441
arities: no need to maintain original codomain (cf. f795c1164708) -- completion happens in axclass.ML;
 
wenzelm 
parents: 
36429 
diff
changeset
 | 
229  | 
| NONE => (c, Ss) :: ars)  | 
| 19514 | 230  | 
end;  | 
231  | 
||
| 
37248
 
8e8e5f9d1441
arities: no need to maintain original codomain (cf. f795c1164708) -- completion happens in axclass.ML;
 
wenzelm 
parents: 
36429 
diff
changeset
 | 
232  | 
fun complete algebra (c, Ss) = map (rpair Ss) (c :: super_classes algebra c);  | 
| 19645 | 233  | 
|
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
234  | 
fun insert pp algebra t (c, Ss) ars =  | 
| 19514 | 235  | 
(case AList.lookup (op =) ars c of  | 
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
236  | 
NONE => coregular pp algebra t (c, Ss) ars  | 
| 
37248
 
8e8e5f9d1441
arities: no need to maintain original codomain (cf. f795c1164708) -- completion happens in axclass.ML;
 
wenzelm 
parents: 
36429 
diff
changeset
 | 
237  | 
| SOME Ss' =>  | 
| 19645 | 238  | 
if sorts_le algebra (Ss, Ss') then ars  | 
| 
37248
 
8e8e5f9d1441
arities: no need to maintain original codomain (cf. f795c1164708) -- completion happens in axclass.ML;
 
wenzelm 
parents: 
36429 
diff
changeset
 | 
239  | 
else if sorts_le algebra (Ss', Ss)  | 
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
240  | 
then coregular pp algebra t (c, Ss) (remove (op =) (c, Ss') ars)  | 
| 
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
241  | 
else err_conflict pp t NONE (c, Ss) (c, Ss'));  | 
| 19514 | 242  | 
|
| 
35975
 
cef3c78ace0a
more efficient merge_algebra for important special cases -- tricky due to required completion if class algebra changes;
 
wenzelm 
parents: 
35961 
diff
changeset
 | 
243  | 
in  | 
| 
 
cef3c78ace0a
more efficient merge_algebra for important special cases -- tricky due to required completion if class algebra changes;
 
wenzelm 
parents: 
35961 
diff
changeset
 | 
244  | 
|
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
245  | 
fun insert_ars pp algebra t = fold_rev (insert pp algebra t);  | 
| 
35975
 
cef3c78ace0a
more efficient merge_algebra for important special cases -- tricky due to required completion if class algebra changes;
 
wenzelm 
parents: 
35961 
diff
changeset
 | 
246  | 
|
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
247  | 
fun insert_complete_ars pp algebra (t, ars) arities =  | 
| 19645 | 248  | 
let val ars' =  | 
249  | 
Symtab.lookup_list arities t  | 
|
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
250  | 
|> fold_rev (insert_ars pp algebra t) (map (complete algebra) ars);  | 
| 19645 | 251  | 
in Symtab.update (t, ars') arities end;  | 
| 19514 | 252  | 
|
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
253  | 
fun add_arities pp arg algebra =  | 
| 
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
254  | 
algebra |> map_arities (insert_complete_ars pp algebra arg);  | 
| 19514 | 255  | 
|
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
256  | 
fun add_arities_table pp algebra =  | 
| 
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
257  | 
Symtab.fold (fn (t, ars) => insert_complete_ars pp algebra (t, ars));  | 
| 19514 | 258  | 
|
259  | 
end;  | 
|
260  | 
||
| 19529 | 261  | 
|
| 19645 | 262  | 
(* classrel *)  | 
263  | 
||
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
264  | 
fun rebuild_arities pp algebra = algebra |> map_arities (fn arities =>  | 
| 19645 | 265  | 
Symtab.empty  | 
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
266  | 
|> add_arities_table pp algebra arities);  | 
| 19645 | 267  | 
|
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
268  | 
fun add_classrel pp rel = rebuild_arities pp o map_classes (fn classes =>  | 
| 19645 | 269  | 
classes |> Graph.add_edge_trans_acyclic rel  | 
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
270  | 
handle Graph.CYCLES css => err_cyclic_classes pp css);  | 
| 19645 | 271  | 
|
272  | 
||
273  | 
(* empty and merge *)  | 
|
274  | 
||
275  | 
val empty_algebra = make_algebra (Graph.empty, Symtab.empty);  | 
|
276  | 
||
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
277  | 
fun merge_algebra pp  | 
| 19645 | 278  | 
   (Algebra {classes = classes1, arities = arities1},
 | 
279  | 
    Algebra {classes = classes2, arities = arities2}) =
 | 
|
280  | 
let  | 
|
281  | 
val classes' = Graph.merge_trans_acyclic (op =) (classes1, classes2)  | 
|
| 
23655
 
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
 
wenzelm 
parents: 
23585 
diff
changeset
 | 
282  | 
handle Graph.DUP c => err_dup_class c  | 
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
283  | 
| Graph.CYCLES css => err_cyclic_classes pp css;  | 
| 19645 | 284  | 
val algebra0 = make_algebra (classes', Symtab.empty);  | 
| 
35975
 
cef3c78ace0a
more efficient merge_algebra for important special cases -- tricky due to required completion if class algebra changes;
 
wenzelm 
parents: 
35961 
diff
changeset
 | 
285  | 
val arities' =  | 
| 
 
cef3c78ace0a
more efficient merge_algebra for important special cases -- tricky due to required completion if class algebra changes;
 
wenzelm 
parents: 
35961 
diff
changeset
 | 
286  | 
(case (pointer_eq (classes1, classes2), pointer_eq (arities1, arities2)) of  | 
| 
 
cef3c78ace0a
more efficient merge_algebra for important special cases -- tricky due to required completion if class algebra changes;
 
wenzelm 
parents: 
35961 
diff
changeset
 | 
287  | 
(true, true) => arities1  | 
| 
 
cef3c78ace0a
more efficient merge_algebra for important special cases -- tricky due to required completion if class algebra changes;
 
wenzelm 
parents: 
35961 
diff
changeset
 | 
288  | 
| (true, false) => (*no completion*)  | 
| 
 
cef3c78ace0a
more efficient merge_algebra for important special cases -- tricky due to required completion if class algebra changes;
 
wenzelm 
parents: 
35961 
diff
changeset
 | 
289  | 
(arities1, arities2) |> Symtab.join (fn t => fn (ars1, ars2) =>  | 
| 
 
cef3c78ace0a
more efficient merge_algebra for important special cases -- tricky due to required completion if class algebra changes;
 
wenzelm 
parents: 
35961 
diff
changeset
 | 
290  | 
if pointer_eq (ars1, ars2) then raise Symtab.SAME  | 
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
291  | 
else insert_ars pp algebra0 t ars2 ars1)  | 
| 
35975
 
cef3c78ace0a
more efficient merge_algebra for important special cases -- tricky due to required completion if class algebra changes;
 
wenzelm 
parents: 
35961 
diff
changeset
 | 
292  | 
| (false, true) => (*unary completion*)  | 
| 
 
cef3c78ace0a
more efficient merge_algebra for important special cases -- tricky due to required completion if class algebra changes;
 
wenzelm 
parents: 
35961 
diff
changeset
 | 
293  | 
Symtab.empty  | 
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
294  | 
|> add_arities_table pp algebra0 arities1  | 
| 
35975
 
cef3c78ace0a
more efficient merge_algebra for important special cases -- tricky due to required completion if class algebra changes;
 
wenzelm 
parents: 
35961 
diff
changeset
 | 
295  | 
| (false, false) => (*binary completion*)  | 
| 
 
cef3c78ace0a
more efficient merge_algebra for important special cases -- tricky due to required completion if class algebra changes;
 
wenzelm 
parents: 
35961 
diff
changeset
 | 
296  | 
Symtab.empty  | 
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
297  | 
|> add_arities_table pp algebra0 arities1  | 
| 
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
298  | 
|> add_arities_table pp algebra0 arities2);  | 
| 19645 | 299  | 
in make_algebra (classes', arities') end;  | 
300  | 
||
| 
21933
 
819ef284720b
classes: more direct way to achieve topological sorting;
 
wenzelm 
parents: 
21926 
diff
changeset
 | 
301  | 
|
| 
37248
 
8e8e5f9d1441
arities: no need to maintain original codomain (cf. f795c1164708) -- completion happens in axclass.ML;
 
wenzelm 
parents: 
36429 
diff
changeset
 | 
302  | 
(* algebra projections *) (* FIXME potentially violates abstract type integrity *)  | 
| 28922 | 303  | 
|
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
304  | 
fun subalgebra pp P sargs (algebra as Algebra {classes, arities}) =
 | 
| 19952 | 305  | 
let  | 
| 24732 | 306  | 
val restrict_sort = minimize_sort algebra o filter P o Graph.all_succs classes;  | 
| 
37248
 
8e8e5f9d1441
arities: no need to maintain original codomain (cf. f795c1164708) -- completion happens in axclass.ML;
 
wenzelm 
parents: 
36429 
diff
changeset
 | 
307  | 
fun restrict_arity t (c, Ss) =  | 
| 
 
8e8e5f9d1441
arities: no need to maintain original codomain (cf. f795c1164708) -- completion happens in axclass.ML;
 
wenzelm 
parents: 
36429 
diff
changeset
 | 
308  | 
if P c then  | 
| 
 
8e8e5f9d1441
arities: no need to maintain original codomain (cf. f795c1164708) -- completion happens in axclass.ML;
 
wenzelm 
parents: 
36429 
diff
changeset
 | 
309  | 
(case sargs (c, t) of  | 
| 
 
8e8e5f9d1441
arities: no need to maintain original codomain (cf. f795c1164708) -- completion happens in axclass.ML;
 
wenzelm 
parents: 
36429 
diff
changeset
 | 
310  | 
SOME sorts =>  | 
| 
 
8e8e5f9d1441
arities: no need to maintain original codomain (cf. f795c1164708) -- completion happens in axclass.ML;
 
wenzelm 
parents: 
36429 
diff
changeset
 | 
311  | 
SOME (c, Ss |> map2 (curry (inter_sort algebra)) sorts |> map restrict_sort)  | 
| 
 
8e8e5f9d1441
arities: no need to maintain original codomain (cf. f795c1164708) -- completion happens in axclass.ML;
 
wenzelm 
parents: 
36429 
diff
changeset
 | 
312  | 
| NONE => NONE)  | 
| 
22181
 
39104d1c43ca
added explicit query function for arities to subalgebra projection
 
haftmann 
parents: 
21933 
diff
changeset
 | 
313  | 
else NONE;  | 
| 
46614
 
165886a4fe64
clarified Graph.restrict (formerly Graph.subgraph) based on public graph operations;
 
wenzelm 
parents: 
45595 
diff
changeset
 | 
314  | 
val classes' = classes |> Graph.restrict P;  | 
| 39020 | 315  | 
val arities' = arities |> Symtab.map (map_filter o restrict_arity);  | 
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
316  | 
in (restrict_sort, rebuild_arities pp (make_algebra (classes', arities'))) end;  | 
| 20465 | 317  | 
|
| 19645 | 318  | 
|
| 19529 | 319  | 
|
320  | 
(** sorts of types **)  | 
|
321  | 
||
| 
35961
 
00e48e1d9afd
Sorts.of_sort_derivation: do not use slow Graph.irreducible_paths here, which not always needed  (SUBTLE CHANGE IN SEMANTICS);
 
wenzelm 
parents: 
35408 
diff
changeset
 | 
322  | 
(* errors -- performance tuning via delayed message composition *)  | 
| 
19578
 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 
wenzelm 
parents: 
19531 
diff
changeset
 | 
323  | 
|
| 26639 | 324  | 
datatype class_error =  | 
| 36105 | 325  | 
No_Classrel of class * class |  | 
326  | 
No_Arity of string * class |  | 
|
327  | 
No_Subsort of sort * sort;  | 
|
| 19529 | 328  | 
|
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
329  | 
fun class_error pp =  | 
| 
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
330  | 
let val ctxt = Syntax.init_pretty pp in  | 
| 
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
331  | 
fn No_Classrel (c1, c2) => "No class relation " ^ Syntax.string_of_classrel ctxt [c1, c2]  | 
| 
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
332  | 
| No_Arity (a, c) => "No type arity " ^ Syntax.string_of_arity ctxt (a, [], [c])  | 
| 
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
333  | 
| No_Subsort (S1, S2) =>  | 
| 
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
334  | 
"Cannot derive subsort relation " ^  | 
| 
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
335  | 
Syntax.string_of_sort ctxt S1 ^ " < " ^ Syntax.string_of_sort ctxt S2  | 
| 
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46614 
diff
changeset
 | 
336  | 
end;  | 
| 19529 | 337  | 
|
| 
19578
 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 
wenzelm 
parents: 
19531 
diff
changeset
 | 
338  | 
exception CLASS_ERROR of class_error;  | 
| 
 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 
wenzelm 
parents: 
19531 
diff
changeset
 | 
339  | 
|
| 
 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 
wenzelm 
parents: 
19531 
diff
changeset
 | 
340  | 
|
| 48272 | 341  | 
(* instances *)  | 
342  | 
||
343  | 
fun has_instance algebra a =  | 
|
344  | 
forall (AList.defined (op =) (Symtab.lookup_list (arities_of algebra) a));  | 
|
| 19529 | 345  | 
|
| 19645 | 346  | 
fun mg_domain algebra a S =  | 
| 19529 | 347  | 
let  | 
| 48272 | 348  | 
val ars = Symtab.lookup_list (arities_of algebra) a;  | 
| 19529 | 349  | 
fun dom c =  | 
| 48272 | 350  | 
(case AList.lookup (op =) ars c of  | 
| 36105 | 351  | 
NONE => raise CLASS_ERROR (No_Arity (a, c))  | 
| 
37248
 
8e8e5f9d1441
arities: no need to maintain original codomain (cf. f795c1164708) -- completion happens in axclass.ML;
 
wenzelm 
parents: 
36429 
diff
changeset
 | 
352  | 
| SOME Ss => Ss);  | 
| 19645 | 353  | 
fun dom_inter c Ss = ListPair.map (inter_sort algebra) (dom c, Ss);  | 
| 19529 | 354  | 
in  | 
355  | 
(case S of  | 
|
356  | 
[] => raise Fail "Unknown domain of empty intersection"  | 
|
357  | 
| c :: cs => fold dom_inter cs (dom c))  | 
|
358  | 
end;  | 
|
359  | 
||
360  | 
||
| 26639 | 361  | 
(* meet_sort *)  | 
362  | 
||
363  | 
fun meet_sort algebra =  | 
|
364  | 
let  | 
|
365  | 
fun inters S S' = inter_sort algebra (S, S');  | 
|
366  | 
fun meet _ [] = I  | 
|
367  | 
| meet (TFree (_, S)) S' =  | 
|
368  | 
if sort_le algebra (S, S') then I  | 
|
| 36105 | 369  | 
else raise CLASS_ERROR (No_Subsort (S, S'))  | 
| 26639 | 370  | 
| meet (TVar (v, S)) S' =  | 
371  | 
if sort_le algebra (S, S') then I  | 
|
372  | 
else Vartab.map_default (v, S) (inters S')  | 
|
373  | 
| meet (Type (a, Ts)) S = fold2 meet Ts (mg_domain algebra a S);  | 
|
374  | 
in uncurry meet end;  | 
|
375  | 
||
| 28665 | 376  | 
fun meet_sort_typ algebra (T, S) =  | 
| 
37248
 
8e8e5f9d1441
arities: no need to maintain original codomain (cf. f795c1164708) -- completion happens in axclass.ML;
 
wenzelm 
parents: 
36429 
diff
changeset
 | 
377  | 
let val tab = meet_sort algebra (T, S) Vartab.empty;  | 
| 
 
8e8e5f9d1441
arities: no need to maintain original codomain (cf. f795c1164708) -- completion happens in axclass.ML;
 
wenzelm 
parents: 
36429 
diff
changeset
 | 
378  | 
in Term.map_type_tvar (fn (v, _) => TVar (v, (the o Vartab.lookup tab) v)) end;  | 
| 28665 | 379  | 
|
| 26639 | 380  | 
|
| 19529 | 381  | 
(* of_sort *)  | 
382  | 
||
| 19645 | 383  | 
fun of_sort algebra =  | 
| 19529 | 384  | 
let  | 
385  | 
fun ofS (_, []) = true  | 
|
| 19645 | 386  | 
| ofS (TFree (_, S), S') = sort_le algebra (S, S')  | 
387  | 
| ofS (TVar (_, S), S') = sort_le algebra (S, S')  | 
|
| 19529 | 388  | 
| ofS (Type (a, Ts), S) =  | 
| 19645 | 389  | 
let val Ss = mg_domain algebra a S in  | 
| 19529 | 390  | 
ListPair.all ofS (Ts, Ss)  | 
| 
19578
 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 
wenzelm 
parents: 
19531 
diff
changeset
 | 
391  | 
end handle CLASS_ERROR _ => false;  | 
| 19529 | 392  | 
in ofS end;  | 
393  | 
||
394  | 
||
| 27498 | 395  | 
(* animating derivations *)  | 
396  | 
||
| 32791 | 397  | 
fun of_sort_derivation algebra {class_relation, type_constructor, type_variable} =
 | 
| 19529 | 398  | 
let  | 
| 27555 | 399  | 
val arities = arities_of algebra;  | 
| 
19578
 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 
wenzelm 
parents: 
19531 
diff
changeset
 | 
400  | 
|
| 
36104
 
fecb587a1d0e
of_sort_derivation: weaken bypasses derivation for identical sort -- accomodate proof terms by krauss/schropp, for example;
 
wenzelm 
parents: 
36103 
diff
changeset
 | 
401  | 
fun weaken T D1 S2 =  | 
| 
 
fecb587a1d0e
of_sort_derivation: weaken bypasses derivation for identical sort -- accomodate proof terms by krauss/schropp, for example;
 
wenzelm 
parents: 
36103 
diff
changeset
 | 
402  | 
let val S1 = map snd D1 in  | 
| 
 
fecb587a1d0e
of_sort_derivation: weaken bypasses derivation for identical sort -- accomodate proof terms by krauss/schropp, for example;
 
wenzelm 
parents: 
36103 
diff
changeset
 | 
403  | 
if S1 = S2 then map fst D1  | 
| 
 
fecb587a1d0e
of_sort_derivation: weaken bypasses derivation for identical sort -- accomodate proof terms by krauss/schropp, for example;
 
wenzelm 
parents: 
36103 
diff
changeset
 | 
404  | 
else  | 
| 
 
fecb587a1d0e
of_sort_derivation: weaken bypasses derivation for identical sort -- accomodate proof terms by krauss/schropp, for example;
 
wenzelm 
parents: 
36103 
diff
changeset
 | 
405  | 
S2 |> map (fn c2 =>  | 
| 
 
fecb587a1d0e
of_sort_derivation: weaken bypasses derivation for identical sort -- accomodate proof terms by krauss/schropp, for example;
 
wenzelm 
parents: 
36103 
diff
changeset
 | 
406  | 
(case D1 |> find_first (fn (_, c1) => class_le algebra (c1, c2)) of  | 
| 
 
fecb587a1d0e
of_sort_derivation: weaken bypasses derivation for identical sort -- accomodate proof terms by krauss/schropp, for example;
 
wenzelm 
parents: 
36103 
diff
changeset
 | 
407  | 
SOME d1 => class_relation T d1 c2  | 
| 36105 | 408  | 
| NONE => raise CLASS_ERROR (No_Subsort (S1, S2))))  | 
| 
36104
 
fecb587a1d0e
of_sort_derivation: weaken bypasses derivation for identical sort -- accomodate proof terms by krauss/schropp, for example;
 
wenzelm 
parents: 
36103 
diff
changeset
 | 
409  | 
end;  | 
| 19529 | 410  | 
|
| 36103 | 411  | 
fun derive (_, []) = []  | 
| 
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: 
43792 
diff
changeset
 | 
412  | 
| derive (Type (a, Us), S) =  | 
| 19529 | 413  | 
let  | 
| 19645 | 414  | 
val Ss = mg_domain algebra a S;  | 
| 36103 | 415  | 
val dom = map2 (fn U => fn S => derive (U, S) ~~ S) Us Ss;  | 
| 19529 | 416  | 
in  | 
417  | 
S |> map (fn c =>  | 
|
418  | 
let  | 
|
| 
37248
 
8e8e5f9d1441
arities: no need to maintain original codomain (cf. f795c1164708) -- completion happens in axclass.ML;
 
wenzelm 
parents: 
36429 
diff
changeset
 | 
419  | 
val Ss' = the (AList.lookup (op =) (Symtab.lookup_list arities a) c);  | 
| 
36102
 
a51d1d154c71
of_sort_derivation: pass-through full type information -- following the version by krauss/schropp;
 
wenzelm 
parents: 
35975 
diff
changeset
 | 
420  | 
val dom' = map (fn ((U, d), S') => weaken U d S' ~~ S') ((Us ~~ dom) ~~ Ss');  | 
| 
37248
 
8e8e5f9d1441
arities: no need to maintain original codomain (cf. f795c1164708) -- completion happens in axclass.ML;
 
wenzelm 
parents: 
36429 
diff
changeset
 | 
421  | 
in type_constructor (a, Us) dom' c end)  | 
| 19529 | 422  | 
end  | 
| 36103 | 423  | 
| derive (T, S) = weaken T (type_variable T) S;  | 
424  | 
in derive end;  | 
|
| 19529 | 425  | 
|
| 
35961
 
00e48e1d9afd
Sorts.of_sort_derivation: do not use slow Graph.irreducible_paths here, which not always needed  (SUBTLE CHANGE IN SEMANTICS);
 
wenzelm 
parents: 
35408 
diff
changeset
 | 
426  | 
fun classrel_derivation algebra class_relation =  | 
| 
 
00e48e1d9afd
Sorts.of_sort_derivation: do not use slow Graph.irreducible_paths here, which not always needed  (SUBTLE CHANGE IN SEMANTICS);
 
wenzelm 
parents: 
35408 
diff
changeset
 | 
427  | 
let  | 
| 
 
00e48e1d9afd
Sorts.of_sort_derivation: do not use slow Graph.irreducible_paths here, which not always needed  (SUBTLE CHANGE IN SEMANTICS);
 
wenzelm 
parents: 
35408 
diff
changeset
 | 
428  | 
fun path (x, c1 :: c2 :: cs) = path (class_relation (x, c1) c2, c2 :: cs)  | 
| 
 
00e48e1d9afd
Sorts.of_sort_derivation: do not use slow Graph.irreducible_paths here, which not always needed  (SUBTLE CHANGE IN SEMANTICS);
 
wenzelm 
parents: 
35408 
diff
changeset
 | 
429  | 
| path (x, _) = x;  | 
| 
 
00e48e1d9afd
Sorts.of_sort_derivation: do not use slow Graph.irreducible_paths here, which not always needed  (SUBTLE CHANGE IN SEMANTICS);
 
wenzelm 
parents: 
35408 
diff
changeset
 | 
430  | 
in  | 
| 
 
00e48e1d9afd
Sorts.of_sort_derivation: do not use slow Graph.irreducible_paths here, which not always needed  (SUBTLE CHANGE IN SEMANTICS);
 
wenzelm 
parents: 
35408 
diff
changeset
 | 
431  | 
fn (x, c1) => fn c2 =>  | 
| 
 
00e48e1d9afd
Sorts.of_sort_derivation: do not use slow Graph.irreducible_paths here, which not always needed  (SUBTLE CHANGE IN SEMANTICS);
 
wenzelm 
parents: 
35408 
diff
changeset
 | 
432  | 
(case Graph.irreducible_paths (classes_of algebra) (c1, c2) of  | 
| 36105 | 433  | 
[] => raise CLASS_ERROR (No_Classrel (c1, c2))  | 
| 
35961
 
00e48e1d9afd
Sorts.of_sort_derivation: do not use slow Graph.irreducible_paths here, which not always needed  (SUBTLE CHANGE IN SEMANTICS);
 
wenzelm 
parents: 
35408 
diff
changeset
 | 
434  | 
| cs :: _ => path (x, cs))  | 
| 
 
00e48e1d9afd
Sorts.of_sort_derivation: do not use slow Graph.irreducible_paths here, which not always needed  (SUBTLE CHANGE IN SEMANTICS);
 
wenzelm 
parents: 
35408 
diff
changeset
 | 
435  | 
end;  | 
| 
 
00e48e1d9afd
Sorts.of_sort_derivation: do not use slow Graph.irreducible_paths here, which not always needed  (SUBTLE CHANGE IN SEMANTICS);
 
wenzelm 
parents: 
35408 
diff
changeset
 | 
436  | 
|
| 19529 | 437  | 
|
438  | 
(* witness_sorts *)  | 
|
439  | 
||
| 19645 | 440  | 
fun witness_sorts algebra types hyps sorts =  | 
| 19529 | 441  | 
let  | 
| 19645 | 442  | 
fun le S1 S2 = sort_le algebra (S1, S2);  | 
| 
31946
 
99ac0321cd47
witness_sorts: proper type witnesses for hyps, not invented "'hyp" variables;
 
wenzelm 
parents: 
30242 
diff
changeset
 | 
443  | 
fun get S2 (T, S1) = if le S1 S2 then SOME (T, S2) else NONE;  | 
| 19645 | 444  | 
fun mg_dom t S = SOME (mg_domain algebra t S) handle CLASS_ERROR _ => NONE;  | 
| 19529 | 445  | 
|
| 
19578
 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 
wenzelm 
parents: 
19531 
diff
changeset
 | 
446  | 
fun witn_sort _ [] solved_failed = (SOME (propT, []), solved_failed)  | 
| 
 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 
wenzelm 
parents: 
19531 
diff
changeset
 | 
447  | 
| witn_sort path S (solved, failed) =  | 
| 
 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 
wenzelm 
parents: 
19531 
diff
changeset
 | 
448  | 
if exists (le S) failed then (NONE, (solved, failed))  | 
| 19529 | 449  | 
else  | 
| 
31946
 
99ac0321cd47
witness_sorts: proper type witnesses for hyps, not invented "'hyp" variables;
 
wenzelm 
parents: 
30242 
diff
changeset
 | 
450  | 
(case get_first (get S) solved of  | 
| 
19578
 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 
wenzelm 
parents: 
19531 
diff
changeset
 | 
451  | 
SOME w => (SOME w, (solved, failed))  | 
| 19529 | 452  | 
| NONE =>  | 
| 
31946
 
99ac0321cd47
witness_sorts: proper type witnesses for hyps, not invented "'hyp" variables;
 
wenzelm 
parents: 
30242 
diff
changeset
 | 
453  | 
(case get_first (get S) hyps of  | 
| 
19578
 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 
wenzelm 
parents: 
19531 
diff
changeset
 | 
454  | 
SOME w => (SOME w, (w :: solved, failed))  | 
| 19584 | 455  | 
| NONE => witn_types path types S (solved, failed)))  | 
| 19529 | 456  | 
|
| 
19578
 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 
wenzelm 
parents: 
19531 
diff
changeset
 | 
457  | 
and witn_sorts path x = fold_map (witn_sort path) x  | 
| 19529 | 458  | 
|
| 
19578
 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 
wenzelm 
parents: 
19531 
diff
changeset
 | 
459  | 
and witn_types _ [] S (solved, failed) = (NONE, (solved, S :: failed))  | 
| 
 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 
wenzelm 
parents: 
19531 
diff
changeset
 | 
460  | 
| witn_types path (t :: ts) S solved_failed =  | 
| 19529 | 461  | 
(case mg_dom t S of  | 
462  | 
SOME SS =>  | 
|
463  | 
(*do not descend into stronger args (achieving termination)*)  | 
|
464  | 
if exists (fn D => le D S orelse exists (le D) path) SS then  | 
|
| 
19578
 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 
wenzelm 
parents: 
19531 
diff
changeset
 | 
465  | 
witn_types path ts S solved_failed  | 
| 19529 | 466  | 
else  | 
| 
19578
 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 
wenzelm 
parents: 
19531 
diff
changeset
 | 
467  | 
let val (ws, (solved', failed')) = witn_sorts (S :: path) SS solved_failed in  | 
| 19529 | 468  | 
if forall is_some ws then  | 
469  | 
let val w = (Type (t, map (#1 o the) ws), S)  | 
|
| 
19578
 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 
wenzelm 
parents: 
19531 
diff
changeset
 | 
470  | 
in (SOME w, (w :: solved', failed')) end  | 
| 
 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 
wenzelm 
parents: 
19531 
diff
changeset
 | 
471  | 
else witn_types path ts S (solved', failed')  | 
| 19529 | 472  | 
end  | 
| 
19578
 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 
wenzelm 
parents: 
19531 
diff
changeset
 | 
473  | 
| NONE => witn_types path ts S solved_failed);  | 
| 19529 | 474  | 
|
| 19584 | 475  | 
in map_filter I (#1 (witn_sorts [] sorts ([], []))) end;  | 
| 19529 | 476  | 
|
| 19514 | 477  | 
end;  |