author | wenzelm |
Fri, 16 Feb 2007 22:13:15 +0100 | |
changeset 22331 | 7df6bc8cf0b0 |
parent 22321 | e5cddafe2629 |
child 22353 | c818c6b836f5 |
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 |
|
22302 | 10 |
val fork_mixfix: bool -> bool -> mixfix -> mixfix * mixfix |
11 |
||
21954 | 12 |
val class: bstring -> class list -> Element.context_i Locale.element list -> theory -> |
20964 | 13 |
string * Proof.context |
22321 | 14 |
val class_cmd: bstring -> string list -> Element.context Locale.element list -> theory -> |
22302 | 15 |
string * Proof.context |
22331 | 16 |
val instance_arity: arity list -> ((bstring * Attrib.src list) * term) list |
22302 | 17 |
-> theory -> Proof.state |
22331 | 18 |
val instance_arity_cmd: (bstring * string list * string) list |
22302 | 19 |
-> ((bstring * Attrib.src list) * string) list |
18575 | 20 |
-> theory -> Proof.state |
22331 | 21 |
val prove_instance_arity: tactic -> arity list |
22302 | 22 |
-> ((bstring * Attrib.src list) * term) list |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
23 |
-> theory -> theory |
22321 | 24 |
val instance_class: class * class -> theory -> Proof.state |
25 |
val instance_class_cmd: string * string -> theory -> Proof.state |
|
21954 | 26 |
val instance_sort: class * sort -> theory -> Proof.state |
22321 | 27 |
val instance_sort_cmd: string * string -> theory -> Proof.state |
19150 | 28 |
val prove_instance_sort: tactic -> class * sort -> theory -> theory |
18168 | 29 |
|
22331 | 30 |
val assume_arities_of_sort: theory -> arity list -> typ * sort -> bool |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
31 |
|
21954 | 32 |
(* experimental class target *) |
22302 | 33 |
val class_of_locale: theory -> string -> class option |
34 |
val add_def_in_class: local_theory -> string |
|
35 |
-> (string * Syntax.mixfix) * ((string * Attrib.src list) * (term * thm)) -> theory -> theory |
|
21553 | 36 |
|
18702 | 37 |
val print_classes: theory -> unit |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
38 |
val intro_classes_tac: thm list -> tactic |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
39 |
val default_intro_classes_tac: thm list -> tactic |
18168 | 40 |
end; |
41 |
||
20182
79c9ff40d760
tactic for prove_instance_arity now gets definition theorems as arguments
haftmann
parents:
20175
diff
changeset
|
42 |
structure ClassPackage : CLASS_PACKAGE = |
18168 | 43 |
struct |
44 |
||
22302 | 45 |
(** auxiliary **) |
18168 | 46 |
|
22302 | 47 |
fun fork_mixfix is_class is_loc mx = |
48 |
let |
|
49 |
val mx' = Syntax.unlocalize_mixfix mx; |
|
50 |
val mx1 = if is_class orelse (is_loc andalso mx = mx') then NoSyn else mx'; |
|
51 |
val mx2 = if is_loc then mx else NoSyn; |
|
52 |
in (mx1, mx2) end; |
|
53 |
||
54 |
||
55 |
(** AxClasses with implicit parameter handling **) |
|
56 |
||
57 |
(* AxClass instances *) |
|
18168 | 58 |
|
21954 | 59 |
local |
19456 | 60 |
|
21954 | 61 |
fun gen_instance mk_prop add_thm after_qed insts thy = |
62 |
let |
|
63 |
fun after_qed' results = |
|
64 |
ProofContext.theory ((fold o fold) add_thm results #> after_qed); |
|
65 |
in |
|
66 |
thy |
|
67 |
|> ProofContext.init |
|
68 |
|> Proof.theorem_i NONE after_qed' ((map (fn t => [(t, [])]) o maps (mk_prop thy)) insts) |
|
69 |
end; |
|
18168 | 70 |
|
21954 | 71 |
in |
72 |
||
73 |
val axclass_instance_arity = |
|
74 |
gen_instance (Logic.mk_arities oo Sign.cert_arity) AxClass.add_arity; |
|
75 |
val axclass_instance_sort = |
|
76 |
gen_instance (single oo (Logic.mk_classrel oo AxClass.cert_classrel)) |
|
77 |
AxClass.add_classrel I o single; |
|
78 |
||
79 |
end; (* local *) |
|
18575 | 80 |
|
19038 | 81 |
|
21954 | 82 |
(* introducing axclasses with implicit parameter handling *) |
19038 | 83 |
|
21954 | 84 |
fun axclass_params (name, raw_superclasses) raw_consts raw_dep_axioms thy = |
21924 | 85 |
let |
21954 | 86 |
val superclasses = map (Sign.certify_class thy) raw_superclasses; |
87 |
val consts = (map o apfst o apsnd) (Sign.certify_typ thy) raw_consts; |
|
88 |
fun add_const ((c, ty), syn) = |
|
89 |
Sign.add_consts_authentic [(c, ty, syn)] #> pair (Sign.full_name thy c, ty); |
|
90 |
fun mk_axioms cs thy = |
|
91 |
raw_dep_axioms thy cs |
|
92 |
|> (map o apsnd o map) (Sign.cert_prop thy) |
|
93 |
|> rpair thy; |
|
94 |
fun add_constraint class (c, ty) = |
|
95 |
Sign.add_const_constraint_i (c, SOME (Term.map_type_tfree (fn (v, _) => TFree (v, [class])) ty)); |
|
21924 | 96 |
in |
21954 | 97 |
thy |
98 |
|> fold_map add_const consts |
|
99 |
|-> (fn cs => mk_axioms cs |
|
22302 | 100 |
#-> (fn axioms_prop => AxClass.define_class_i (name, superclasses) (map fst cs) axioms_prop |
21954 | 101 |
#-> (fn class => `(fn thy => AxClass.get_definition thy class) |
102 |
#-> (fn {intro, axioms, ...} => fold (add_constraint class) cs |
|
22321 | 103 |
#> pair (class, ((intro, (maps snd axioms_prop, axioms)), cs)))))) |
21924 | 104 |
end; |
105 |
||
18168 | 106 |
|
21954 | 107 |
(* contexts with arity assumptions *) |
20175 | 108 |
|
109 |
fun assume_arities_of_sort thy arities ty_sort = |
|
110 |
let |
|
111 |
val pp = Sign.pp thy; |
|
112 |
val algebra = Sign.classes_of thy |
|
22331 | 113 |
|> fold (fn (tyco, asorts, sort) => |
20175 | 114 |
Sorts.add_arities pp (tyco, map (fn class => (class, asorts)) sort)) arities; |
115 |
in Sorts.of_sort algebra ty_sort end; |
|
116 |
||
20455 | 117 |
|
21954 | 118 |
(* instances with implicit parameter handling *) |
19038 | 119 |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
120 |
local |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
121 |
|
20175 | 122 |
fun gen_read_def thy prep_att read_def ((raw_name, raw_atts), raw_t) = |
19038 | 123 |
let |
19957 | 124 |
val (_, t) = read_def thy (raw_name, raw_t); |
125 |
val ((c, ty), _) = Sign.cert_def (Sign.pp thy) t; |
|
126 |
val atts = map (prep_att thy) raw_atts; |
|
20175 | 127 |
val insts = (Consts.typargs (Sign.consts_of thy) (c, ty)) |
19957 | 128 |
val name = case raw_name |
20385 | 129 |
of "" => NONE |
130 |
| _ => SOME raw_name; |
|
20175 | 131 |
in (c, (insts, ((name, t), atts))) end; |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
132 |
|
22321 | 133 |
fun read_def_cmd thy = gen_read_def thy Attrib.intern_src read_axm; |
21954 | 134 |
fun read_def thy = gen_read_def thy (K I) (K I); |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
135 |
|
22302 | 136 |
fun gen_instance_arity prep_arity read_def do_proof raw_arities raw_defs theory = |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
137 |
let |
22331 | 138 |
val arities = map (prep_arity theory) raw_arities; |
20175 | 139 |
val _ = if null arities then error "at least one arity must be given" else (); |
140 |
val _ = case (duplicates (op =) o map #1) arities |
|
141 |
of [] => () |
|
142 |
| dupl_tycos => error ("type constructors occur more than once in arities: " |
|
143 |
^ (commas o map quote) dupl_tycos); |
|
21954 | 144 |
val super_sort = (Graph.all_succs o #classes o Sorts.rep_algebra o Sign.classes_of) theory |
20175 | 145 |
fun get_consts_class tyco ty class = |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
146 |
let |
21954 | 147 |
val cs = (these o Option.map snd o try (AxClass.params_of_class theory)) class; |
21463
42dd50268c8b
completed class parameter handling in axclass.ML
haftmann
parents:
21444
diff
changeset
|
148 |
val subst_ty = map_type_tfree (K ty); |
19957 | 149 |
in |
21894 | 150 |
map (fn (c, ty) => (c, ((tyco, class), subst_ty ty))) cs |
19957 | 151 |
end; |
20175 | 152 |
fun get_consts_sort (tyco, asorts, sort) = |
19957 | 153 |
let |
20192
956cd30ef3be
renamed Name.give_names to Name.names and moved Name.alphanum to Symbol.alphanum
haftmann
parents:
20191
diff
changeset
|
154 |
val ty = Type (tyco, map (fn (v, sort) => TVar ((v, 0), sort)) (Name.names Name.context "'a" asorts)) |
21954 | 155 |
in maps (get_consts_class tyco ty) (super_sort sort) end; |
20175 | 156 |
val cs = maps get_consts_sort arities; |
20385 | 157 |
fun mk_typnorm thy (ty, ty_sc) = |
158 |
case try (Sign.typ_match thy (Logic.varifyT ty_sc, ty)) Vartab.empty |
|
159 |
of SOME env => SOME (Logic.varifyT #> Envir.typ_subst_TVars env #> Logic.unvarifyT) |
|
160 |
| NONE => NONE; |
|
20175 | 161 |
fun read_defs defs cs thy_read = |
162 |
let |
|
19957 | 163 |
fun read raw_def cs = |
164 |
let |
|
20385 | 165 |
val (c, (inst, ((name_opt, t), atts))) = read_def thy_read raw_def; |
166 |
val ty = Consts.instance (Sign.consts_of thy_read) (c, inst); |
|
20175 | 167 |
val ((tyco, class), ty') = case AList.lookup (op =) cs c |
19957 | 168 |
of NONE => error ("superfluous definition for constant " ^ quote c) |
19966 | 169 |
| SOME class_ty => class_ty; |
20385 | 170 |
val name = case name_opt |
20628 | 171 |
of NONE => Thm.def_name (Logic.name_arity (tyco, [], c)) |
20385 | 172 |
| SOME name => name; |
173 |
val t' = case mk_typnorm thy_read (ty', ty) |
|
19966 | 174 |
of NONE => error ("superfluous definition for constant " ^ |
175 |
quote c ^ "::" ^ Sign.string_of_typ thy_read ty) |
|
20548
8ef25fe585a8
renamed Term.map_term_types to Term.map_types (cf. Term.fold_types);
wenzelm
parents:
20465
diff
changeset
|
176 |
| SOME norm => map_types norm t |
20175 | 177 |
in (((class, tyco), ((name, t'), atts)), AList.delete (op =) c cs) end; |
19957 | 178 |
in fold_map read defs cs end; |
22331 | 179 |
val (defs, _) = read_defs raw_defs cs |
180 |
(fold Sign.primitive_arity arities (Theory.copy theory)); |
|
181 |
||
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
182 |
fun get_remove_contraint c thy = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
183 |
let |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
184 |
val ty = Sign.the_const_constraint thy c; |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
185 |
in |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
186 |
thy |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
187 |
|> Sign.add_const_constraint_i (c, NONE) |
20385 | 188 |
|> pair (c, Logic.unvarifyT ty) |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
189 |
end; |
19966 | 190 |
fun add_defs defs thy = |
191 |
thy |
|
22302 | 192 |
|> PureThy.add_defs_i true (map ((apsnd o map) (Attrib.attribute thy) o snd) defs) |
19966 | 193 |
|-> (fn thms => pair (map fst defs ~~ thms)); |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
194 |
fun after_qed cs thy = |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
195 |
thy |
19412 | 196 |
|> fold Sign.add_const_constraint_i (map (apsnd SOME) cs); |
19038 | 197 |
in |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
198 |
theory |
20385 | 199 |
|> fold_map get_remove_contraint (map fst cs |> distinct (op =)) |
19966 | 200 |
||>> add_defs defs |
21954 | 201 |
|-> (fn (cs, _) => do_proof (after_qed cs) arities) |
19038 | 202 |
end; |
203 |
||
22321 | 204 |
fun instance_arity_e' do_proof = gen_instance_arity Sign.read_arity read_def_cmd do_proof; |
22302 | 205 |
fun instance_arity' do_proof = gen_instance_arity Sign.cert_arity read_def do_proof; |
21924 | 206 |
fun tactic_proof tac after_qed arities = |
207 |
fold (fn arity => AxClass.prove_arity arity tac) arities |
|
20175 | 208 |
#> after_qed; |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
209 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
210 |
in |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
211 |
|
22321 | 212 |
val instance_arity_cmd = instance_arity_e' axclass_instance_arity; |
21954 | 213 |
val instance_arity = instance_arity' axclass_instance_arity; |
214 |
val prove_instance_arity = instance_arity' o tactic_proof; |
|
215 |
||
216 |
end; (* local *) |
|
217 |
||
218 |
||
219 |
||
220 |
(** combining locales and axclasses **) |
|
221 |
||
222 |
(* theory data *) |
|
223 |
||
224 |
datatype class_data = ClassData of { |
|
225 |
locale: string, |
|
226 |
consts: (string * string) list |
|
227 |
(*locale parameter ~> toplevel theory constant*), |
|
22302 | 228 |
witness: Element.witness list, |
21954 | 229 |
propnames: string list, |
22302 | 230 |
(*for ad-hoc instance_in*) |
231 |
primdefs: thm list |
|
21954 | 232 |
(*for experimental class target*) |
233 |
}; |
|
234 |
||
235 |
fun rep_classdata (ClassData c) = c; |
|
236 |
||
22302 | 237 |
fun merge_pair f1 f2 ((x1, y1), (x2, y2)) = (f1 (x1, x2), f2 (y1, y2)); |
238 |
||
21954 | 239 |
structure ClassData = TheoryDataFun ( |
240 |
struct |
|
241 |
val name = "Pure/classes"; |
|
22302 | 242 |
type T = class_data Graph.T * class Symtab.table (*locale name ~> class name*); |
243 |
val empty = (Graph.empty, Symtab.empty); |
|
21954 | 244 |
val copy = I; |
245 |
val extend = I; |
|
22302 | 246 |
fun merge _ = merge_pair (Graph.merge (K true)) (Symtab.merge (K true)); |
21954 | 247 |
fun print _ _ = (); |
248 |
end |
|
249 |
); |
|
250 |
||
251 |
val _ = Context.add_setup ClassData.init; |
|
252 |
||
253 |
||
254 |
(* queries *) |
|
255 |
||
22302 | 256 |
val lookup_class_data = Option.map rep_classdata oo try o Graph.get_node o fst o ClassData.get; |
257 |
fun class_of_locale thy = Symtab.lookup ((snd o ClassData.get) thy); |
|
258 |
||
21954 | 259 |
fun the_class_data thy class = |
260 |
case lookup_class_data thy class |
|
261 |
of NONE => error ("undeclared class " ^ quote class) |
|
262 |
| SOME data => data; |
|
263 |
||
22302 | 264 |
val ancestry = Graph.all_succs o fst o ClassData.get; |
21954 | 265 |
|
22331 | 266 |
fun param_map thy = |
21954 | 267 |
let |
22302 | 268 |
fun params thy class = |
269 |
let |
|
270 |
val const_typs = (#params o AxClass.get_definition thy) class; |
|
271 |
val const_names = (#consts o the_class_data thy) class; |
|
272 |
in |
|
273 |
(map o apsnd) (fn c => (c, (the o AList.lookup (op =) const_typs) c)) const_names |
|
274 |
end; |
|
275 |
in maps (params thy) o ancestry thy end; |
|
21954 | 276 |
|
22302 | 277 |
val the_witness = #witness oo the_class_data; |
21954 | 278 |
val the_propnames = #propnames oo the_class_data; |
279 |
||
280 |
fun print_classes thy = |
|
281 |
let |
|
282 |
val algebra = Sign.classes_of thy; |
|
283 |
val arities = |
|
284 |
Symtab.empty |
|
285 |
|> Symtab.fold (fn (tyco, arities) => fold (fn (class, _) => |
|
286 |
Symtab.map_default (class, []) (insert (op =) tyco)) arities) |
|
287 |
((#arities o Sorts.rep_algebra) algebra); |
|
288 |
val the_arities = these o Symtab.lookup arities; |
|
289 |
fun mk_arity class tyco = |
|
290 |
let |
|
291 |
val Ss = Sorts.mg_domain algebra tyco [class]; |
|
292 |
in Sign.pretty_arity thy (tyco, Ss, [class]) end; |
|
293 |
fun mk_param (c, ty) = Pretty.str (Sign.extern_const thy c ^ " :: " |
|
294 |
^ setmp show_sorts false (Sign.string_of_typ thy o Type.strip_sorts) ty); |
|
295 |
fun mk_entry class = (Pretty.block o Pretty.fbreaks o map_filter I) [ |
|
296 |
(SOME o Pretty.str) ("class " ^ class ^ ":"), |
|
297 |
(SOME o Pretty.block) [Pretty.str "supersort: ", |
|
298 |
(Sign.pretty_sort thy o Sign.certify_sort thy o Sign.super_classes thy) class], |
|
299 |
Option.map (Pretty.str o prefix "locale: " o #locale) (lookup_class_data thy class), |
|
300 |
((fn [] => NONE | ps => (SOME o Pretty.block o Pretty.fbreaks) (Pretty.str "parameters:" :: ps)) o map mk_param |
|
301 |
o these o Option.map #params o try (AxClass.get_definition thy)) class, |
|
302 |
(SOME o Pretty.block o Pretty.breaks) [ |
|
303 |
Pretty.str "instances:", |
|
304 |
Pretty.list "" "" (map (mk_arity class) (the_arities class)) |
|
305 |
] |
|
306 |
] |
|
307 |
in |
|
308 |
(Pretty.writeln o Pretty.chunks o separate (Pretty.str "") o map mk_entry o Sorts.all_classes) |
|
309 |
algebra |
|
310 |
end; |
|
311 |
||
312 |
||
313 |
(* updaters *) |
|
314 |
||
22302 | 315 |
fun add_class_data ((class, superclasses), (locale, consts, witness, propnames)) = |
316 |
ClassData.map (fn (gr, tab) => ( |
|
317 |
gr |
|
318 |
|> Graph.new_node (class, ClassData { |
|
21954 | 319 |
locale = locale, |
320 |
consts = consts, |
|
22302 | 321 |
witness = witness, |
21954 | 322 |
propnames = propnames, |
22302 | 323 |
primdefs = []}) |
324 |
|> fold (curry Graph.add_edge class) superclasses, |
|
325 |
tab |
|
326 |
|> Symtab.update (locale, class) |
|
327 |
)); |
|
328 |
||
329 |
fun add_primdef (class, thm) thy = |
|
330 |
(ClassData.map o apfst o Graph.map_node class) |
|
331 |
(fn ClassData { locale, consts, witness, propnames, primdefs } => ClassData { locale = locale, |
|
332 |
consts = consts, witness = witness, propnames = propnames, primdefs = thm :: primdefs }); |
|
333 |
||
21954 | 334 |
|
22302 | 335 |
(* exporting terms and theorems to global toplevel *) |
336 |
(*FIXME should also be used when introducing classes*) |
|
337 |
||
338 |
fun globalize thy classes = |
|
339 |
let |
|
340 |
val consts = param_map thy classes; |
|
341 |
val constrain_sort = curry (Sorts.inter_sort (Sign.classes_of thy)) classes; |
|
342 |
val subst_typ = Term.map_type_tfree (fn var as (v, sort) => |
|
343 |
if v = AxClass.param_tyvarname then TFree (v, constrain_sort sort) else TFree var) |
|
344 |
fun subst_aterm (t as Free (v, ty)) = (case AList.lookup (op =) consts v |
|
345 |
of SOME (c, _) => Const (c, ty) |
|
346 |
| NONE => t) |
|
347 |
| subst_aterm t = t; |
|
348 |
in (subst_typ, map_types subst_typ #> Term.map_aterms subst_aterm) end; |
|
349 |
||
350 |
val global_term = snd oo globalize |
|
21954 | 351 |
|
352 |
||
353 |
(* tactics and methods *) |
|
354 |
||
22321 | 355 |
(*FIXME adjust these when minimal intro rules are at hand*) |
21954 | 356 |
fun intro_classes_tac facts st = |
22182 | 357 |
let |
358 |
val thy = Thm.theory_of_thm st; |
|
359 |
val ctxt = ProofContext.init thy; |
|
360 |
val intro_classes_tac = ALLGOALS |
|
361 |
(Method.insert_tac facts THEN' REPEAT_ALL_NEW (resolve_tac (AxClass.class_intros thy))) |
|
362 |
THEN Tactic.distinct_subgoals_tac; |
|
22321 | 363 |
val intro_locales_tac = SOMEGOAL (SELECT_GOAL (Locale.intro_locales_tac true ctxt facts)) |
364 |
THEN Tactic.distinct_subgoals_tac; |
|
22182 | 365 |
in |
22321 | 366 |
(intro_classes_tac THEN REPEAT (intro_locales_tac ORELSE intro_locales_tac)) st |
22182 | 367 |
end; |
21954 | 368 |
|
369 |
fun default_intro_classes_tac [] = intro_classes_tac [] |
|
370 |
| default_intro_classes_tac _ = Tactical.no_tac; (*no error message!*) |
|
371 |
||
372 |
fun default_tac rules ctxt facts = |
|
373 |
HEADGOAL (Method.some_rule_tac rules ctxt facts) ORELSE |
|
374 |
default_intro_classes_tac facts; |
|
375 |
||
376 |
val _ = Context.add_setup (Method.add_methods |
|
377 |
[("intro_classes", Method.no_args (Method.METHOD intro_classes_tac), |
|
378 |
"back-chain introduction rules of classes"), |
|
379 |
("default", Method.thms_ctxt_args (Method.METHOD oo default_tac), |
|
380 |
"apply some intro/elim rule")]); |
|
381 |
||
22321 | 382 |
(* tactical interfaces to locale commands *) |
21954 | 383 |
|
22321 | 384 |
fun prove_interpretation tac prfx_atts expr insts thy = |
385 |
thy |
|
386 |
|> Locale.interpretation_i I prfx_atts expr insts |
|
387 |
|> Proof.global_terminal_proof (Method.Basic (fn _ => Method.SIMPLE_METHOD tac), NONE) |
|
388 |
|> ProofContext.theory_of; |
|
21954 | 389 |
|
22069 | 390 |
fun prove_interpretation_in tac after_qed (name, expr) thy = |
391 |
thy |
|
392 |
|> Locale.interpretation_in_locale (ProofContext.theory after_qed) (name, expr) |
|
393 |
|> Proof.global_terminal_proof (Method.Basic (fn _ => Method.SIMPLE_METHOD tac), NONE) |
|
394 |
|> ProofContext.theory_of; |
|
395 |
||
396 |
fun instance_sort' do_proof (class, sort) theory = |
|
397 |
let |
|
398 |
val loc_name = (#locale o the_class_data theory) class; |
|
399 |
val loc_expr = |
|
400 |
(Locale.Merge o map (Locale.Locale o #locale o the_class_data theory)) sort; |
|
401 |
val const_names = (map (NameSpace.base o snd) |
|
402 |
o maps (#consts o the_class_data theory) |
|
22302 | 403 |
o ancestry theory) [class]; |
22069 | 404 |
fun get_thms thy = |
22302 | 405 |
ancestry thy sort |
22069 | 406 |
|> AList.make (the_propnames thy) |
407 |
|> map (apsnd (map (NameSpace.append (Logic.const_of_class loc_name)))) |
|
408 |
|> map_filter (fn (superclass, thm_names) => |
|
409 |
case map_filter (try (PureThy.get_thm thy o Name)) thm_names |
|
410 |
of [] => NONE |
|
411 |
| thms => SOME (superclass, thms)); |
|
412 |
fun after_qed thy = |
|
413 |
thy |
|
414 |
|> `get_thms |
|
415 |
|-> fold (fn (supclass, thms) => I |
|
416 |
AxClass.prove_classrel (class, supclass) |
|
417 |
(ALLGOALS (K (intro_classes_tac [])) THEN |
|
418 |
(ALLGOALS o ProofContext.fact_tac) thms)) |
|
419 |
in |
|
420 |
theory |
|
421 |
|> do_proof after_qed (loc_name, loc_expr) |
|
422 |
end; |
|
423 |
||
22074 | 424 |
val prove_instance_sort = instance_sort' o prove_interpretation_in; |
425 |
||
22069 | 426 |
|
427 |
(* classes and instances *) |
|
428 |
||
429 |
local |
|
430 |
||
431 |
fun add_axclass (name, supsort) params axs thy = |
|
432 |
let |
|
433 |
val (c, thy') = thy |
|
434 |
|> AxClass.define_class_i (name, supsort) params axs; |
|
435 |
val {intro, axioms, ...} = AxClass.get_definition thy' c; |
|
436 |
in ((c, (intro, axioms)), thy') end; |
|
437 |
||
438 |
fun certify_class thy class = |
|
439 |
tap (the_class_data thy) (Sign.certify_class thy class); |
|
440 |
||
441 |
fun read_class thy = |
|
442 |
certify_class thy o Sign.intern_class thy; |
|
443 |
||
21954 | 444 |
fun gen_class add_locale prep_class bname raw_supclasses raw_elems thy = |
445 |
let |
|
22321 | 446 |
(*FIXME need proper concept for reading locale statements*) |
447 |
fun subst_classtyvar (_, _) = |
|
448 |
TFree (AxClass.param_tyvarname, []) |
|
449 |
| subst_classtyvar (v, sort) = |
|
450 |
error ("Sort constraint illegal in type class, for type variable " ^ v ^ "::" ^ Sign.string_of_sort thy sort); |
|
451 |
(*val subst_classtyvars = Element.map_ctxt {name = I, var = I, term = I, |
|
452 |
typ = Term.map_type_tfree subst_classtyvar, fact = I, attrib = I};*) |
|
453 |
val elems = fold_rev (fn Locale.Elem e => cons e | _ => I) |
|
454 |
raw_elems []; (*FIXME make include possible here?*) |
|
21954 | 455 |
val supclasses = map (prep_class thy) raw_supclasses; |
456 |
val supsort = |
|
457 |
supclasses |
|
458 |
|> Sign.certify_sort thy |
|
22302 | 459 |
|> (fn [] => Sign.defaultS thy | S => S); (*FIXME Why syntax defaultS?*) |
22321 | 460 |
val supexpr = Locale.Merge |
461 |
(map (Locale.Locale o #locale o the_class_data thy) supclasses); |
|
462 |
val supparams = (map fst o Locale.parameters_of_expr thy) supexpr; |
|
21954 | 463 |
val supconsts = AList.make |
22321 | 464 |
(the o AList.lookup (op =) (param_map thy supclasses)) (map fst supparams); |
465 |
(*FIXME include proper*) |
|
466 |
(*val elems_constrains = map |
|
467 |
(Element.Constrains o apsnd (Term.map_type_tfree subst_classtyvar)) supparams;*) |
|
468 |
fun extract_params thy name_locale = |
|
21954 | 469 |
let |
22321 | 470 |
val params = Locale.parameters_of thy name_locale; |
471 |
in |
|
472 |
(map (fst o fst) params, params |
|
473 |
|> (map o apfst o apsnd o Term.map_type_tfree) (K (TFree (AxClass.param_tyvarname, Sign.defaultS thy))) |
|
474 |
|> (map o apsnd) (fork_mixfix false true #> fst) |
|
475 |
|> chop (length supconsts) |
|
476 |
|> snd) |
|
21954 | 477 |
end; |
478 |
fun extract_assumes name_locale params thy cs = |
|
479 |
let |
|
480 |
val consts = supconsts @ (map (fst o fst) params ~~ cs); |
|
481 |
fun subst (Free (c, ty)) = |
|
22048 | 482 |
Const ((fst o the o AList.lookup (op =) consts) c, ty) |
483 |
| subst t = t; |
|
21954 | 484 |
fun prep_asm ((name, atts), ts) = |
22321 | 485 |
(*FIXME: name handling?*) |
21954 | 486 |
((NameSpace.base name, map (Attrib.attribute thy) atts), (map o map_aterms) subst ts); |
487 |
in |
|
22182 | 488 |
Locale.global_asms_of thy name_locale |
21954 | 489 |
|> map prep_asm |
490 |
end; |
|
22048 | 491 |
fun extract_axiom_names thy name_locale = |
492 |
name_locale |
|
22182 | 493 |
|> Locale.global_asms_of thy |
22321 | 494 |
|> map (NameSpace.base o fst o fst) (*FIXME - is finally dropped*) |
495 |
fun mk_instT class = Symtab.empty |
|
496 |
|> Symtab.update (AxClass.param_tyvarname, TFree (AxClass.param_tyvarname, [class])); |
|
497 |
fun mk_inst class param_names cs = |
|
498 |
Symtab.empty |
|
499 |
|> fold2 (fn v => fn (c, ty) => Symtab.update (v, Const |
|
500 |
(c, Term.map_type_tfree (fn (v, _) => TFree (v, [class])) ty))) param_names cs; |
|
501 |
fun make_witness t thm = Element.make_witness t (Goal.protect thm); |
|
21954 | 502 |
in |
503 |
thy |
|
22321 | 504 |
|> add_locale bname supexpr ((*elems_constrains @*) elems) |
21954 | 505 |
|-> (fn name_locale => ProofContext.theory_result ( |
22321 | 506 |
`(fn thy => extract_params thy name_locale) |
507 |
#-> (fn (param_names, params) => |
|
21954 | 508 |
axclass_params (bname, supsort) params (extract_assumes name_locale params) |
22302 | 509 |
#-> (fn (name_axclass, ((_, (ax_terms, ax_axioms)), consts)) => |
22048 | 510 |
`(fn thy => extract_axiom_names thy name_locale) |
511 |
#-> (fn axiom_names => |
|
512 |
add_class_data ((name_axclass, supclasses), |
|
22302 | 513 |
(name_locale, map (fst o fst) params ~~ map fst consts, |
22321 | 514 |
map2 make_witness ax_terms ax_axioms, axiom_names)) |
515 |
#> prove_interpretation ((ALLGOALS o ProofContext.fact_tac) ax_axioms) |
|
516 |
(Logic.const_of_class bname, []) (Locale.Locale name_locale) |
|
517 |
(mk_instT name_axclass, mk_inst name_axclass param_names (map snd supconsts @ consts)) |
|
21954 | 518 |
#> pair name_axclass |
22048 | 519 |
))))) |
21954 | 520 |
end; |
521 |
||
522 |
in |
|
523 |
||
22321 | 524 |
val class_cmd = gen_class (Locale.add_locale true) read_class; |
22182 | 525 |
val class = gen_class (Locale.add_locale_i true) certify_class; |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
526 |
|
20455 | 527 |
end; (*local*) |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
528 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
529 |
local |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
530 |
|
21954 | 531 |
fun gen_instance_sort prep_class prep_sort (raw_class, raw_sort) theory = |
532 |
let |
|
533 |
val class = prep_class theory raw_class; |
|
534 |
val sort = prep_sort theory raw_sort; |
|
22321 | 535 |
in |
21954 | 536 |
theory |
537 |
|> instance_sort' (Locale.interpretation_in_locale o ProofContext.theory) (class, sort) |
|
22321 | 538 |
end; |
539 |
||
540 |
fun gen_instance_class prep_class (raw_class, raw_superclass) theory = |
|
541 |
let |
|
542 |
val class = prep_class theory raw_class; |
|
543 |
val superclass = prep_class theory raw_superclass; |
|
544 |
in |
|
545 |
theory |
|
546 |
|> axclass_instance_sort (class, superclass) |
|
21954 | 547 |
end; |
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
548 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
549 |
in |
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
550 |
|
22321 | 551 |
val instance_sort_cmd = gen_instance_sort Sign.read_class Sign.read_sort; |
21954 | 552 |
val instance_sort = gen_instance_sort Sign.certify_class Sign.certify_sort; |
22321 | 553 |
val instance_class_cmd = gen_instance_class Sign.read_class; |
554 |
val instance_class = gen_instance_class Sign.certify_class; |
|
19110
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
555 |
|
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
haftmann
parents:
19102
diff
changeset
|
556 |
end; (* local *) |
19038 | 557 |
|
19957 | 558 |
|
21954 | 559 |
(** experimental class target **) |
560 |
||
22302 | 561 |
fun add_def_in_class lthy loc ((c, syn), ((name, atts), (rhs, eq))) thy = |
21954 | 562 |
let |
22302 | 563 |
val SOME class = class_of_locale thy loc; |
564 |
val rhs' = global_term thy [class] rhs; |
|
565 |
val (syn', _) = fork_mixfix true true syn; |
|
566 |
val ty = Term.fastype_of rhs'; |
|
567 |
fun add_const (c, ty, syn) = |
|
568 |
Sign.add_consts_authentic [(c, ty, syn)] |
|
569 |
#> pair (Sign.full_name thy c, ty); |
|
570 |
fun add_def ((name, atts), prop) thy = |
|
571 |
thy |
|
572 |
|> PureThy.add_defs_i false [((name, prop), map (Attrib.attribute thy) atts)] |
|
573 |
|>> the_single; |
|
574 |
(*val _ = Output.info "------ DEF ------"; |
|
575 |
val _ = Output.info c; |
|
576 |
val _ = Output.info name; |
|
577 |
val _ = (Output.info o Sign.string_of_term thy) rhs'; |
|
578 |
val eq' = Morphism.thm (LocalTheory.target_morphism lthy) eq; |
|
579 |
val _ = (Output.info o string_of_thm) eq; |
|
580 |
val _ = (Output.info o string_of_thm) eq'; |
|
581 |
val witness = the_witness thy class; |
|
582 |
val _ = Output.info "------ WIT ------"; |
|
583 |
fun print_wit (t, thm) = "(" ^ Sign.string_of_term thy t ^ ", " ^ Display.string_of_thm thm ^ ")" |
|
584 |
val _ = (Output.info o cat_lines o map (print_wit o Element.dest_witness)) witness; |
|
585 |
val _ = (Output.info o string_of_thm) eq'; |
|
586 |
val eq'' = Element.satisfy_thm witness eq'; |
|
587 |
val _ = (Output.info o string_of_thm) eq'';*) |
|
21954 | 588 |
in |
22302 | 589 |
thy |
590 |
(*|> add_const (c, ty, syn') |
|
591 |
|-> (fn c => add_def ((name, atts), Logic.mk_equals (Const c, rhs'))) |
|
592 |
|-> (fn global_def_thm => tap (fn _ => (Output.info o string_of_thm) global_def_thm))*) |
|
21954 | 593 |
end; |
594 |
||
20455 | 595 |
end; |