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