| author | wenzelm | 
| Thu, 01 Jan 2009 10:42:48 +0100 | |
| changeset 29283 | f4743512b12d | 
| parent 29269 | 5c25a2012975 | 
| child 29972 | aee7610106fd | 
| child 30240 | 5b25fee0362c | 
| 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 | |
| 28623 | 17 | val make: sort list -> sort OrdList.T | 
| 28374 | 18 | val subset: sort OrdList.T * sort OrdList.T -> bool | 
| 28354 | 19 | val union: sort OrdList.T -> sort OrdList.T -> sort OrdList.T | 
| 20 | val subtract: sort OrdList.T -> sort OrdList.T -> sort OrdList.T | |
| 21 | val remove_sort: sort -> sort OrdList.T -> sort OrdList.T | |
| 22 | val insert_sort: sort -> sort OrdList.T -> sort OrdList.T | |
| 23 | val insert_typ: typ -> sort OrdList.T -> sort OrdList.T | |
| 24 | val insert_typs: typ list -> sort OrdList.T -> sort OrdList.T | |
| 25 | val insert_term: term -> sort OrdList.T -> sort OrdList.T | |
| 26 | val insert_terms: term list -> sort OrdList.T -> sort OrdList.T | |
| 19645 | 27 | type algebra | 
| 28 | val rep_algebra: algebra -> | |
| 20573 | 29 |    {classes: serial Graph.T,
 | 
| 19645 | 30 | arities: (class * (class * sort list)) list Symtab.table} | 
| 21933 
819ef284720b
classes: more direct way to achieve topological sorting;
 wenzelm parents: 
21926diff
changeset | 31 | val all_classes: algebra -> class list | 
| 19645 | 32 | val super_classes: algebra -> class -> class list | 
| 33 | val class_less: algebra -> class * class -> bool | |
| 34 | val class_le: algebra -> class * class -> bool | |
| 35 | val sort_eq: algebra -> sort * sort -> bool | |
| 36 | val sort_le: algebra -> sort * sort -> bool | |
| 37 | val sorts_le: algebra -> sort list * sort list -> bool | |
| 38 | val inter_sort: algebra -> sort * sort -> sort | |
| 24732 | 39 | val minimize_sort: algebra -> sort -> sort | 
| 40 | val complete_sort: algebra -> sort -> sort | |
| 28623 | 41 | val minimal_sorts: algebra -> sort list -> sort OrdList.T | 
| 19645 | 42 | val certify_class: algebra -> class -> class (*exception TYPE*) | 
| 43 | val certify_sort: algebra -> sort -> sort (*exception TYPE*) | |
| 44 | val add_class: Pretty.pp -> class * class list -> algebra -> algebra | |
| 45 | val add_classrel: Pretty.pp -> class * class -> algebra -> algebra | |
| 46 | val add_arities: Pretty.pp -> string * (class * sort list) list -> algebra -> algebra | |
| 47 | val empty_algebra: algebra | |
| 48 | val merge_algebra: Pretty.pp -> algebra * algebra -> algebra | |
| 28922 | 49 | val classrels_of: algebra -> (class * class list) list | 
| 50 | val instances_of: algebra -> (string * class) list | |
| 22181 
39104d1c43ca
added explicit query function for arities to subalgebra projection
 haftmann parents: 
21933diff
changeset | 51 | val subalgebra: Pretty.pp -> (class -> bool) -> (class * string -> sort list) | 
| 
39104d1c43ca
added explicit query function for arities to subalgebra projection
 haftmann parents: 
21933diff
changeset | 52 | -> algebra -> (sort -> sort) * algebra | 
| 19578 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 wenzelm parents: 
19531diff
changeset | 53 | type class_error | 
| 26639 | 54 | val class_error: Pretty.pp -> class_error -> string | 
| 19578 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 wenzelm parents: 
19531diff
changeset | 55 | exception CLASS_ERROR of class_error | 
| 19645 | 56 | val mg_domain: algebra -> string -> sort -> sort list (*exception CLASS_ERROR*) | 
| 28665 | 57 | val meet_sort: algebra -> typ * sort | 
| 58 | -> sort Vartab.table -> sort Vartab.table (*exception CLASS_ERROR*) | |
| 59 | val meet_sort_typ: algebra -> typ * sort -> typ -> typ (*exception CLASS_ERROR*) | |
| 19645 | 60 | val of_sort: algebra -> typ * sort -> bool | 
| 27555 | 61 |   val weaken: algebra -> ('a * class -> class -> 'a) -> 'a * class -> class -> 'a
 | 
