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