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