| 19645 | 62 | val of_sort_derivation: Pretty.pp -> algebra -> | 
| 22570 
f166a5416b3f
renamed of_sort_derivation record fields (avoid clash with Alice keywords);
 wenzelm parents: 
22364diff
changeset | 63 |     {class_relation: 'a * class -> class -> 'a,
 | 
| 
f166a5416b3f
renamed of_sort_derivation record fields (avoid clash with Alice keywords);
 wenzelm parents: 
22364diff
changeset | 64 |      type_constructor: string -> ('a * class) list list -> class -> 'a,
 | 
| 
f166a5416b3f
renamed of_sort_derivation record fields (avoid clash with Alice keywords);
 wenzelm parents: 
22364diff
changeset | 65 |      type_variable: typ -> ('a * class) list} ->
 | 
| 19584 | 66 | typ * sort -> 'a list (*exception CLASS_ERROR*) | 
| 19645 | 67 | val witness_sorts: algebra -> string list -> 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 | |
| 29269 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 wenzelm parents: 
28922diff
changeset | 76 | val make = OrdList.make TermOrd.sort_ord; | 
| 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 wenzelm parents: 
28922diff
changeset | 77 | val op subset = OrdList.subset TermOrd.sort_ord; | 
| 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 wenzelm parents: 
28922diff
changeset | 78 | val op union = OrdList.union TermOrd.sort_ord; | 
| 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 wenzelm parents: 
28922diff
changeset | 79 | val subtract = OrdList.subtract TermOrd.sort_ord; | 
| 14782 
d6ce35a1c386
incorporate sort ops from term.ML; use Graph.T; misc cleanup;
 wenzelm parents: 
9281diff
changeset | 80 | |
| 29269 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 wenzelm parents: 
28922diff
changeset | 81 | val remove_sort = OrdList.remove TermOrd.sort_ord; | 
| 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 wenzelm parents: 
28922diff
changeset | 82 | val insert_sort = OrdList.insert TermOrd.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 | 
| 110 | (c, (c0, Ss)) of ars represents the arity t::(Ss)c being derived | |
| 111 | via c0 <= c. "Coregularity" of the arities structure requires | |
| 112 | that for any two declarations t::(Ss1)c1 and t::(Ss2)c2 such that | |
| 113 | c1 <= c2 holds Ss1 <= Ss2. | |
| 2956 | 114 | *) | 
| 115 | ||
| 19645 | 116 | datatype algebra = Algebra of | 
| 20573 | 117 |  {classes: serial Graph.T,
 | 
| 19645 | 118 | arities: (class * (class * sort list)) list Symtab.table}; | 
| 119 | ||
| 120 | fun rep_algebra (Algebra args) = args; | |
| 121 | ||
| 122 | val classes_of = #classes o rep_algebra; | |
| 123 | val arities_of = #arities o rep_algebra; | |
| 124 | ||
| 125 | fun make_algebra (classes, arities) = | |
| 126 |   Algebra {classes = classes, arities = arities};
 | |
| 127 | ||
| 128 | fun map_classes f (Algebra {classes, arities}) = make_algebra (f classes, arities);
 | |
| 129 | fun map_arities f (Algebra {classes, arities}) = make_algebra (classes, f arities);
 | |
| 130 | ||
| 131 | ||
| 132 | (* classes *) | |
| 133 | ||
| 21933 
819ef284720b
classes: more direct way to achieve topological sorting;
 wenzelm parents: 
21926diff
changeset | 134 | 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 | 135 | |
| 19645 | 136 | val super_classes = Graph.imm_succs o classes_of; | 
| 2956 | 137 | |
| 138 | ||
| 19529 | 139 | (* class relations *) | 
| 2956 | 140 | |
| 19645 | 141 | val class_less = Graph.is_edge o classes_of; | 
| 142 | fun class_le algebra (c1, c2) = c1 = c2 orelse class_less algebra (c1, c2); | |
| 2956 | 143 | |
| 144 | ||
| 19529 | 145 | (* sort relations *) | 
| 2956 | 146 | |
| 19645 | 147 | fun sort_le algebra (S1, S2) = | 
| 23585 | 148 | S1 = S2 orelse forall (fn c2 => exists (fn c1 => class_le algebra (c1, c2)) S1) S2; | 
| 2956 | 149 | |
| 19645 | 150 | fun sorts_le algebra (Ss1, Ss2) = | 
| 151 | ListPair.all (sort_le algebra) (Ss1, Ss2); | |
| 2956 | 152 | |
| 19645 | 153 | fun sort_eq algebra (S1, S2) = | 
| 154 | sort_le algebra (S1, S2) andalso sort_le algebra (S2, S1); | |
| 2956 | 155 | |
| 156 | ||
| 19529 | 157 | (* intersection *) | 
| 2956 | 158 | |
| 19645 | 159 | fun inter_class algebra c S = | 
| 2956 | 160 | let | 
| 161 | fun intr [] = [c] | |
| 162 | | intr (S' as c' :: c's) = | |
| 19645 | 163 | if class_le algebra (c', c) then S' | 
| 164 | else if class_le algebra (c, c') then intr c's | |
| 2956 | 165 | else c' :: intr c's | 
| 166 | in intr S end; | |
| 167 | ||
| 19645 | 168 | fun inter_sort algebra (S1, S2) = | 
| 169 | sort_strings (fold (inter_class algebra) S1 S2); | |
| 2956 | 170 | |
| 171 | ||
| 24732 | 172 | (* normal forms *) | 
| 2956 | 173 | |
| 24732 | 174 | fun minimize_sort _ [] = [] | 
| 175 | | minimize_sort _ (S as [_]) = S | |
| 176 | | minimize_sort algebra S = | |
| 19645 | 177 | filter (fn c => not (exists (fn c' => class_less algebra (c', c)) S)) S | 
| 19529 | 178 | |> sort_distinct string_ord; | 
| 2990 | 179 | |
| 24732 | 180 | fun complete_sort algebra = | 
| 181 | Graph.all_succs (classes_of algebra) o minimize_sort algebra; | |
| 182 | ||
| 28623 | 183 | fun minimal_sorts algebra raw_sorts = | 
| 184 | let | |
| 185 | fun le S1 S2 = sort_le algebra (S1, S2); | |
| 186 | val sorts = make (map (minimize_sort algebra) raw_sorts); | |
| 187 | in sorts |> filter_out (fn S => exists (fn S' => le S' S andalso not (le S S')) sorts) end; | |
| 188 | ||
| 2990 | 189 | |
| 19645 | 190 | (* certify *) | 
| 191 | ||
| 192 | fun certify_class algebra c = | |
| 193 | if can (Graph.get_node (classes_of algebra)) c then c | |
| 194 |   else raise TYPE ("Undeclared class: " ^ quote c, [], []);
 | |
| 195 | ||
| 24732 | 196 | fun certify_sort classes = minimize_sort classes o map (certify_class classes); | 
| 19645 | 197 | |
| 198 | ||
| 2956 | 199 | |
| 19529 | 200 | (** build algebras **) | 
| 19514 | 201 | |
| 202 | (* classes *) | |
| 203 | ||
| 23655 
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
 wenzelm parents: 
23585diff
changeset | 204 | fun err_dup_class c = error ("Duplicate declaration of class: " ^ quote c);
 | 
| 19514 | 205 | |
| 206 | fun err_cyclic_classes pp css = | |
| 207 | error (cat_lines (map (fn cs => | |
| 208 | "Cycle in class relation: " ^ Pretty.string_of_classrel pp cs) css)); | |
| 209 | ||
| 19645 | 210 | fun add_class pp (c, cs) = map_classes (fn classes => | 
| 19514 | 211 | let | 
| 20573 | 212 | val classes' = classes |> Graph.new_node (c, serial ()) | 
| 23655 
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
 wenzelm parents: 
23585diff
changeset | 213 | handle Graph.DUP dup => err_dup_class dup; | 
| 19514 | 214 | val classes'' = classes' |> fold Graph.add_edge_trans_acyclic (map (pair c) cs) | 
| 215 | handle Graph.CYCLES css => err_cyclic_classes pp css; | |
| 19645 | 216 | in classes'' end); | 
| 19514 | 217 | |
| 218 | ||
| 219 | (* arities *) | |
| 220 | ||
| 221 | local | |
| 222 | ||
| 223 | fun for_classes _ NONE = "" | |
| 224 | | for_classes pp (SOME (c1, c2)) = | |
| 225 | " for classes " ^ Pretty.string_of_classrel pp [c1, c2]; | |
| 226 | ||
| 227 | fun err_conflict pp t cc (c, Ss) (c', Ss') = | |
| 228 |   error ("Conflict of type arities" ^ for_classes pp cc ^ ":\n  " ^
 | |
| 229 | Pretty.string_of_arity pp (t, Ss, [c]) ^ " and\n " ^ | |
| 230 | Pretty.string_of_arity pp (t, Ss', [c'])); | |
| 231 | ||
| 19645 | 232 | fun coregular pp algebra t (c, (c0, Ss)) ars = | 
| 19514 | 233 | let | 
| 19524 | 234 | fun conflict (c', (_, Ss')) = | 
| 19645 | 235 | if class_le algebra (c, c') andalso not (sorts_le algebra (Ss, Ss')) then | 
| 19514 | 236 | SOME ((c, c'), (c', Ss')) | 
| 19645 | 237 | else if class_le algebra (c', c) andalso not (sorts_le algebra (Ss', Ss)) then | 
| 19514 | 238 | SOME ((c', c), (c', Ss')) | 
| 239 | else NONE; | |
| 240 | in | |
| 241 | (case get_first conflict ars of | |
| 242 | SOME ((c1, c2), (c', Ss')) => err_conflict pp t (SOME (c1, c2)) (c, Ss) (c', Ss') | |
| 19524 | 243 | | NONE => (c, (c0, Ss)) :: ars) | 
| 19514 | 244 | end; | 
| 245 | ||
| 19645 | 246 | fun complete algebra (c0, Ss) = map (rpair (c0, Ss)) (c0 :: super_classes algebra c0); | 
| 247 | ||
| 248 | fun insert pp algebra t (c, (c0, Ss)) ars = | |
| 19514 | 249 | (case AList.lookup (op =) ars c of | 
| 19645 | 250 | NONE => coregular pp algebra t (c, (c0, Ss)) ars | 
| 19524 | 251 | | SOME (_, Ss') => | 
| 19645 | 252 | if sorts_le algebra (Ss, Ss') then ars | 
| 253 | else if sorts_le algebra (Ss', Ss) then | |
| 254 | coregular pp algebra t (c, (c0, Ss)) | |
| 19524 | 255 | (filter_out (fn (c'', (_, Ss'')) => c = c'' andalso Ss'' = Ss') ars) | 
| 19514 | 256 | else err_conflict pp t NONE (c, Ss) (c, Ss')); | 
| 257 | ||
| 19645 | 258 | fun insert_ars pp algebra (t, ars) arities = | 
| 259 | let val ars' = | |
| 260 | Symtab.lookup_list arities t | |
| 261 | |> fold_rev (fold_rev (insert pp algebra t)) (map (complete algebra) ars) | |
| 262 | in Symtab.update (t, ars') arities end; | |
| 19514 | 263 | |
| 264 | in | |
| 265 | ||
| 19645 | 266 | fun add_arities pp arg algebra = algebra |> map_arities (insert_ars pp algebra arg); | 
| 19514 | 267 | |
| 19645 | 268 | fun add_arities_table pp algebra = | 
| 269 | Symtab.fold (fn (t, ars) => insert_ars pp algebra (t, map snd ars)); | |
| 19514 | 270 | |
| 271 | end; | |
| 272 | ||
| 19529 | 273 | |
| 19645 | 274 | (* classrel *) | 
| 275 | ||
| 276 | fun rebuild_arities pp algebra = algebra |> map_arities (fn arities => | |
| 277 | Symtab.empty | |
| 278 | |> add_arities_table pp algebra arities); | |
| 279 | ||
| 280 | fun add_classrel pp rel = rebuild_arities pp o map_classes (fn classes => | |
| 281 | classes |> Graph.add_edge_trans_acyclic rel | |
| 282 | handle Graph.CYCLES css => err_cyclic_classes pp css); | |
| 283 | ||
| 284 | ||
| 285 | (* empty and merge *) | |
| 286 | ||
| 287 | val empty_algebra = make_algebra (Graph.empty, Symtab.empty); | |
| 288 | ||
| 289 | fun merge_algebra pp | |
| 290 |    (Algebra {classes = classes1, arities = arities1},
 | |
| 291 |     Algebra {classes = classes2, arities = arities2}) =
 | |
| 292 | let | |
| 293 | val classes' = Graph.merge_trans_acyclic (op =) (classes1, classes2) | |
| 23655 
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
 wenzelm parents: 
23585diff
changeset | 294 | handle Graph.DUP c => err_dup_class c | 
| 19645 | 295 | | Graph.CYCLES css => err_cyclic_classes pp css; | 
| 296 | val algebra0 = make_algebra (classes', Symtab.empty); | |
| 297 | val arities' = Symtab.empty | |
| 298 | |> add_arities_table pp algebra0 arities1 | |
| 299 | |> add_arities_table pp algebra0 arities2; | |
| 300 | in make_algebra (classes', arities') end; | |
| 301 | ||
| 21933 
819ef284720b
classes: more direct way to achieve topological sorting;
 wenzelm parents: 
21926diff
changeset | 302 | |
| 28922 | 303 | (* algebra projections *) | 
| 304 | ||
| 305 | fun classrels_of (Algebra {classes, ...}) =
 | |
| 306 | map (fn [c] => (c, Graph.imm_succs classes c)) (rev (Graph.strong_conn classes)); | |
| 307 | ||
| 308 | fun instances_of (Algebra {arities, ...}) =
 | |
| 309 | Symtab.fold (fn (a, cs) => append (map (pair a o fst) cs)) arities []; | |
| 21933 
819ef284720b
classes: more direct way to achieve topological sorting;
 wenzelm parents: 
21926diff
changeset | 310 | |
| 22181 
39104d1c43ca
added explicit query function for arities to subalgebra projection
 haftmann parents: 
21933diff
changeset | 311 | fun subalgebra pp P sargs (algebra as Algebra {classes, arities}) =
 | 
| 19952 | 312 | let | 
| 24732 | 313 | val restrict_sort = minimize_sort algebra o filter P o Graph.all_succs classes; | 
| 22181 
39104d1c43ca
added explicit query function for arities to subalgebra projection
 haftmann parents: 
21933diff
changeset | 314 | fun restrict_arity tyco (c, (_, Ss)) = | 
| 
39104d1c43ca
added explicit query function for arities to subalgebra projection
 haftmann parents: 
21933diff
changeset | 315 | if P c then | 
| 
39104d1c43ca
added explicit query function for arities to subalgebra projection
 haftmann parents: 
21933diff
changeset | 316 | SOME (c, (c, Ss |> map2 (curry (inter_sort algebra)) (sargs (c, tyco)) | 
| 
39104d1c43ca
added explicit query function for arities to subalgebra projection
 haftmann parents: 
21933diff
changeset | 317 | |> map restrict_sort)) | 
| 
39104d1c43ca
added explicit query function for arities to subalgebra projection
 haftmann parents: 
21933diff
changeset | 318 | else NONE; | 
| 21933 
819ef284720b
classes: more direct way to achieve topological sorting;
 wenzelm parents: 
21926diff
changeset | 319 | val classes' = classes |> Graph.subgraph P; | 
| 22181 
39104d1c43ca
added explicit query function for arities to subalgebra projection
 haftmann parents: 
21933diff
changeset | 320 | val arities' = arities |> Symtab.map' (map_filter o restrict_arity); | 
| 21933 
819ef284720b
classes: more direct way to achieve topological sorting;
 wenzelm parents: 
21926diff
changeset | 321 | in (restrict_sort, rebuild_arities pp (make_algebra (classes', arities'))) end; | 
| 20465 | 322 | |
| 19645 | 323 | |
| 19529 | 324 | |
| 325 | (** sorts of types **) | |
| 326 | ||
| 26639 | 327 | (* errors -- delayed message composition *) | 
| 19578 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 wenzelm parents: 
19531diff
changeset | 328 | |
| 26639 | 329 | datatype class_error = | 
| 330 | NoClassrel of class * class | | |
| 331 | NoArity of string * class | | |
| 332 | NoSubsort of sort * sort; | |
| 19529 | 333 | |
| 26639 | 334 | fun class_error pp (NoClassrel (c1, c2)) = | 
| 22196 | 335 | "No class relation " ^ Pretty.string_of_classrel pp [c1, c2] | 
| 26639 | 336 | | class_error pp (NoArity (a, c)) = | 
| 26326 | 337 | "No type arity " ^ Pretty.string_of_arity pp (a, [], [c]) | 
| 26994 | 338 | | class_error pp (NoSubsort (S1, S2)) = | 
| 339 | "Cannot derive subsort relation " ^ Pretty.string_of_sort pp S1 | |
| 340 | ^ " < " ^ Pretty.string_of_sort pp S2; | |
| 19529 | 341 | |
| 19578 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 wenzelm parents: 
19531diff
changeset | 342 | exception CLASS_ERROR of class_error; | 
| 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 wenzelm parents: 
19531diff
changeset | 343 | |
| 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 wenzelm parents: 
19531diff
changeset | 344 | |
| 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 wenzelm parents: 
19531diff
changeset | 345 | (* mg_domain *) | 
| 19529 | 346 | |
| 19645 | 347 | fun mg_domain algebra a S = | 
| 19529 | 348 | let | 
| 19645 | 349 | val arities = arities_of algebra; | 
| 19529 | 350 | fun dom c = | 
| 351 | (case AList.lookup (op =) (Symtab.lookup_list arities a) c of | |
| 19578 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 wenzelm parents: 
19531diff
changeset | 352 | NONE => raise CLASS_ERROR (NoArity (a, c)) | 
| 19529 | 353 | | SOME (_, Ss) => Ss); | 
| 19645 | 354 | fun dom_inter c Ss = ListPair.map (inter_sort algebra) (dom c, Ss); | 
| 19529 | 355 | in | 
| 356 | (case S of | |
| 357 | [] => raise Fail "Unknown domain of empty intersection" | |
| 358 | | c :: cs => fold dom_inter cs (dom c)) | |
| 359 | end; | |
| 360 | ||
| 361 | ||
| 26639 | 362 | (* meet_sort *) | 
| 363 | ||
| 364 | fun meet_sort algebra = | |
| 365 | let | |
| 366 | fun inters S S' = inter_sort algebra (S, S'); | |
| 367 | fun meet _ [] = I | |
| 368 | | meet (TFree (_, S)) S' = | |
| 369 | if sort_le algebra (S, S') then I | |
| 370 | else raise CLASS_ERROR (NoSubsort (S, S')) | |
| 371 | | meet (TVar (v, S)) S' = | |
| 372 | if sort_le algebra (S, S') then I | |
| 373 | else Vartab.map_default (v, S) (inters S') | |
| 374 | | meet (Type (a, Ts)) S = fold2 meet Ts (mg_domain algebra a S); | |
| 375 | in uncurry meet end; | |
| 376 | ||
| 28665 | 377 | fun meet_sort_typ algebra (T, S) = | 
| 378 | let | |
| 379 | val tab = meet_sort algebra (T, S) Vartab.empty; | |
| 380 | in Term.map_type_tvar (fn (v, _) => | |
| 381 | TVar (v, (the o Vartab.lookup tab) v)) | |
| 382 | end; | |
| 383 | ||
| 26639 | 384 | |
| 19529 | 385 | (* of_sort *) | 
| 386 | ||
| 19645 | 387 | fun of_sort algebra = | 
| 19529 | 388 | let | 
| 389 | fun ofS (_, []) = true | |
| 19645 | 390 | | ofS (TFree (_, S), S') = sort_le algebra (S, S') | 
| 391 | | ofS (TVar (_, S), S') = sort_le algebra (S, S') | |
| 19529 | 392 | | ofS (Type (a, Ts), S) = | 
| 19645 | 393 | let val Ss = mg_domain algebra a S in | 
| 19529 | 394 | ListPair.all ofS (Ts, Ss) | 
| 19578 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 wenzelm parents: 
19531diff
changeset | 395 | end handle CLASS_ERROR _ => false; | 
| 19529 | 396 | in ofS end; | 
| 397 | ||
| 398 | ||
| 27498 | 399 | (* animating derivations *) | 
| 400 | ||
| 27555 | 401 | fun weaken algebra class_relation = | 
| 402 | let | |
| 403 | fun path (x, c1 :: c2 :: cs) = path (class_relation (x, c1) c2, c2 :: cs) | |
| 404 | | path (x, _) = x; | |
| 405 | in fn (x, c1) => fn c2 => | |
| 406 | (case Graph.irreducible_paths (classes_of algebra) (c1, c2) of | |
| 407 | [] => raise CLASS_ERROR (NoClassrel (c1, c2)) | |
| 408 | | cs :: _ => path (x, cs)) | |
| 409 | end; | |
| 19529 | 410 | |
| 22570 
f166a5416b3f
renamed of_sort_derivation record fields (avoid clash with Alice keywords);
 wenzelm parents: 
22364diff
changeset | 411 | fun of_sort_derivation pp algebra {class_relation, type_constructor, type_variable} =
 | 
| 19529 | 412 | let | 
| 27555 | 413 | val weaken = weaken algebra class_relation; | 
| 414 | val arities = arities_of algebra; | |
| 19578 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 wenzelm parents: 
19531diff
changeset | 415 | |
| 19529 | 416 | fun weakens S1 S2 = S2 |> map (fn c2 => | 
| 19645 | 417 | (case S1 |> find_first (fn (_, c1) => class_le algebra (c1, c2)) of | 
| 19529 | 418 | SOME d1 => weaken d1 c2 | 
| 26994 | 419 | | NONE => raise CLASS_ERROR (NoSubsort (map #2 S1, S2)))); | 
| 19529 | 420 | |
| 421 | fun derive _ [] = [] | |
| 422 | | derive (Type (a, Ts)) S = | |
| 423 | let | |
| 19645 | 424 | val Ss = mg_domain algebra a S; | 
| 19529 | 425 | val dom = map2 (fn T => fn S => derive T S ~~ S) Ts Ss; | 
| 426 | in | |
| 427 | S |> map (fn c => | |
| 428 | let | |
| 429 | val (c0, Ss') = the (AList.lookup (op =) (Symtab.lookup_list arities a) c); | |
| 430 | val dom' = map2 (fn d => fn S' => weakens d S' ~~ S') dom Ss'; | |
| 22570 
f166a5416b3f
renamed of_sort_derivation record fields (avoid clash with Alice keywords);
 wenzelm parents: 
22364diff
changeset | 431 | in weaken (type_constructor a dom' c0, c0) c end) | 
| 19529 | 432 | end | 
| 22570 
f166a5416b3f
renamed of_sort_derivation record fields (avoid clash with Alice keywords);
 wenzelm parents: 
22364diff
changeset | 433 | | derive T S = weakens (type_variable T) S; | 
| 19529 | 434 | in uncurry derive end; | 
| 435 | ||
| 436 | ||
| 437 | (* witness_sorts *) | |
| 438 | ||
| 19645 | 439 | fun witness_sorts algebra types hyps sorts = | 
| 19529 | 440 | let | 
| 19645 | 441 | fun le S1 S2 = sort_le algebra (S1, S2); | 
| 19529 | 442 | fun get_solved S2 (T, S1) = if le S1 S2 then SOME (T, S2) else NONE; | 
| 443 |     fun get_hyp S2 S1 = if le S1 S2 then SOME (TFree ("'hyp", S1), 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: 
19531diff
changeset | 446 | fun witn_sort _ [] solved_failed = (SOME (propT, []), solved_failed) | 
| 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 wenzelm parents: 
19531diff
changeset | 447 | | witn_sort path S (solved, failed) = | 
| 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 wenzelm parents: 
19531diff
changeset | 448 | if exists (le S) failed then (NONE, (solved, failed)) | 
| 19529 | 449 | else | 
| 450 | (case get_first (get_solved S) solved of | |
| 19578 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 wenzelm parents: 
19531diff
changeset | 451 | SOME w => (SOME w, (solved, failed)) | 
| 19529 | 452 | | NONE => | 
| 453 | (case get_first (get_hyp S) hyps of | |
| 19578 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 wenzelm parents: 
19531diff
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: 
19531diff
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: 
19531diff
changeset | 459 | and witn_types _ [] S (solved, failed) = (NONE, (solved, S :: failed)) | 
| 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 wenzelm parents: 
19531diff
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: 
19531diff
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: 
19531diff
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: 
19531diff
changeset | 470 | in (SOME w, (w :: solved', failed')) end | 
| 
f93b7637a5e6
added class_error and exception CLASS_ERROR (supercedes DOMAIN);
 wenzelm parents: 
19531diff
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: 
19531diff
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; |