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