18168
|
1 |
(* Title: Pure/Tools/class_package.ML
|
|
2 |
ID: $Id$
|
|
3 |
Author: Florian Haftmann, TU Muenchen
|
|
4 |
|
18755
|
5 |
Type classes, simulated by locales.
|
18168
|
6 |
*)
|
|
7 |
|
|
8 |
signature CLASS_PACKAGE =
|
|
9 |
sig
|
18755
|
10 |
val add_class: bstring -> class list -> Element.context list -> theory
|
18515
|
11 |
-> ProofContext.context * theory
|
18755
|
12 |
val add_class_i: bstring -> class list -> Element.context_i list -> theory
|
18515
|
13 |
-> ProofContext.context * theory
|
18575
|
14 |
val add_instance_arity: (xstring * string list) * string
|
|
15 |
-> ((bstring * string) * Attrib.src list) list
|
|
16 |
-> theory -> Proof.state
|
|
17 |
val add_instance_arity_i: (string * sort list) * sort
|
18728
|
18 |
-> ((bstring * term) * attribute list) list
|
18575
|
19 |
-> theory -> Proof.state
|
18515
|
20 |
val add_classentry: class -> xstring list -> xstring list -> theory -> theory
|
18168
|
21 |
|
18755
|
22 |
val intern_class: theory -> xstring -> string
|
|
23 |
val extern_class: theory -> string -> xstring
|
|
24 |
val certify_class: theory -> class -> class
|
18360
|
25 |
val syntactic_sort_of: theory -> sort -> sort
|
18702
|
26 |
val the_superclasses: theory -> class -> class list
|
|
27 |
val the_consts_sign: theory -> class -> string * (string * typ) list
|
18168
|
28 |
val lookup_const_class: theory -> string -> class option
|
18702
|
29 |
val the_instances: theory -> class -> (string * string) list
|
|
30 |
val the_inst_sign: theory -> class * string -> (string * sort) list * (string * typ) list
|
18168
|
31 |
val get_classtab: theory -> (string list * (string * string) list) Symtab.table
|
18702
|
32 |
val print_classes: theory -> unit
|
18168
|
33 |
|
|
34 |
type sortcontext = (string * sort) list
|
|
35 |
datatype sortlookup = Instance of (class * string) * sortlookup list list
|
|
36 |
| Lookup of class list * (string * int)
|
|
37 |
val extract_sortctxt: theory -> typ -> sortcontext
|
18702
|
38 |
val extract_sortlookup: theory -> string * typ -> sortlookup list list
|
18168
|
39 |
end;
|
|
40 |
|
|
41 |
structure ClassPackage: CLASS_PACKAGE =
|
|
42 |
struct
|
|
43 |
|
|
44 |
|
18708
|
45 |
(* theory data *)
|
18168
|
46 |
|
|
47 |
type class_data = {
|
18575
|
48 |
superclasses: class list,
|
|
49 |
name_locale: string,
|
|
50 |
name_axclass: string,
|
|
51 |
var: string,
|
|
52 |
consts: (string * typ) list,
|
|
53 |
insts: (string * string) list
|
18168
|
54 |
};
|
|
55 |
|
18575
|
56 |
structure ClassData = TheoryDataFun (
|
18168
|
57 |
struct
|
|
58 |
val name = "Pure/classes";
|
|
59 |
type T = class_data Symtab.table * class Symtab.table;
|
|
60 |
val empty = (Symtab.empty, Symtab.empty);
|
|
61 |
val copy = I;
|
|
62 |
val extend = I;
|
|
63 |
fun merge _ ((t1, r1), (t2, r2))=
|
|
64 |
(Symtab.merge (op =) (t1, t2),
|
|
65 |
Symtab.merge (op =) (r1, r2));
|
18575
|
66 |
fun print thy (tab, _) =
|
|
67 |
let
|
|
68 |
fun pretty_class (name, {superclasses, name_locale, name_axclass, var, consts, insts}) =
|
|
69 |
(Pretty.block o Pretty.fbreaks) [
|
|
70 |
Pretty.str ("class " ^ name ^ ":"),
|
|
71 |
(Pretty.block o Pretty.fbreaks) (
|
|
72 |
Pretty.str "superclasses: "
|
|
73 |
:: map Pretty.str superclasses
|
|
74 |
),
|
|
75 |
Pretty.str ("locale: " ^ name_locale),
|
|
76 |
Pretty.str ("axclass: " ^ name_axclass),
|
|
77 |
Pretty.str ("class variable: " ^ var),
|
|
78 |
(Pretty.block o Pretty.fbreaks) (
|
|
79 |
Pretty.str "constants: "
|
|
80 |
:: map (fn (c, ty) => Pretty.str (c ^ " :: " ^ Sign.string_of_typ thy ty)) consts
|
|
81 |
),
|
|
82 |
(Pretty.block o Pretty.fbreaks) (
|
|
83 |
Pretty.str "instances: "
|
|
84 |
:: map (fn (tyco, thyname) => Pretty.str (tyco ^ ", in theory " ^ thyname)) insts
|
|
85 |
)
|
|
86 |
]
|
|
87 |
in
|
|
88 |
(Pretty.writeln o Pretty.chunks o map pretty_class o Symtab.dest) tab
|
|
89 |
end;
|
18168
|
90 |
end
|
|
91 |
);
|
|
92 |
|
18708
|
93 |
val _ = Context.add_setup ClassData.init;
|
18575
|
94 |
val print_classes = ClassData.print;
|
|
95 |
|
|
96 |
val lookup_class_data = Symtab.lookup o fst o ClassData.get;
|
|
97 |
val lookup_const_class = Symtab.lookup o snd o ClassData.get;
|
18168
|
98 |
|
|
99 |
fun get_class_data thy class =
|
|
100 |
case lookup_class_data thy class
|
18755
|
101 |
of NONE => error ("undeclared operational class " ^ quote class)
|
18168
|
102 |
| SOME data => data;
|
|
103 |
|
18575
|
104 |
fun add_class_data (class, (superclasses, name_locale, name_axclass, classvar, consts)) =
|
|
105 |
ClassData.map (fn (classtab, consttab) => (
|
|
106 |
classtab
|
|
107 |
|> Symtab.update (class, {
|
|
108 |
superclasses = superclasses,
|
|
109 |
name_locale = name_locale,
|
|
110 |
name_axclass = name_axclass,
|
|
111 |
var = classvar,
|
|
112 |
consts = consts,
|
|
113 |
insts = []
|
|
114 |
}),
|
|
115 |
consttab
|
|
116 |
|> fold (fn (c, _) => Symtab.update (c, class)) consts
|
|
117 |
));
|
|
118 |
|
|
119 |
fun add_inst_data (class, inst) =
|
|
120 |
(ClassData.map o apfst o Symtab.map_entry class)
|
|
121 |
(fn {superclasses, name_locale, name_axclass, var, consts, insts}
|
|
122 |
=> {
|
|
123 |
superclasses = superclasses,
|
|
124 |
name_locale = name_locale,
|
|
125 |
name_axclass = name_axclass,
|
|
126 |
var = var,
|
|
127 |
consts = consts,
|
|
128 |
insts = insts @ [inst]
|
|
129 |
});
|
|
130 |
|
18168
|
131 |
|
18755
|
132 |
(* name handling *)
|
|
133 |
|
|
134 |
fun certify_class thy class =
|
|
135 |
(get_class_data thy class; class);
|
|
136 |
|
|
137 |
fun intern_class thy raw_class =
|
|
138 |
certify_class thy (Sign.intern_class thy raw_class);
|
|
139 |
|
|
140 |
fun extern_class thy class =
|
|
141 |
Sign.extern_class thy (certify_class thy class);
|
|
142 |
|
|
143 |
|
18575
|
144 |
(* classes and instances *)
|
18168
|
145 |
|
18702
|
146 |
fun subst_clsvar v ty_subst =
|
|
147 |
map_type_tfree (fn u as (w, _) =>
|
|
148 |
if w = v then ty_subst else TFree u);
|
|
149 |
|
18168
|
150 |
local
|
|
151 |
|
18755
|
152 |
fun gen_add_class add_locale prep_class bname raw_supclasses raw_body thy =
|
18168
|
153 |
let
|
18755
|
154 |
val supclasses = map (prep_class thy) raw_supclasses;
|
|
155 |
val supsort = case map (#name_axclass o get_class_data thy) supclasses
|
|
156 |
of [] => Sign.defaultS thy
|
|
157 |
| sort => Sorts.certify_sort (Sign.classes_of thy) sort;
|
|
158 |
val locexpr = case supclasses
|
|
159 |
of [] => Locale.empty
|
|
160 |
| _ => (Locale.Merge o map (Locale.Locale o #name_locale o get_class_data thy))
|
|
161 |
supclasses;
|
18575
|
162 |
fun extract_assumes c_adds elems =
|
|
163 |
let
|
|
164 |
fun subst_free ts =
|
|
165 |
let
|
|
166 |
val get_ty = the o AList.lookup (op =) (fold Term.add_frees ts []);
|
|
167 |
val subst_map = map (fn (c, (c', _)) =>
|
|
168 |
(Free (c, get_ty c), Const (c', get_ty c))) c_adds;
|
|
169 |
in map (subst_atomic subst_map) ts end;
|
|
170 |
in
|
|
171 |
elems
|
|
172 |
|> (map o List.mapPartial)
|
18755
|
173 |
(fn (Element.Assumes asms) => (SOME o map (map fst o snd)) asms
|
18575
|
174 |
| _ => NONE)
|
|
175 |
|> Library.flat o Library.flat o Library.flat
|
|
176 |
|> subst_free
|
|
177 |
end;
|
18515
|
178 |
fun extract_tyvar_name thy tys =
|
|
179 |
fold (curry add_typ_tfrees) tys []
|
18575
|
180 |
|> (fn [(v, sort)] =>
|
18755
|
181 |
if Sorts.sort_le (Sign.classes_of thy) (swap (sort, supsort))
|
18515
|
182 |
then v
|
18575
|
183 |
else error ("illegal sort constraint on class type variable: " ^ Sign.string_of_sort thy sort)
|
18515
|
184 |
| [] => error ("no class type variable")
|
|
185 |
| vs => error ("more than one type variable: " ^ (commas o map (Sign.string_of_typ thy o TFree)) vs))
|
18755
|
186 |
fun extract_tyvar_consts thy (import_elems, body_elems) =
|
|
187 |
let
|
|
188 |
fun get_elems elems =
|
|
189 |
elems
|
|
190 |
|> Library.flat
|
|
191 |
|> List.mapPartial
|
|
192 |
(fn (Element.Fixes consts) => SOME consts
|
|
193 |
| _ => NONE)
|
|
194 |
|> Library.flat
|
|
195 |
|> map (fn (c, ty, syn) =>
|
|
196 |
((c, the ty), (Syntax.unlocalize_mixfix o Syntax.fix_mixfix c) syn))
|
|
197 |
val import_consts = get_elems import_elems;
|
|
198 |
val body_consts = get_elems body_elems;
|
|
199 |
val v = extract_tyvar_name thy (map (snd o fst) (import_consts @ body_consts));
|
|
200 |
fun subst_ty cs =
|
|
201 |
map ((apfst o apsnd) (subst_clsvar v (TFree (v, [])))) cs;
|
|
202 |
val import_cs = subst_ty import_consts;
|
|
203 |
val body_cs = subst_ty body_consts;
|
|
204 |
in (v, (import_cs, body_cs)) end;
|
18575
|
205 |
fun add_global_const v ((c, ty), syn) thy =
|
18515
|
206 |
thy
|
18575
|
207 |
|> Sign.add_consts_i [(c, ty |> subst_clsvar v (TFree (v, Sign.defaultS thy)), syn)]
|
|
208 |
|> `(fn thy => (c, (Sign.intern_const thy c, ty)))
|
|
209 |
fun add_global_constraint v class (_, (c, ty)) thy =
|
18515
|
210 |
thy
|
18575
|
211 |
|> Sign.add_const_constraint_i (c, subst_clsvar v (TVar ((v, 0), [class])) ty);
|
18755
|
212 |
fun interpret name_locale cs ax_axioms thy =
|
|
213 |
thy
|
|
214 |
|> Locale.interpretation (NameSpace.base name_locale, [])
|
|
215 |
(Locale.Locale name_locale)
|
|
216 |
(map (fn ((c, _), _) => SOME (Sign.intern_const thy c)) cs)
|
|
217 |
|> Proof.global_terminal_proof (Method.Basic (Method.fact ax_axioms), NONE)
|
|
218 |
|> swap;
|
18515
|
219 |
fun print_ctxt ctxt elem =
|
|
220 |
map Pretty.writeln (Element.pretty_ctxt ctxt elem)
|
18168
|
221 |
in
|
|
222 |
thy
|
18755
|
223 |
|> add_locale bname locexpr raw_body
|
18575
|
224 |
|-> (fn ((import_elems, body_elems), ctxt) =>
|
|
225 |
`(fn thy => Locale.intern thy bname)
|
18515
|
226 |
#-> (fn name_locale =>
|
18755
|
227 |
`(fn thy => extract_tyvar_consts thy (import_elems, body_elems))
|
|
228 |
#-> (fn (v, (c_global, c_defs)) =>
|
18575
|
229 |
fold_map (add_global_const v) c_defs
|
|
230 |
#-> (fn c_adds =>
|
18755
|
231 |
AxClass.add_axclass_i (bname, supsort)
|
18575
|
232 |
(map (Thm.no_attributes o pair "") (extract_assumes c_adds (import_elems @ body_elems)))
|
18755
|
233 |
#-> (fn { axioms = ax_axioms : thm list, ...} =>
|
18575
|
234 |
`(fn thy => Sign.intern_class thy bname)
|
18515
|
235 |
#-> (fn name_axclass =>
|
18575
|
236 |
fold (add_global_constraint v name_axclass) c_adds
|
18755
|
237 |
#> add_class_data (name_locale, (supclasses, name_locale, name_axclass, v, map snd c_adds))
|
18575
|
238 |
#> tap (fn _ => (map o map) (print_ctxt ctxt) import_elems)
|
|
239 |
#> tap (fn _ => (map o map) (print_ctxt ctxt) body_elems)
|
18755
|
240 |
#> interpret name_locale (c_global @ c_defs) ax_axioms
|
18575
|
241 |
))))))
|
18168
|
242 |
end;
|
|
243 |
|
|
244 |
in
|
|
245 |
|
18755
|
246 |
val add_class = gen_add_class (Locale.add_locale_context true) intern_class;
|
|
247 |
val add_class_i = gen_add_class (Locale.add_locale_context_i true) certify_class;
|
18168
|
248 |
|
|
249 |
end; (* local *)
|
|
250 |
|
18575
|
251 |
fun gen_instance_arity prep_arity add_defs tap_def raw_arity raw_defs thy =
|
|
252 |
let
|
|
253 |
val dest_def = Theory.dest_def (Sign.pp thy) handle TERM (msg, _) => error msg;
|
|
254 |
val arity as (tyco, asorts, sort) = prep_arity thy ((fn ((x, y), z) => (x, y, z)) raw_arity);
|
|
255 |
val ty_inst = Type (tyco, map2 (curry TVar o rpair 0) (Term.invent_names [] "'a" (length asorts)) asorts)
|
|
256 |
fun get_c_req class =
|
|
257 |
let
|
|
258 |
val data = get_class_data thy class;
|
|
259 |
val subst_ty = map_type_tfree (fn (var as (v, _)) =>
|
|
260 |
if #var data = v then ty_inst else TFree var)
|
|
261 |
in (map (apsnd subst_ty) o #consts) data end;
|
|
262 |
val c_req = (Library.flat o map get_c_req) sort;
|
|
263 |
fun get_remove_contraint c thy =
|
|
264 |
let
|
|
265 |
val ty1 = Sign.the_const_constraint thy c;
|
|
266 |
val ty2 = Sign.the_const_type thy c;
|
|
267 |
in
|
|
268 |
thy
|
|
269 |
|> Sign.add_const_constraint_i (c, ty2)
|
|
270 |
|> pair (c, ty1)
|
|
271 |
end;
|
|
272 |
fun get_c_given thy = map (fst o dest_def o snd o tap_def thy o fst) raw_defs;
|
|
273 |
fun check_defs c_given c_req thy =
|
|
274 |
let
|
18702
|
275 |
fun eq_c ((c1, ty1), (c2, ty2)) = c1 = c2
|
|
276 |
andalso Sign.typ_instance thy (ty1, ty2)
|
18755
|
277 |
andalso Sign.typ_instance thy (ty2, ty1);
|
|
278 |
val _ = case fold (remove eq_c) c_req c_given
|
|
279 |
of [] => ()
|
|
280 |
| cs => error ("superfluous definition(s) given for "
|
|
281 |
^ (commas o map (fn (c, ty) => quote (c ^ "::" ^ Sign.string_of_typ thy ty))) cs);
|
18575
|
282 |
val _ = case fold (remove eq_c) c_given c_req
|
|
283 |
of [] => ()
|
18755
|
284 |
| cs => error ("no definition(s) given for "
|
18575
|
285 |
^ (commas o map (fn (c, ty) => quote (c ^ "::" ^ Sign.string_of_typ thy ty))) cs);
|
|
286 |
in thy end;
|
18755
|
287 |
fun after_qed cs =
|
|
288 |
fold Sign.add_const_constraint_i cs
|
|
289 |
#> fold (fn class => add_inst_data (class, (tyco, Context.theory_name thy))) sort;
|
18575
|
290 |
in
|
|
291 |
thy
|
|
292 |
|> fold_map get_remove_contraint (map fst c_req)
|
18755
|
293 |
||> tap (fn thy => check_defs (get_c_given thy) c_req thy)
|
18575
|
294 |
||> add_defs (true, raw_defs)
|
18755
|
295 |
|-> (fn cs => AxClass.instance_arity_i (after_qed cs) arity)
|
18575
|
296 |
end;
|
|
297 |
|
|
298 |
val add_instance_arity = fn x => gen_instance_arity (AxClass.read_arity) IsarThy.add_defs read_axm x;
|
|
299 |
val add_instance_arity_i = fn x => gen_instance_arity (AxClass.cert_arity) IsarThy.add_defs_i (K I) x;
|
|
300 |
|
18168
|
301 |
|
18702
|
302 |
(* queries *)
|
18168
|
303 |
|
18702
|
304 |
fun is_class thy cls =
|
|
305 |
lookup_class_data thy cls
|
|
306 |
|> Option.map (not o null o #consts)
|
|
307 |
|> the_default false;
|
18168
|
308 |
|
18360
|
309 |
fun syntactic_sort_of thy sort =
|
|
310 |
let
|
|
311 |
val classes = Sign.classes_of thy;
|
|
312 |
fun get_sort cls =
|
|
313 |
if is_class thy cls
|
|
314 |
then [cls]
|
|
315 |
else syntactic_sort_of thy (Sorts.superclasses classes cls);
|
|
316 |
in
|
|
317 |
map get_sort sort
|
|
318 |
|> Library.flat
|
|
319 |
|> Sorts.norm_sort classes
|
|
320 |
end;
|
18168
|
321 |
|
18702
|
322 |
fun the_superclasses thy class =
|
18515
|
323 |
if is_class thy class
|
|
324 |
then
|
|
325 |
Sorts.superclasses (Sign.classes_of thy) class
|
|
326 |
|> syntactic_sort_of thy
|
|
327 |
else
|
|
328 |
error ("no syntactic class: " ^ class);
|
18168
|
329 |
|
18702
|
330 |
fun the_consts_sign thy class =
|
18168
|
331 |
let
|
18702
|
332 |
val data = (the oo Symtab.lookup) ((fst o ClassData.get) thy) class
|
|
333 |
in (#var data, #consts data) end;
|
|
334 |
|
|
335 |
fun lookup_const_class thy =
|
|
336 |
Symtab.lookup ((snd o ClassData.get) thy);
|
|
337 |
|
|
338 |
fun the_instances thy class =
|
|
339 |
(#insts o the o Symtab.lookup ((fst o ClassData.get) thy)) class;
|
18168
|
340 |
|
18702
|
341 |
fun the_inst_sign thy (class, tyco) =
|
18168
|
342 |
let
|
18702
|
343 |
val _ = if is_class thy class then () else error ("no syntactic class: " ^ class);
|
|
344 |
val arity =
|
|
345 |
Sorts.mg_domain (Sign.classes_arities_of thy) tyco [class]
|
|
346 |
|> map (syntactic_sort_of thy);
|
|
347 |
val clsvar = (#var o the o Symtab.lookup ((fst o ClassData.get) thy)) class;
|
|
348 |
val const_sign = (snd o the_consts_sign thy) class;
|
|
349 |
fun add_var sort used =
|
|
350 |
let
|
|
351 |
val v = hd (Term.invent_names used "'a" 1)
|
|
352 |
in ((v, sort), v::used) end;
|
|
353 |
val (vsorts, _) =
|
|
354 |
[]
|
|
355 |
|> fold (fn (_, ty) => curry (gen_union (op =))
|
|
356 |
((map (fst o fst) o typ_tvars) ty @ (map fst o typ_tfrees) ty)) const_sign
|
|
357 |
|> fold_map add_var arity;
|
|
358 |
val ty_inst = Type (tyco, map (fn (v, sort) => TVar ((v, 0), sort)) vsorts);
|
|
359 |
val inst_signs = map (apsnd (subst_clsvar clsvar ty_inst)) const_sign;
|
|
360 |
in (vsorts, inst_signs) end;
|
18168
|
361 |
|
|
362 |
fun get_classtab thy =
|
|
363 |
Symtab.fold
|
18575
|
364 |
(fn (class, { consts = consts, insts = insts, ... }) =>
|
|
365 |
Symtab.update_new (class, (map fst consts, insts)))
|
18702
|
366 |
((fst o ClassData.get) thy) Symtab.empty;
|
18168
|
367 |
|
|
368 |
|
|
369 |
(* extracting dictionary obligations from types *)
|
|
370 |
|
|
371 |
type sortcontext = (string * sort) list;
|
|
372 |
|
18335
|
373 |
fun extract_sortctxt thy ty =
|
18702
|
374 |
(typ_tfrees o fst o Type.freeze_thaw_type) ty
|
18360
|
375 |
|> map (apsnd (syntactic_sort_of thy))
|
18168
|
376 |
|> filter (not o null o snd);
|
|
377 |
|
|
378 |
datatype sortlookup = Instance of (class * string) * sortlookup list list
|
|
379 |
| Lookup of class list * (string * int)
|
|
380 |
|
18702
|
381 |
fun extract_sortlookup thy (c, raw_typ_use) =
|
18168
|
382 |
let
|
18702
|
383 |
val raw_typ_def = Sign.the_const_constraint thy c;
|
18168
|
384 |
val typ_def = Type.varifyT raw_typ_def;
|
|
385 |
val typ_use = Type.varifyT raw_typ_use;
|
|
386 |
val match_tab = Sign.typ_match thy (typ_def, typ_use) Vartab.empty;
|
|
387 |
fun tab_lookup vname = (the o Vartab.lookup match_tab) (vname, 0);
|
|
388 |
fun get_superclass_derivation (subclasses, superclass) =
|
|
389 |
(the oo get_first) (fn subclass =>
|
|
390 |
Sorts.class_le_path (Sign.classes_of thy) (subclass, superclass)
|
|
391 |
) subclasses;
|
|
392 |
fun mk_class_deriv thy subclasses superclass =
|
|
393 |
case get_superclass_derivation (subclasses, superclass)
|
18515
|
394 |
of (subclass::deriv) =>
|
|
395 |
((rev o filter (is_class thy)) deriv, find_index_eq subclass subclasses);
|
18168
|
396 |
fun mk_lookup (sort_def, (Type (tycon, tys))) =
|
|
397 |
let
|
18330
|
398 |
val arity_lookup = map2 (curry mk_lookup)
|
18360
|
399 |
(map (syntactic_sort_of thy) (Sorts.mg_domain (Sign.classes_arities_of thy) tycon sort_def)) tys
|
18168
|
400 |
in map (fn class => Instance ((class, tycon), arity_lookup)) sort_def end
|
|
401 |
| mk_lookup (sort_def, TVar ((vname, _), sort_use)) =
|
|
402 |
let
|
|
403 |
fun mk_look class =
|
18515
|
404 |
let val (deriv, classindex) = mk_class_deriv thy (syntactic_sort_of thy sort_use) class
|
18168
|
405 |
in Lookup (deriv, (vname, classindex)) end;
|
|
406 |
in map mk_look sort_def end;
|
18702
|
407 |
fun reorder_sortctxt ctxt =
|
|
408 |
case lookup_const_class thy c
|
|
409 |
of NONE => ctxt
|
|
410 |
| SOME class =>
|
|
411 |
let
|
|
412 |
val data = (the o Symtab.lookup ((fst o ClassData.get) thy)) class;
|
|
413 |
val sign = (Type.varifyT o the o AList.lookup (op =) (#consts data)) c;
|
|
414 |
val match_tab = Sign.typ_match thy (sign, typ_def) Vartab.empty;
|
|
415 |
val v : string = case Vartab.lookup match_tab (#var data, 0)
|
|
416 |
of SOME (_, TVar ((v, _), _)) => v;
|
|
417 |
in
|
|
418 |
(v, (the o AList.lookup (op =) ctxt) v) :: AList.delete (op =) v ctxt
|
|
419 |
end;
|
18168
|
420 |
in
|
18335
|
421 |
extract_sortctxt thy ((fst o Type.freeze_thaw_type) raw_typ_def)
|
18702
|
422 |
|> reorder_sortctxt
|
18168
|
423 |
|> map (tab_lookup o fst)
|
18360
|
424 |
|> map (apfst (syntactic_sort_of thy))
|
18168
|
425 |
|> filter (not o null o fst)
|
|
426 |
|> map mk_lookup
|
|
427 |
end;
|
|
428 |
|
|
429 |
|
18335
|
430 |
(* intermediate auxiliary *)
|
18168
|
431 |
|
18575
|
432 |
fun add_classentry raw_class raw_cs raw_insts thy =
|
18168
|
433 |
let
|
|
434 |
val class = Sign.intern_class thy raw_class;
|
18702
|
435 |
val cs_proto =
|
|
436 |
raw_cs
|
|
437 |
|> map (Sign.intern_const thy)
|
|
438 |
|> map (fn c => (c, Sign.the_const_constraint thy c));
|
|
439 |
val used =
|
|
440 |
[]
|
|
441 |
|> fold (fn (_, ty) => curry (gen_union (op =))
|
|
442 |
((map (fst o fst) o typ_tvars) ty @ (map fst o typ_tfrees) ty)) cs_proto
|
|
443 |
val v = hd (Term.invent_names used "'a" 1);
|
|
444 |
val cs =
|
|
445 |
cs_proto
|
|
446 |
|> map (fn (c, ty) => (c, map_type_tvar (fn var as ((tvar', _), sort) =>
|
|
447 |
if Sorts.sort_eq (Sign.classes_of thy) ([class], sort)
|
|
448 |
then TFree (v, [])
|
|
449 |
else TVar var
|
|
450 |
) ty));
|
18575
|
451 |
val insts = map (rpair (Context.theory_name thy) o Sign.intern_type thy) raw_insts;
|
18168
|
452 |
in
|
|
453 |
thy
|
18702
|
454 |
|> add_class_data (class, ([], "", class, v, cs))
|
18575
|
455 |
|> fold (curry add_inst_data class) insts
|
18335
|
456 |
end;
|
18515
|
457 |
|
|
458 |
|
|
459 |
(* toplevel interface *)
|
|
460 |
|
|
461 |
local
|
|
462 |
|
|
463 |
structure P = OuterParse
|
|
464 |
and K = OuterKeyword
|
|
465 |
|
|
466 |
in
|
|
467 |
|
18575
|
468 |
val (classK, instanceK) = ("class", "class_instance")
|
18515
|
469 |
|
|
470 |
val classP =
|
|
471 |
OuterSyntax.command classK "operational type classes" K.thy_decl
|
18755
|
472 |
(P.name --| P.$$$ "="
|
|
473 |
-- Scan.optional (Scan.repeat1 (P.name --| P.$$$ "+")) []
|
|
474 |
-- Scan.optional (P.!!! (Scan.repeat1 P.context_element)) []
|
18515
|
475 |
>> (Toplevel.theory_context
|
18755
|
476 |
o (fn f => swap o f) o (fn ((bname, supclasses), elems) => add_class bname supclasses elems)));
|
18515
|
477 |
|
18575
|
478 |
val instanceP =
|
18755
|
479 |
OuterSyntax.command instanceK "declare instance for operational type class" K.thy_goal
|
18575
|
480 |
(P.xname -- (P.$$$ "::" |-- P.!!! P.arity)
|
|
481 |
-- Scan.repeat1 P.spec_name
|
18755
|
482 |
>> (Toplevel.print oo Toplevel.theory_to_proof
|
18575
|
483 |
o (fn ((tyco, (asorts, sort)), defs) => add_instance_arity ((tyco, asorts), sort) defs)));
|
|
484 |
|
|
485 |
val _ = OuterSyntax.add_parsers [classP, instanceP];
|
18515
|
486 |
|
|
487 |
end; (* local *)
|
|
488 |
|
18168
|
489 |
end; (* struct *)
|