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