author | haftmann |
Fri, 30 Jun 2006 12:03:36 +0200 | |
changeset 19966 | 88bbe97ed0b0 |
parent 19957 | 91ba241a1678 |
child 20076 | def4ad161528 |
permissions | -rw-r--r-- |
18168 | 1 |
(* Title: Pure/Tools/class_package.ML |
2 |
ID: $Id$ |
|
3 |
Author: Florian Haftmann, TU Muenchen |
|
4 |
||
19468 | 5 |
Type classes derived from primitive axclasses and locales. |
18168 | 6 |
*) |
7 |
||
8 |
signature CLASS_PACKAGE = |
|
9 |
sig |
|
19150 | 10 |
val class: bstring -> class list -> Element.context list -> theory |
18515 | 11 |
-> ProofContext.context * theory |
19150 | 12 |
val class_i: bstring -> class list -> Element.context_i list -> theory |
18515 | 13 |
-> ProofContext.context * theory |
19150 | 14 |
val instance_arity: (xstring * string list) * string |
19136 | 15 |
-> bstring * Attrib.src list -> ((bstring * Attrib.src list) * string) list |
18575 | 16 |
-> theory -> Proof.state |
19150 | 17 |
val instance_arity_i: (string * sort list) * sort |
19136 | 18 |
-> bstring * attribute list -> ((bstring * attribute list) * term) list |
18575 | 19 |
-> theory -> Proof.state |
19150 | 20 |
val prove_instance_arity: tactic -> (string * sort list) * sort |
19136 | 21 |
-> bstring * attribute list -> ((bstring * attribute list) * term) list |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
22 |
-> theory -> theory |
19150 | 23 |
val instance_sort: string * string -> theory -> Proof.state |
24 |
val instance_sort_i: class * sort -> theory -> Proof.state |
|
25 |
val prove_instance_sort: tactic -> class * sort -> theory -> theory |
|
18168 | 26 |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
27 |
val intern_class: theory -> xstring -> class |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
28 |
val intern_sort: theory -> sort -> sort |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
29 |
val extern_class: theory -> class -> xstring |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
30 |
val extern_sort: theory -> sort -> sort |
18755 | 31 |
val certify_class: theory -> class -> class |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
32 |
val certify_sort: theory -> sort -> sort |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
33 |
val read_sort: theory -> string -> sort |
18884 | 34 |
val operational_sort_of: theory -> sort -> sort |
18702 | 35 |
val the_superclasses: theory -> class -> class list |
36 |
val the_consts_sign: theory -> class -> string * (string * typ) list |
|
37 |
val the_inst_sign: theory -> class * string -> (string * sort) list * (string * typ) list |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
38 |
|
18702 | 39 |
val print_classes: theory -> unit |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
40 |
val intro_classes_tac: thm list -> tactic |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
41 |
val default_intro_classes_tac: thm list -> tactic |
18168 | 42 |
|
43 |
type sortcontext = (string * sort) list |
|
18884 | 44 |
datatype classlookup = Instance of (class * string) * classlookup list list |
19253 | 45 |
| Lookup of class list * (string * (int * int)) |
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19929
diff
changeset
|
46 |
val sortcontext_of_typ: theory -> typ -> sortcontext |
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19929
diff
changeset
|
47 |
val sortlookup: theory -> sort * typ -> classlookup list |
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19929
diff
changeset
|
48 |
val sortlookups_const: theory -> string * typ -> classlookup list list |
19966 | 49 |
|
50 |
val use_instance2: bool ref; |
|
51 |
val the_propnames: theory -> class -> string list |
|
18168 | 52 |
end; |
53 |
||
54 |
structure ClassPackage: CLASS_PACKAGE = |
|
55 |
struct |
|
56 |
||
57 |
||
19957 | 58 |
(* auxiliary *) |
59 |
||
60 |
fun instantiations_of thy (ty, ty') = |
|
61 |
let |
|
62 |
val vartab = typ_tvars ty; |
|
63 |
fun prep_vartab (v, (_, ty)) = |
|
64 |
case (the o AList.lookup (op =) vartab) v |
|
65 |
of [] => NONE |
|
66 |
| sort => SOME ((v, sort), ty); |
|
67 |
in case try (Sign.typ_match thy (ty, ty')) Vartab.empty |
|
68 |
of NONE => NONE |
|
69 |
| SOME vartab => |
|
70 |
SOME ((map_filter prep_vartab o Vartab.dest) vartab) |
|
71 |
end; |
|
72 |
||
73 |
||
18708 | 74 |
(* theory data *) |
18168 | 75 |
|
19456 | 76 |
datatype class_data = ClassData of { |
18575 | 77 |
name_locale: string, |
78 |
name_axclass: string, |
|
79 |
var: string, |
|
19966 | 80 |
consts: (string * (string * typ)) list, |
19957 | 81 |
(*locale parameter ~> toplevel constant*) |
19966 | 82 |
propnames: string list |
83 |
} * thm list Symtab.table; |
|
18168 | 84 |
|
19456 | 85 |
fun rep_classdata (ClassData c) = c; |
86 |
||
18575 | 87 |
structure ClassData = TheoryDataFun ( |
18168 | 88 |
struct |
89 |
val name = "Pure/classes"; |
|
19966 | 90 |
type T = class_data Symtab.table; |
91 |
val empty = Symtab.empty; |
|
18168 | 92 |
val copy = I; |
93 |
val extend = I; |
|
19966 | 94 |
fun merge _ = Symtab.join (fn _ => fn (ClassData (classd, instd1), ClassData (_, instd2)) => |
95 |
(ClassData (classd, Symtab.merge (K true) (instd1, instd2)))); |
|
96 |
fun print thy data = |
|
18575 | 97 |
let |
19966 | 98 |
fun pretty_class (name, ClassData ({name_locale, name_axclass, var, consts, ...}, _)) = |
18575 | 99 |
(Pretty.block o Pretty.fbreaks) [ |
100 |
Pretty.str ("class " ^ name ^ ":"), |
|
101 |
Pretty.str ("locale: " ^ name_locale), |
|
102 |
Pretty.str ("axclass: " ^ name_axclass), |
|
103 |
Pretty.str ("class variable: " ^ var), |
|
104 |
(Pretty.block o Pretty.fbreaks) ( |
|
105 |
Pretty.str "constants: " |
|
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
106 |
:: map (fn (_, (c, ty)) => Pretty.str (c ^ " :: " ^ Sign.string_of_typ thy ty)) consts |
18575 | 107 |
) |
108 |
] |
|
109 |
in |
|
19966 | 110 |
(Pretty.writeln o Pretty.chunks o map pretty_class o Symtab.dest) data |
18575 | 111 |
end; |
18168 | 112 |
end |
113 |
); |
|
114 |
||
18708 | 115 |
val _ = Context.add_setup ClassData.init; |
18575 | 116 |
val print_classes = ClassData.print; |
117 |
||
19038 | 118 |
|
119 |
(* queries *) |
|
120 |
||
19966 | 121 |
val lookup_class_data = Option.map rep_classdata oo Symtab.lookup o ClassData.get; |
18168 | 122 |
|
19038 | 123 |
fun the_class_data thy class = |
18168 | 124 |
case lookup_class_data thy class |
18755 | 125 |
of NONE => error ("undeclared operational class " ^ quote class) |
18168 | 126 |
| SOME data => data; |
127 |
||
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
128 |
val is_class = is_some oo lookup_class_data; |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
129 |
|
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
130 |
fun is_operational_class thy cls = |
18702 | 131 |
lookup_class_data thy cls |
19966 | 132 |
|> Option.map (not o null o #consts o fst) |
18702 | 133 |
|> the_default false; |
18168 | 134 |
|
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19929
diff
changeset
|
135 |
fun operational_sort_of thy = |
18360 | 136 |
let |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
137 |
fun get_sort class = |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
138 |
if is_operational_class thy class |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
139 |
then [class] |
19412 | 140 |
else operational_sort_of thy (Sign.super_classes thy class); |
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19929
diff
changeset
|
141 |
in Sign.certify_sort thy o maps get_sort end; |
18168 | 142 |
|
18702 | 143 |
fun the_superclasses thy class = |
18515 | 144 |
if is_class thy class |
145 |
then |
|
19412 | 146 |
Sign.super_classes thy class |
18884 | 147 |
|> operational_sort_of thy |
18515 | 148 |
else |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
149 |
error ("no class: " ^ class); |
18168 | 150 |
|
19038 | 151 |
fun the_ancestry thy classes = |
152 |
let |
|
153 |
fun ancestry class anc = |
|
154 |
anc |
|
155 |
|> cons class |
|
156 |
|> fold ancestry (the_superclasses thy class); |
|
157 |
in fold ancestry classes [] end; |
|
158 |
||
159 |
fun subst_clsvar v ty_subst = |
|
160 |
map_type_tfree (fn u as (w, _) => |
|
161 |
if w = v then ty_subst else TFree u); |
|
162 |
||
19966 | 163 |
val the_parm_map = #consts o fst oo the_class_data; |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
164 |
|
18702 | 165 |
fun the_consts_sign thy class = |
18168 | 166 |
let |
19966 | 167 |
val data = (fst o the_class_data thy) class |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
168 |
in (#var data, (map snd o #consts) data) end; |
18702 | 169 |
|
170 |
fun the_inst_sign thy (class, tyco) = |
|
18168 | 171 |
let |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
172 |
val _ = if is_operational_class thy class then () else error ("no operational class: " ^ class); |
19957 | 173 |
val asorts = Sign.arity_sorts thy tyco [class]; |
19966 | 174 |
val (clsvar, const_sign) = the_consts_sign thy class; |
18702 | 175 |
fun add_var sort used = |
176 |
let |
|
177 |
val v = hd (Term.invent_names used "'a" 1) |
|
178 |
in ((v, sort), v::used) end; |
|
179 |
val (vsorts, _) = |
|
19213 | 180 |
[clsvar] |
18702 | 181 |
|> fold (fn (_, ty) => curry (gen_union (op =)) |
182 |
((map (fst o fst) o typ_tvars) ty @ (map fst o typ_tfrees) ty)) const_sign |
|
19957 | 183 |
|> fold_map add_var asorts; |
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19929
diff
changeset
|
184 |
val ty_inst = Type (tyco, map TFree vsorts); |
18702 | 185 |
val inst_signs = map (apsnd (subst_clsvar clsvar ty_inst)) const_sign; |
186 |
in (vsorts, inst_signs) end; |
|
18168 | 187 |
|
19966 | 188 |
val the_propnames = #propnames o fst oo the_class_data; |
189 |
||
18168 | 190 |
|
19038 | 191 |
(* updaters *) |
192 |
||
19966 | 193 |
fun add_class_data (class, (name_locale, name_axclass, var, consts, propnames)) = |
19957 | 194 |
ClassData.map ( |
19966 | 195 |
Symtab.update_new (class, ClassData ({ |
19957 | 196 |
name_locale = name_locale, |
197 |
name_axclass = name_axclass, |
|
198 |
var = var, |
|
19966 | 199 |
consts = consts, |
200 |
propnames = propnames}, Symtab.empty)) |
|
19957 | 201 |
); |
19038 | 202 |
|
19966 | 203 |
fun add_inst_def ((class, tyco), thm) = |
204 |
ClassData.map ( |
|
205 |
Symtab.map_entry class (fn ClassData (classd, instd) => |
|
206 |
ClassData (classd, Symtab.insert_list eq_thm (tyco, thm) instd)) |
|
207 |
); |
|
19038 | 208 |
|
209 |
(* name handling *) |
|
210 |
||
211 |
fun certify_class thy class = |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
212 |
(fn class => (the_class_data thy class; class)) (Sign.certify_class thy class); |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
213 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
214 |
fun certify_sort thy sort = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
215 |
map (fn class => (the_class_data thy class; class)) (Sign.certify_sort thy sort); |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
216 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
217 |
fun intern_class thy = |
19957 | 218 |
certify_class thy o Sign.intern_class thy; |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
219 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
220 |
fun intern_sort thy = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
221 |
certify_sort thy o Sign.intern_sort thy; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
222 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
223 |
fun extern_class thy = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
224 |
Sign.extern_class thy o certify_class thy; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
225 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
226 |
fun extern_sort thy = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
227 |
Sign.extern_sort thy o certify_sort thy; |
19038 | 228 |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
229 |
fun read_sort thy = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
230 |
certify_sort thy o Sign.read_sort thy; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
231 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
232 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
233 |
(* tactics and methods *) |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
234 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
235 |
fun intro_classes_tac facts st = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
236 |
(ALLGOALS (Method.insert_tac facts THEN' |
19928 | 237 |
REPEAT_ALL_NEW (resolve_tac (AxClass.class_intros (Thm.theory_of_thm st)))) |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
238 |
THEN Tactic.distinct_subgoals_tac) st; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
239 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
240 |
fun default_intro_classes_tac [] = intro_classes_tac [] |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
241 |
| default_intro_classes_tac _ = Tactical.no_tac; (*no error message!*) |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
242 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
243 |
fun default_tac rules ctxt facts = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
244 |
HEADGOAL (Method.some_rule_tac rules ctxt facts) ORELSE |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
245 |
default_intro_classes_tac facts; |
19038 | 246 |
|
19468 | 247 |
val _ = Context.add_setup (Method.add_methods |
248 |
[("intro_classes", Method.no_args (Method.METHOD intro_classes_tac), |
|
249 |
"back-chain introduction rules of classes"), |
|
250 |
("default", Method.thms_ctxt_args (Method.METHOD oo default_tac), |
|
251 |
"apply some intro/elim rule")]); |
|
252 |
||
19038 | 253 |
|
19246 | 254 |
(* axclass instances *) |
255 |
||
256 |
local |
|
257 |
||
19412 | 258 |
fun gen_instance mk_prop add_thm after_qed inst thy = |
19246 | 259 |
thy |
260 |
|> ProofContext.init |
|
19412 | 261 |
|> Proof.theorem_i PureThy.internalK NONE (after_qed oo (fold o fold) add_thm) NONE ("", []) |
19585 | 262 |
(map (fn t => (("", []), [(t, [])])) (mk_prop thy inst)); |
19246 | 263 |
|
264 |
in |
|
265 |
||
19966 | 266 |
val axclass_instance_sort = |
19412 | 267 |
gen_instance (single oo (Logic.mk_classrel oo AxClass.read_classrel)) AxClass.add_classrel I; |
19246 | 268 |
val axclass_instance_arity = |
19412 | 269 |
gen_instance (Logic.mk_arities oo Sign.read_arity) AxClass.add_arity; |
19246 | 270 |
val axclass_instance_arity_i = |
19412 | 271 |
gen_instance (Logic.mk_arities oo Sign.cert_arity) AxClass.add_arity; |
19246 | 272 |
|
273 |
end; |
|
274 |
||
275 |
||
19038 | 276 |
(* classes and instances *) |
277 |
||
278 |
local |
|
279 |
||
19957 | 280 |
fun add_axclass_i (name, supsort) params axs thy = |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
281 |
let |
19395 | 282 |
val (c, thy') = thy |
19957 | 283 |
|> AxClass.define_class_i (name, supsort) params axs; |
19518 | 284 |
val {intro, axioms, ...} = AxClass.get_definition thy' c; |
19395 | 285 |
in ((c, (intro, axioms)), thy') end; |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
286 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
287 |
fun prove_interpretation_i (prfx, atts) expr insts tac thy = |
19038 | 288 |
let |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
289 |
fun ad_hoc_term NONE = NONE |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
290 |
| ad_hoc_term (SOME (Const (c, ty))) = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
291 |
let |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
292 |
val p = setmp show_types true (setmp show_sorts true (setmp print_mode [] (Sign.pretty_typ thy))) ty; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
293 |
val s = c ^ "::" ^ Pretty.output p; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
294 |
in SOME s end |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
295 |
| ad_hoc_term (SOME t) = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
296 |
let |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
297 |
val p = setmp show_types true (setmp show_sorts true (setmp print_mode [] (Sign.pretty_term thy))) t; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
298 |
val s = Pretty.output p; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
299 |
in SOME s end; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
300 |
in |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
301 |
thy |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
302 |
|> Locale.interpretation (prfx, atts) expr (map ad_hoc_term insts) |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
303 |
|> Proof.global_terminal_proof (Method.Basic (fn _ => Method.SIMPLE_METHOD tac), NONE) |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
304 |
|-> (fn _ => I) |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
305 |
end; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
306 |
|
19150 | 307 |
fun gen_class add_locale prep_class bname raw_supclasses raw_elems thy = |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
308 |
let |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
309 |
val supclasses = map (prep_class thy) raw_supclasses; |
19038 | 310 |
val supsort = |
311 |
supclasses |
|
19966 | 312 |
|> map (#name_axclass o fst o the_class_data thy) |
19468 | 313 |
|> Sign.certify_sort thy |
19038 | 314 |
|> null ? K (Sign.defaultS thy); |
19966 | 315 |
val expr = (Locale.Merge o map (Locale.Locale o #name_locale o fst o the_class_data thy)) supclasses; |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
316 |
val mapp_sup = AList.make |
19468 | 317 |
(the o AList.lookup (op =) ((flat o map (the_parm_map thy) o the_ancestry thy) supclasses)) |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
318 |
((map (fst o fst) o Locale.parameters_of_expr thy) expr); |
19038 | 319 |
fun extract_tyvar_consts thy name_locale = |
320 |
let |
|
321 |
fun extract_tyvar_name thy tys = |
|
322 |
fold (curry add_typ_tfrees) tys [] |
|
323 |
|> (fn [(v, sort)] => |
|
19468 | 324 |
if Sign.subsort thy (supsort, sort) |
19038 | 325 |
then v |
326 |
else error ("illegal sort constraint on class type variable: " ^ Sign.string_of_sort thy sort) |
|
327 |
| [] => error ("no class type variable") |
|
328 |
| vs => error ("more than one type variable: " ^ (commas o map (Sign.string_of_typ thy o TFree)) vs)) |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
329 |
val consts1 = |
19038 | 330 |
Locale.parameters_of thy name_locale |
331 |
|> map (apsnd Syntax.unlocalize_mixfix) |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
332 |
val v = (extract_tyvar_name thy o map (snd o fst)) consts1; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
333 |
val consts2 = map ((apfst o apsnd) (subst_clsvar v (TFree (v, [])))) consts1; |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
334 |
in (v, chop (length mapp_sup) consts2) end; |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
335 |
fun add_consts v raw_cs_sup raw_cs_this thy = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
336 |
let |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
337 |
fun add_global_const ((c, ty), syn) thy = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
338 |
thy |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
339 |
|> Sign.add_consts_i [(c, ty |> subst_clsvar v (TFree (v, Sign.defaultS thy)), syn)] |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
340 |
|> `(fn thy => (c, (Sign.intern_const thy c, ty))) |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
341 |
in |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
342 |
thy |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
343 |
|> fold_map add_global_const raw_cs_this |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
344 |
end; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
345 |
fun extract_assumes thy name_locale cs_mapp = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
346 |
let |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
347 |
val subst_assume = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
348 |
map_aterms (fn Free (c, ty) => Const ((fst o the o AList.lookup (op =) cs_mapp) c, ty) |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
349 |
| t => t) |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
350 |
fun prep_asm ((name, atts), ts) = |
19957 | 351 |
((NameSpace.base name, map (Attrib.attribute thy) atts), map subst_assume ts) |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
352 |
in |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
353 |
(map prep_asm o Locale.local_asms_of thy) name_locale |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
354 |
end; |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
355 |
fun add_global_constraint v class (_, (c, ty)) thy = |
19038 | 356 |
thy |
19136 | 357 |
|> Sign.add_const_constraint_i (c, SOME (subst_clsvar v (TFree (v, [class])) ty)); |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
358 |
fun mk_const thy class v (c, ty) = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
359 |
Const (c, subst_clsvar v (TFree (v, [class])) ty); |
19038 | 360 |
in |
361 |
thy |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
362 |
|> add_locale bname expr raw_elems |
19928 | 363 |
|-> (fn (name_locale, ctxt) => |
19038 | 364 |
`(fn thy => extract_tyvar_consts thy name_locale) |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
365 |
#-> (fn (v, (raw_cs_sup, raw_cs_this)) => |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
366 |
add_consts v raw_cs_sup raw_cs_this |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
367 |
#-> (fn mapp_this => |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
368 |
`(fn thy => extract_assumes thy name_locale (mapp_sup @ mapp_this)) |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
369 |
#-> (fn loc_axioms => |
19957 | 370 |
add_axclass_i (bname, supsort) (map (fst o snd) mapp_this) loc_axioms |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
371 |
#-> (fn (name_axclass, (_, ax_axioms)) => |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
372 |
fold (add_global_constraint v name_axclass) mapp_this |
19966 | 373 |
#> add_class_data (name_locale, (name_locale, name_axclass, v, mapp_this, |
374 |
map (fst o fst) loc_axioms)) |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
375 |
#> prove_interpretation_i (NameSpace.base name_locale, []) |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
376 |
(Locale.Locale name_locale) (map (SOME o mk_const thy name_axclass v) (map snd (mapp_sup @ mapp_this))) |
19929 | 377 |
((ALLGOALS o ProofContext.fact_tac) ax_axioms) |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
378 |
#> pair ctxt |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
379 |
))))) |
19038 | 380 |
end; |
381 |
||
382 |
in |
|
383 |
||
19928 | 384 |
val class = gen_class (Locale.add_locale false) intern_class; |
385 |
val class_i = gen_class (Locale.add_locale_i false) certify_class; |
|
19038 | 386 |
|
387 |
end; (* local *) |
|
388 |
||
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
389 |
local |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
390 |
|
19957 | 391 |
fun gen_read_def thy prep_att read_def tyco ((raw_name, raw_atts), raw_t) = |
19038 | 392 |
let |
19957 | 393 |
val (_, t) = read_def thy (raw_name, raw_t); |
394 |
val ((c, ty), _) = Sign.cert_def (Sign.pp thy) t; |
|
395 |
val atts = map (prep_att thy) raw_atts; |
|
396 |
val name = case raw_name |
|
397 |
of "" => Thm.def_name (NameSpace.base c ^ "_" ^ NameSpace.base tyco) |
|
398 |
| _ => raw_name; |
|
399 |
in (c, (Logic.varifyT ty, ((name, t), atts))) end; |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
400 |
|
19957 | 401 |
fun read_def thy = gen_read_def thy Attrib.attribute read_axm; |
402 |
fun read_def_i thy = gen_read_def thy (K I) (K I); |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
403 |
|
19957 | 404 |
fun gen_instance_arity prep_arity prep_att read_def do_proof raw_arity (raw_name, raw_atts) raw_defs theory = |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
405 |
let |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
406 |
val pp = Sign.pp theory; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
407 |
val arity as (tyco, asorts, sort) = prep_arity theory ((fn ((x, y), z) => (x, y, z)) raw_arity); |
19038 | 408 |
val ty_inst = Type (tyco, map2 (curry TVar o rpair 0) (Term.invent_names [] "'a" (length asorts)) asorts) |
19136 | 409 |
val name = case raw_name |
410 |
of "" => Thm.def_name ((space_implode "_" o map NameSpace.base) sort ^ "_" ^ NameSpace.base tyco) |
|
411 |
| _ => raw_name; |
|
412 |
val atts = map (prep_att theory) raw_atts; |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
413 |
fun get_consts class = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
414 |
let |
19966 | 415 |
val data = (fst o the_class_data theory) class; |
19957 | 416 |
fun defined c = |
417 |
is_some (find_first (fn (_, { lhs = [ty], ...}) => |
|
418 |
Sign.typ_instance theory (ty, ty_inst) orelse Sign.typ_instance theory (ty_inst, ty)) |
|
419 |
(Defs.specifications_of (Theory.defs_of theory) c)) |
|
420 |
val subst_ty = map_type_tfree (fn (v, sort) => |
|
421 |
if #var data = v then ty_inst else TVar ((v, 0), sort)); |
|
422 |
in |
|
423 |
(map_filter (fn (_, (c, ty)) => |
|
19966 | 424 |
if defined c then NONE else SOME ((c, (class, subst_ty ty)))) o #consts) data |
19957 | 425 |
end; |
426 |
val cs = (maps get_consts o the_ancestry theory) sort; |
|
427 |
fun read_defs defs cs = |
|
428 |
let |
|
429 |
val thy_read = (Sign.primitive_arity (tyco, asorts, sort) o Theory.copy) theory; |
|
430 |
fun read raw_def cs = |
|
431 |
let |
|
432 |
val (c, (ty, def)) = read_def thy_read tyco raw_def; |
|
19966 | 433 |
val (class, ty') = case AList.lookup (op =) cs c |
19957 | 434 |
of NONE => error ("superfluous definition for constant " ^ quote c) |
19966 | 435 |
| SOME class_ty => class_ty; |
436 |
val def' = case instantiations_of thy_read (ty, ty') |
|
437 |
of NONE => error ("superfluous definition for constant " ^ |
|
438 |
quote c ^ "::" ^ Sign.string_of_typ thy_read ty) |
|
439 |
| SOME insttab => |
|
440 |
(apfst o apsnd o map_term_types) |
|
441 |
(Logic.unvarifyT o Term.instantiateT insttab o Logic.varifyT) def |
|
442 |
in ((class, def'), AList.delete (op =) c cs) end; |
|
19957 | 443 |
in fold_map read defs cs end; |
444 |
val (defs, _) = read_defs raw_defs cs; |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
445 |
fun get_remove_contraint c thy = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
446 |
let |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
447 |
val ty = Sign.the_const_constraint thy c; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
448 |
in |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
449 |
thy |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
450 |
|> Sign.add_const_constraint_i (c, NONE) |
19806 | 451 |
|> pair (c, Logic.legacy_unvarifyT ty) |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
452 |
end; |
19966 | 453 |
fun add_defs defs thy = |
454 |
thy |
|
455 |
|> PureThy.add_defs_i true (map snd defs) |
|
456 |
|-> (fn thms => pair (map fst defs ~~ thms)); |
|
457 |
fun register_def (class, thm) thy = |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
458 |
thy |
19966 | 459 |
|> add_inst_def ((class, tyco), thm); |
460 |
fun note_all thy = |
|
461 |
let |
|
462 |
val thms = maps (fn class => Symtab.lookup_list |
|
463 |
((snd o the_class_data thy) class) tyco) (the_ancestry thy sort); |
|
464 |
in |
|
465 |
thy |
|
466 |
|> PureThy.note_thmss_i PureThy.internalK [((name, atts), [(thms, [])])] |
|
467 |
|> snd |
|
468 |
end; |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
469 |
fun after_qed cs thy = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
470 |
thy |
19412 | 471 |
|> fold Sign.add_const_constraint_i (map (apsnd SOME) cs); |
19038 | 472 |
in |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
473 |
theory |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
474 |
|> fold_map get_remove_contraint (map fst cs) |
19966 | 475 |
||>> add_defs defs |
476 |
|-> (fn (cs, def_thms) => |
|
477 |
fold register_def def_thms |
|
478 |
#> note_all |
|
479 |
#> do_proof (after_qed cs) arity) |
|
19038 | 480 |
end; |
481 |
||
19957 | 482 |
fun instance_arity' do_proof = gen_instance_arity Sign.read_arity Attrib.attribute |
483 |
read_def do_proof; |
|
484 |
fun instance_arity_i' do_proof = gen_instance_arity Sign.cert_arity (K I) |
|
485 |
read_def_i do_proof; |
|
19246 | 486 |
fun tactic_proof tac after_qed arity = AxClass.prove_arity arity tac #> after_qed; |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
487 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
488 |
in |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
489 |
|
19957 | 490 |
val instance_arity = instance_arity' axclass_instance_arity_i; |
491 |
val instance_arity_i = instance_arity_i' axclass_instance_arity_i; |
|
19150 | 492 |
val prove_instance_arity = instance_arity_i' o tactic_proof; |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
493 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
494 |
end; (* local *) |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
495 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
496 |
local |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
497 |
|
19966 | 498 |
fun add_interpretation_in (after_qed : theory -> theory) (name, expr) thy = |
19136 | 499 |
thy |
500 |
|> Locale.interpretation_in_locale (name, expr); |
|
501 |
||
19966 | 502 |
fun prove_interpretation_in tac (after_qed : theory -> theory) (name, expr) thy = |
19136 | 503 |
thy |
504 |
|> Locale.interpretation_in_locale (name, expr) |
|
505 |
|> Proof.global_terminal_proof (Method.Basic (fn _ => Method.SIMPLE_METHOD tac), NONE) |
|
19966 | 506 |
|> snd |
507 |
|> after_qed; |
|
19038 | 508 |
|
19150 | 509 |
fun gen_instance_sort prep_class prep_sort do_proof (raw_class, raw_sort) theory = |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
510 |
let |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
511 |
val class = prep_class theory raw_class; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
512 |
val sort = prep_sort theory raw_sort; |
19966 | 513 |
val loc_name = (#name_locale o fst o the_class_data theory) class; |
19468 | 514 |
val loc_expr = |
19966 | 515 |
(Locale.Merge o map (Locale.Locale o #name_locale o fst o the_class_data theory)) sort; |
516 |
val const_names = (map (NameSpace.base o fst o snd) |
|
517 |
o maps (#consts o fst o the_class_data theory) |
|
518 |
o the_ancestry theory) [class]; |
|
519 |
val prop_tab = AList.make (the_propnames theory) |
|
520 |
(the_ancestry theory sort); |
|
521 |
fun mk_thm_names (superclass, prop_names) = |
|
522 |
let |
|
523 |
val thm_name_base = NameSpace.append "local" (space_implode "_" const_names); |
|
524 |
val export_name = class ^ "_" ^ superclass; |
|
525 |
in (export_name, map (Name o NameSpace.append thm_name_base) prop_names) end; |
|
526 |
val notes_tab_proto = map mk_thm_names prop_tab; |
|
527 |
fun test_note thy thmref = |
|
528 |
can (Locale.note_thmss PureThy.corollaryK loc_name |
|
529 |
[(("", []), [(thmref, [])])]) (Theory.copy thy); |
|
530 |
val notes_tab = map_filter (fn (export_name, thm_names) => case filter (test_note theory) thm_names |
|
531 |
of [] => NONE |
|
532 |
| thm_names' => SOME (export_name, thm_names')) notes_tab_proto; |
|
533 |
val _ = writeln ("fishing for "); |
|
534 |
val _ = print notes_tab; |
|
535 |
fun after_qed thy = thy; |
|
536 |
fun after_qed''' thy = |
|
537 |
fold (fn supclass => |
|
19412 | 538 |
AxClass.prove_classrel (class, supclass) |
19136 | 539 |
(ALLGOALS (K (intro_classes_tac [])) THEN |
19966 | 540 |
(ALLGOALS o resolve_tac o flat) []) |
541 |
) sort thy; |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
542 |
in |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
543 |
theory |
19136 | 544 |
|> do_proof after_qed (loc_name, loc_expr) |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
545 |
end; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
546 |
|
19150 | 547 |
fun instance_sort' do_proof = gen_instance_sort intern_class read_sort do_proof; |
548 |
fun instance_sort_i' do_proof = gen_instance_sort certify_class certify_sort do_proof; |
|
19136 | 549 |
val setup_proof = add_interpretation_in; |
550 |
val tactic_proof = prove_interpretation_in; |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
551 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
552 |
in |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
553 |
|
19150 | 554 |
val instance_sort = instance_sort' setup_proof; |
555 |
val instance_sort_i = instance_sort_i' setup_proof; |
|
556 |
val prove_instance_sort = instance_sort_i' o tactic_proof; |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
557 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
558 |
end; (* local *) |
19038 | 559 |
|
19957 | 560 |
|
18168 | 561 |
(* extracting dictionary obligations from types *) |
562 |
||
563 |
type sortcontext = (string * sort) list; |
|
564 |
||
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19929
diff
changeset
|
565 |
fun sortcontext_of_typ thy ty = |
18702 | 566 |
(typ_tfrees o fst o Type.freeze_thaw_type) ty |
18884 | 567 |
|> map (apsnd (operational_sort_of thy)) |
18168 | 568 |
|> filter (not o null o snd); |
569 |
||
18884 | 570 |
datatype classlookup = Instance of (class * string) * classlookup list list |
19253 | 571 |
| Lookup of class list * (string * (int * int)) |
19213 | 572 |
|
573 |
fun pretty_lookup' (Instance ((class, tyco), lss)) = |
|
574 |
(Pretty.block o Pretty.breaks) ( |
|
575 |
Pretty.enum "," "{" "}" [Pretty.str class, Pretty.str tyco] |
|
576 |
:: map pretty_lookup lss |
|
577 |
) |
|
19253 | 578 |
| pretty_lookup' (Lookup (classes, (v, (i, j)))) = |
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19929
diff
changeset
|
579 |
Pretty.enum " <" "[" "]" (map Pretty.str classes @ |
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19929
diff
changeset
|
580 |
[Pretty.str (v ^ "!" ^ string_of_int i ^ "/" ^ string_of_int j)]) |
19213 | 581 |
and pretty_lookup ls = (Pretty.enum "," "(" ")" o map pretty_lookup') ls; |
18168 | 582 |
|
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19929
diff
changeset
|
583 |
fun sortlookup thy (sort_decl, typ_ctxt) = |
18168 | 584 |
let |
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19929
diff
changeset
|
585 |
val pp = Sign.pp thy; |
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19929
diff
changeset
|
586 |
val algebra = Sorts.project_algebra pp (is_operational_class thy) |
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19929
diff
changeset
|
587 |
(Sign.classes_of thy); |
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19929
diff
changeset
|
588 |
fun classrel (l as Lookup (classes, p), _) class = |
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19929
diff
changeset
|
589 |
Lookup (class :: classes, p) |
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19929
diff
changeset
|
590 |
| classrel (Instance ((_, tyco), lss), _) class = |
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19929
diff
changeset
|
591 |
Instance ((class, tyco), lss); |
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19929
diff
changeset
|
592 |
fun constructor tyco lss class = |
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19929
diff
changeset
|
593 |
Instance ((class, tyco), (map o map) fst lss) |
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19929
diff
changeset
|
594 |
fun variable (TFree (v, sort)) = |
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19929
diff
changeset
|
595 |
map_index (fn (n, class) => (Lookup ([], (v, (n, length sort))), class)) |
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19929
diff
changeset
|
596 |
(operational_sort_of thy sort) |
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19929
diff
changeset
|
597 |
| variable (TVar _) = error "TVar encountered while deriving sortlookup"; |
18864 | 598 |
in |
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19929
diff
changeset
|
599 |
Sorts.of_sort_derivation pp algebra |
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19929
diff
changeset
|
600 |
{classrel = classrel, constructor = constructor, variable = variable} |
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19929
diff
changeset
|
601 |
(typ_ctxt, operational_sort_of thy sort_decl) |
18864 | 602 |
end; |
603 |
||
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19929
diff
changeset
|
604 |
fun sortlookups_const thy (c, typ_ctxt) = |
18864 | 605 |
let |
19957 | 606 |
val typ_decl = case AxClass.class_of thy c |
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19929
diff
changeset
|
607 |
of NONE => Sign.the_const_type thy c |
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19929
diff
changeset
|
608 |
| SOME class => case the_consts_sign thy class of (v, cs) => |
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19929
diff
changeset
|
609 |
(Logic.legacy_varifyT o subst_clsvar v (TFree (v, [class]))) |
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19929
diff
changeset
|
610 |
((the o AList.lookup (op =) cs) c) |
18168 | 611 |
in |
19957 | 612 |
instantiations_of thy (typ_decl, typ_ctxt) |
613 |
|> the |
|
614 |
|> map (fn ((_, sort), ty) => sortlookup thy (sort, ty)) |
|
19953
2f54a51f1801
class package refinements, slight code generation refinements
haftmann
parents:
19929
diff
changeset
|
615 |
|> filter_out null |
18168 | 616 |
end; |
617 |
||
18884 | 618 |
|
18515 | 619 |
(* toplevel interface *) |
620 |
||
621 |
local |
|
622 |
||
623 |
structure P = OuterParse |
|
624 |
and K = OuterKeyword |
|
625 |
||
626 |
in |
|
627 |
||
18849
05a16861d3a5
added three times overloaded Isar instance command
haftmann
parents:
18829
diff
changeset
|
628 |
val (classK, instanceK) = ("class", "instance") |
18515 | 629 |
|
19966 | 630 |
val use_instance2 = ref false; |
19253 | 631 |
|
19966 | 632 |
fun wrap_add_instance_sort (class, sort) thy = |
633 |
if ! use_instance2 |
|
634 |
andalso forall (is_some o lookup_class_data thy) (Sign.read_sort thy sort) |
|
635 |
then |
|
636 |
instance_sort (class, sort) thy |
|
637 |
else |
|
638 |
axclass_instance_sort (class, sort) thy |
|
19136 | 639 |
|
18911 | 640 |
val parse_inst = |
19136 | 641 |
(Scan.optional (P.$$$ "(" |-- P.!!! (P.list1 P.sort --| P.$$$ ")")) [] -- P.xname --| P.$$$ "::" -- P.sort) |
642 |
>> (fn ((asorts, tyco), sort) => ((tyco, asorts), sort)) |
|
643 |
|| (P.xname --| P.$$$ "::" -- P.!!! P.arity) |
|
644 |
>> (fn (tyco, (asorts, sort)) => ((tyco, asorts), sort)); |
|
18911 | 645 |
|
19038 | 646 |
val locale_val = |
647 |
(P.locale_expr -- |
|
648 |
Scan.optional (P.$$$ "+" |-- P.!!! (Scan.repeat1 P.context_element)) [] || |
|
649 |
Scan.repeat1 P.context_element >> pair Locale.empty); |
|
650 |
||
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
651 |
val class_subP = P.name -- Scan.repeat (P.$$$ "+" |-- P.name) >> (op ::); |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
652 |
val class_bodyP = P.!!! (Scan.repeat1 P.context_element); |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
653 |
|
18515 | 654 |
val classP = |
18849
05a16861d3a5
added three times overloaded Isar instance command
haftmann
parents:
18829
diff
changeset
|
655 |
OuterSyntax.command classK "operational type classes" K.thy_decl ( |
05a16861d3a5
added three times overloaded Isar instance command
haftmann
parents:
18829
diff
changeset
|
656 |
P.name --| P.$$$ "=" |
19280
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
657 |
-- ( |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
658 |
class_subP --| P.$$$ "+" -- class_bodyP |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
659 |
|| class_subP >> rpair [] |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
660 |
|| class_bodyP >> pair [] |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
661 |
) >> (Toplevel.theory_context |
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
haftmann
parents:
19253
diff
changeset
|
662 |
o (fn (bname, (supclasses, elems)) => class bname supclasses elems))); |
18515 | 663 |
|
18575 | 664 |
val instanceP = |
18849
05a16861d3a5
added three times overloaded Isar instance command
haftmann
parents:
18829
diff
changeset
|
665 |
OuterSyntax.command instanceK "prove type arity or subclass relation" K.thy_goal (( |
19966 | 666 |
P.xname -- ((P.$$$ "\\<subseteq>" || P.$$$ "<") |-- P.!!! P.xname) >> wrap_add_instance_sort |
19136 | 667 |
|| P.opt_thm_name ":" -- (parse_inst -- Scan.repeat (P.opt_thm_name ":" -- P.prop)) |
19246 | 668 |
>> (fn (("", []), (((tyco, asorts), sort), [])) => axclass_instance_arity I (tyco, asorts, sort) |
19150 | 669 |
| (natts, (inst, defs)) => instance_arity inst natts defs) |
18849
05a16861d3a5
added three times overloaded Isar instance command
haftmann
parents:
18829
diff
changeset
|
670 |
) >> (Toplevel.print oo Toplevel.theory_to_proof)); |
18575 | 671 |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
672 |
val _ = OuterSyntax.add_parsers [classP, instanceP]; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
673 |
|
18515 | 674 |
end; (* local *) |
675 |
||
18168 | 676 |
end; (* struct *) |