1 (* Title: Pure/axclass.ML
3 Author: Markus Wenzel, TU Muenchen
4 License: GPL (GNU GENERAL PUBLIC LICENSE)
6 Axiomatic type class package.
11 val quiet_mode: bool ref
12 val print_axclasses: theory -> unit
13 val add_classrel_thms: thm list -> theory -> theory
14 val add_arity_thms: thm list -> theory -> theory
15 val add_axclass: bclass * xclass list -> ((bstring * string) * Args.src list) list
16 -> theory -> theory * {intro: thm, axioms: thm list}
17 val add_axclass_i: bclass * class list -> ((bstring * term) * theory attribute list) list
18 -> theory -> theory * {intro: thm, axioms: thm list}
19 val add_inst_subclass_x: xclass * xclass -> string list -> thm list
20 -> tactic option -> theory -> theory
21 val add_inst_subclass: xclass * xclass -> tactic -> theory -> theory
22 val add_inst_subclass_i: class * class -> tactic -> theory -> theory
23 val add_inst_arity_x: xstring * string list * string -> string list
24 -> thm list -> tactic option -> theory -> theory
25 val add_inst_arity: xstring * string list * string -> tactic -> theory -> theory
26 val add_inst_arity_i: string * sort list * sort -> tactic -> theory -> theory
27 val default_intro_classes_tac: thm list -> int -> tactic
28 val axclass_tac: thm list -> tactic
29 val instance_subclass_proof: (xclass * xclass) * Comment.text -> bool -> theory -> ProofHistory.T
30 val instance_subclass_proof_i: (class * class) * Comment.text -> bool -> theory -> ProofHistory.T
31 val instance_arity_proof: (xstring * string list * xclass) * Comment.text
32 -> bool -> theory -> ProofHistory.T
33 val instance_arity_proof_i: (string * sort list * class) * Comment.text
34 -> bool -> theory -> ProofHistory.T
35 val setup: (theory -> theory) list
38 structure AxClass : AX_CLASS =
46 val quiet_mode = ref false;
47 fun message s = if ! quiet_mode then () else writeln s;
52 fun map_typ_frees f (Type (t, tys)) = Type (t, map (map_typ_frees f) tys)
53 | map_typ_frees f (TFree a) = f a
54 | map_typ_frees _ a = a;
56 val map_term_tfrees = map_term_types o map_typ_frees;
58 fun aT S = TFree ("'a", S);
60 fun dest_varT (TFree (x, S)) = ((x, ~1), S)
61 | dest_varT (TVar xi_S) = xi_S
62 | dest_varT T = raise TYPE ("dest_varT", [T], []);
65 (* get axioms and theorems *)
67 val is_def = Logic.is_equals o #prop o rep_thm;
69 fun witnesses thy names thms =
70 PureThy.get_thmss thy names @ thms @ filter is_def (map snd (axioms_of thy));
74 (** abstract syntax operations **)
78 fun intro_name c = c ^ "I";
80 val axiomsN = "axioms";
83 (* subclass relations as terms *)
85 fun mk_classrel (c1, c2) = Logic.mk_inclass (aT [c1], c2);
87 fun dest_classrel tm =
89 fun err () = raise TERM ("dest_classrel", [tm]);
91 val (ty, c2) = Logic.dest_inclass tm handle TERM _ => err ();
92 val c1 = (case dest_varT ty of (_, [c]) => c | _ => err ())
93 handle TYPE _ => err ();
97 (* arities as terms *)
99 fun mk_arity (t, ss, c) =
101 val tfrees = ListPair.map TFree (Term.invent_type_names [] (length ss), ss);
102 in Logic.mk_inclass (Type (t, tfrees), c) end;
106 fun err () = raise TERM ("dest_arity", [tm]);
108 val (ty, c) = Logic.dest_inclass tm handle TERM _ => err ();
111 Type (t, tys) => (t, map dest_varT tys handle TYPE _ => err ())
114 if null (gen_duplicates eq_fst tvars)
115 then map snd tvars else err ();
120 (** add theorems as axioms **)
122 fun prep_thm_axm thy thm =
124 fun err msg = raise THM ("prep_thm_axm: " ^ msg, 0, [thm]);
126 val {sign, hyps, prop, ...} = rep_thm thm;
128 if not (Sign.subsig (sign, Theory.sign_of thy)) then
129 err "theorem not of same theory"
130 else if not (null (extra_shyps thm)) orelse not (null hyps) then
131 err "theorem may not contain hypotheses"
135 (*theorems expressing class relations*)
136 fun add_classrel_thms thms thy =
140 val prop = prep_thm_axm thy thm;
141 val (c1, c2) = dest_classrel prop handle TERM _ =>
142 raise THM ("add_classrel_thms: theorem is not a class relation", 0, [thm]);
144 in Theory.add_classrel_i (map prep_thm thms) thy end;
146 (*theorems expressing arities*)
147 fun add_arity_thms thms thy =
151 val prop = prep_thm_axm thy thm;
152 val (t, ss, c) = dest_arity prop handle TERM _ =>
153 raise THM ("add_arity_thms: theorem is not an arity", 0, [thm]);
155 in Theory.add_arities_i (map prep_thm thms) thy end;
161 (* data kind 'Pure/axclasses' *)
164 {super_classes: class list,
168 structure AxclassesArgs =
170 val name = "Pure/axclasses";
171 type T = axclass_info Symtab.table;
173 val empty = Symtab.empty;
176 fun merge (tab1, tab2) = Symtab.merge (K true) (tab1, tab2);
180 val ext_class = Sign.cond_extern sg Sign.classK;
181 val ext_thm = PureThy.cond_extern_thm_sg sg;
183 fun pretty_class c cs = Pretty.block
184 (Pretty.str (ext_class c) :: Pretty.str " <" :: Pretty.brk 1 ::
185 Pretty.breaks (map (Pretty.str o ext_class) cs));
187 fun pretty_thms name thms = Pretty.big_list (name ^ ":")
188 (map (Display.pretty_thm_sg sg) thms);
190 fun pretty_axclass (name, {super_classes, intro, axioms}) = Pretty.block (Pretty.fbreaks
191 [pretty_class name super_classes, pretty_thms introN [intro], pretty_thms axiomsN axioms]);
192 in Pretty.writeln (Pretty.chunks (map pretty_axclass (Symtab.dest tab))) end;
195 structure AxclassesData = TheoryDataFun(AxclassesArgs);
196 val print_axclasses = AxclassesData.print;
199 (* get and put data *)
201 fun lookup_axclass_info_sg sg c = Symtab.lookup (AxclassesData.get_sg sg, c);
203 fun get_axclass_info thy c =
204 (case lookup_axclass_info_sg (Theory.sign_of thy) c of
205 None => error ("Unknown axclass " ^ quote c)
206 | Some info => info);
208 fun put_axclass_info c info thy =
209 thy |> AxclassesData.put (Symtab.update ((c, info), AxclassesData.get thy));
213 (** add axiomatic type classes **)
217 fun err_not_logic c =
218 error ("Axiomatic class " ^ quote c ^ " not subclass of " ^ quote logicC);
220 fun err_bad_axsort ax c =
221 error ("Sort constraint in axiom " ^ quote ax ^ " not supersort of " ^ quote c);
223 fun err_bad_tfrees ax =
224 error ("More than one type variable in axiom " ^ quote ax);
229 fun ext_axclass prep_class prep_axm prep_att (bclass, raw_super_classes) raw_axioms_atts thy =
231 val sign = Theory.sign_of thy;
233 val class = Sign.full_name sign bclass;
234 val super_classes = map (prep_class sign) raw_super_classes;
235 val axms = map (prep_axm sign o fst) raw_axioms_atts;
236 val atts = map (map (prep_att thy) o snd) raw_axioms_atts;
240 thy |> Theory.add_classes_i [(bclass, super_classes)];
241 val class_sign = Theory.sign_of class_thy;
243 (*prepare abstract axioms*)
245 if null (term_tfrees ax) then
246 Logic.mk_implies (Logic.mk_inclass (aT logicS, class), ax)
247 else map_term_tfrees (K (aT [class])) ax;
248 val abs_axms = map (abs_axm o #2) axms;
250 (*prepare introduction rule*)
251 val _ = if Sign.subsort class_sign ([class], logicS) then () else err_not_logic class;
253 fun axm_sort (name, ax) =
254 (case term_tfrees ax of
256 | [(_, S)] => if Sign.subsort class_sign ([class], S) then S else err_bad_axsort name class
257 | _ => err_bad_tfrees name);
258 val axS = Sign.norm_sort class_sign (logicC :: flat (map axm_sort axms));
260 val int_axm = Logic.close_form o map_term_tfrees (K (aT axS));
261 fun inclass c = Logic.mk_inclass (aT axS, c);
263 val intro_axm = Logic.list_implies
264 (map inclass super_classes @ map (int_axm o #2) axms, inclass class);
266 (*declare axioms and rule*)
267 val (axms_thy, ([intro], [axioms])) =
269 |> Theory.add_path bclass
270 |> PureThy.add_axioms_i [Thm.no_attributes (introN, intro_axm)]
271 |>>> PureThy.add_axiomss_i [Thm.no_attributes (axiomsN, abs_axms)];
272 val info = {super_classes = super_classes, intro = intro, axioms = axioms};
277 |> (#1 o PureThy.add_thms ((map #1 axms ~~ axioms) ~~ atts))
278 |> Theory.parent_path
279 |> (#1 o PureThy.add_thms [Thm.no_attributes (intro_name bclass, intro)])
280 |> put_axclass_info class info;
281 in (final_thy, {intro = intro, axioms = axioms}) end;
284 (* external interfaces *)
286 val add_axclass = ext_axclass Sign.intern_class Theory.read_axm Attrib.global_attribute;
287 val add_axclass_i = ext_axclass (K I) Theory.cert_axm (K I);
291 (** prove class relations and type arities **)
295 fun class_axms sign =
296 let val classes = Sign.classes sign in
297 map (Thm.class_triv sign) classes @
298 mapfilter (apsome #intro o lookup_axclass_info_sg sign) classes
304 fun intro_classes_tac facts i st =
305 ((Method.insert_tac facts THEN'
306 REPEAT_ALL_NEW (resolve_tac (class_axms (Thm.sign_of_thm st)))) i
307 THEN Tactic.distinct_subgoals_tac) st; (*affects all subgoals!?*)
309 val intro_classes_method =
310 ("intro_classes", Method.no_args (Method.METHOD (HEADGOAL o intro_classes_tac)),
311 "back-chain introduction rules of axiomatic type classes");
316 fun default_intro_classes_tac [] i = intro_classes_tac [] i
317 | default_intro_classes_tac _ _ = Tactical.no_tac; (*no error message!*)
319 fun default_tac rules ctxt facts =
320 HEADGOAL (Method.some_rule_tac rules ctxt facts ORELSE' default_intro_classes_tac facts);
323 ("default", Method.thms_ctxt_args (Method.METHOD oo default_tac), "apply some intro/elim rule");
328 (*(1) repeatedly resolve goals of form "OFCLASS(ty, c_class)",
329 try class_trivs first, then "cI" axioms
330 (2) rewrite goals using user supplied definitions
331 (3) repeatedly resolve goals with user supplied non-definitions*)
333 fun axclass_tac thms =
335 val defs = filter is_def thms;
336 val non_defs = filter_out is_def thms;
338 HEADGOAL (intro_classes_tac []) THEN
339 TRY (rewrite_goals_tac defs) THEN
340 TRY (REPEAT_FIRST (fn i => assume_tac i ORELSE resolve_tac non_defs i))
346 fun prove mk_prop str_of sign prop thms usr_tac =
347 (Tactic.prove sign [] [] (mk_prop prop) (K (axclass_tac thms THEN (if_none usr_tac all_tac)))
348 handle ERROR => error ("The error(s) above occurred while trying to prove " ^
349 quote (str_of sign prop))) |> Drule.standard;
352 prove mk_classrel (fn sg => fn c1_c2 => Sign.str_of_classrel sg c1_c2);
355 prove mk_arity (fn sg => fn (t, Ss, c) => Sign.str_of_arity sg (t, Ss, [c]));
359 (** add proved subclass relations and arities **)
361 (* prepare classes and arities *)
363 fun read_class sg c = Sign.certify_class sg (Sign.intern_class sg c);
365 fun cert_classrel sg cc = Library.pairself (Sign.certify_class sg) cc;
366 fun read_classrel sg cc = Library.pairself (read_class sg) cc;
368 fun check_tycon sg t =
369 let val {tycons, abbrs, ...} = Type.rep_tsig (Sign.tsig_of sg) in
370 if is_some (Symtab.lookup (abbrs, t)) then
371 error ("Illegal type abbreviation: " ^ quote t)
372 else if is_none (Symtab.lookup (tycons, t)) then
373 error ("Undeclared type constructor: " ^ quote t)
377 fun prep_arity prep_tycon prep_sort prep_x sg (t, Ss, x) =
378 (check_tycon sg (prep_tycon sg t), map (prep_sort sg) Ss, prep_x sg x);
380 val read_arity = prep_arity Sign.intern_tycon Sign.read_sort Sign.read_sort;
381 val cert_arity = prep_arity (K I) Sign.certify_sort Sign.certify_sort;
382 val read_simple_arity = prep_arity Sign.intern_tycon Sign.read_sort Sign.intern_class;
383 fun cert_simple_arity arg = prep_arity (K I) Sign.certify_sort (K I) arg;
386 (* old-style instance declarations *)
388 fun ext_inst_subclass prep_classrel raw_c1_c2 names thms usr_tac thy =
390 val sign = Theory.sign_of thy;
391 val c1_c2 = prep_classrel sign raw_c1_c2;
393 message ("Proving class inclusion " ^ quote (Sign.str_of_classrel sign c1_c2) ^ " ...");
394 thy |> add_classrel_thms [prove_subclass sign c1_c2 (witnesses thy names thms) usr_tac]
397 fun ext_inst_arity prep_arity (raw_t, raw_Ss, raw_cs) names thms usr_tac thy =
399 val sign = Theory.sign_of thy;
400 val (t, Ss, cs) = prep_arity sign (raw_t, raw_Ss, raw_cs);
401 val wthms = witnesses thy names thms;
403 (message ("Proving type arity " ^
404 quote (Sign.str_of_arity sign (t, Ss, [c])) ^ " ...");
405 prove_arity sign (t, Ss, c) wthms usr_tac);
406 in add_arity_thms (map prove cs) thy end;
408 fun sane_inst_subclass prep sub tac = ext_inst_subclass prep sub [] [] (Some tac);
409 fun sane_inst_arity prep arity tac = ext_inst_arity prep arity [] [] (Some tac);
411 val add_inst_subclass_x = ext_inst_subclass read_classrel;
412 val add_inst_subclass = sane_inst_subclass read_classrel;
413 val add_inst_subclass_i = sane_inst_subclass cert_classrel;
414 val add_inst_arity_x = ext_inst_arity read_arity;
415 val add_inst_arity = sane_inst_arity read_arity;
416 val add_inst_arity_i = sane_inst_arity cert_arity;
420 (** instance proof interfaces **)
422 fun inst_attribute add_thms (thy, thm) = (add_thms [thm] thy, thm);
424 fun inst_proof mk_prop add_thms (sig_prop, comment) int thy =
426 |> IsarThy.theorem_i Drule.internalK (None, []) ((("", [inst_attribute add_thms]),
427 (mk_prop (Theory.sign_of thy) sig_prop, ([], []))), comment) int;
429 val instance_subclass_proof = inst_proof (mk_classrel oo read_classrel) add_classrel_thms;
430 val instance_subclass_proof_i = inst_proof (mk_classrel oo cert_classrel) add_classrel_thms;
431 val instance_arity_proof = inst_proof (mk_arity oo read_simple_arity) add_arity_thms;
432 val instance_arity_proof_i = inst_proof (mk_arity oo cert_simple_arity) add_arity_thms;
436 (** package setup **)
442 Method.add_methods [intro_classes_method, default_method]];
447 local structure P = OuterParse and K = OuterSyntax.Keyword in
450 OuterSyntax.command "axclass" "define axiomatic type class" K.thy_decl
451 (((P.name -- Scan.optional ((P.$$$ "\\<subseteq>" || P.$$$ "<") |--
452 P.!!! (P.list1 P.xname)) []) --| P.marg_comment)
453 -- Scan.repeat (P.spec_name --| P.marg_comment)
454 >> (fn (cls, axs) => Toplevel.theory (#1 o add_axclass cls axs)));
457 OuterSyntax.command "instance" "prove type arity or subclass relation" K.thy_goal
458 ((P.xname -- ((P.$$$ "\\<subseteq>" || P.$$$ "<") |-- P.xname)
459 -- P.marg_comment >> instance_subclass_proof ||
460 (P.xname -- (P.$$$ "::" |-- P.simple_arity) >> P.triple2) -- P.marg_comment
461 >> instance_arity_proof)
462 >> (Toplevel.print oo Toplevel.theory_to_proof));
464 val _ = OuterSyntax.add_parsers [axclassP, instanceP];