1 (* Title: Pure/Isar/class_target.ML
2 Author: Florian Haftmann, TU Muenchen
4 Type classes derived from primitive axclasses and locales - mechanisms.
7 signature CLASS_TARGET =
10 val register: class -> class list -> ((string * typ) * (string * typ)) list
11 -> sort -> morphism -> thm option -> thm option -> thm
13 val register_subclass: class * class -> morphism option -> Element.witness option
14 -> morphism -> theory -> theory
16 val is_class: theory -> class -> bool
17 val base_sort: theory -> class -> sort
18 val rules: theory -> class -> thm option * thm
19 val these_params: theory -> sort -> (string * (class * (string * typ))) list
20 val these_defs: theory -> sort -> thm list
21 val these_operations: theory -> sort -> (string * (class * (typ * term))) list
22 val print_classes: theory -> unit
24 val begin: class list -> sort -> Proof.context -> Proof.context
25 val init: class -> theory -> Proof.context
26 val declare: class -> Properties.T
27 -> (binding * mixfix) * term -> theory -> theory
28 val abbrev: class -> Syntax.mode -> Properties.T
29 -> (binding * mixfix) * term -> theory -> theory
30 val class_prefix: string -> string
31 val refresh_syntax: class -> Proof.context -> Proof.context
32 val redeclare_operations: theory -> sort -> Proof.context -> Proof.context
35 val init_instantiation: string list * (string * sort) list * sort
36 -> theory -> local_theory
37 val instantiation_instance: (local_theory -> local_theory)
38 -> local_theory -> Proof.state
39 val prove_instantiation_instance: (Proof.context -> tactic)
40 -> local_theory -> local_theory
41 val prove_instantiation_exit: (Proof.context -> tactic)
42 -> local_theory -> theory
43 val prove_instantiation_exit_result: (morphism -> 'a -> 'b)
44 -> (Proof.context -> 'b -> tactic) -> 'a -> local_theory -> 'b * theory
45 val conclude_instantiation: local_theory -> local_theory
46 val instantiation_param: local_theory -> binding -> string option
47 val confirm_declaration: binding -> local_theory -> local_theory
48 val pretty_instantiation: local_theory -> Pretty.T
49 val type_name: string -> string
51 val intro_classes_tac: thm list -> tactic
52 val default_intro_tac: Proof.context -> thm list -> tactic
55 val axclass_cmd: binding * xstring list
56 -> (Attrib.binding * string list) list
57 -> theory -> class * theory
58 val classrel_cmd: xstring * xstring -> theory -> Proof.state
59 val instance_arity: (theory -> theory) -> arity -> theory -> Proof.state
60 val instance_arity_cmd: xstring * xstring list * xstring -> theory -> Proof.state
63 structure Class_Target : CLASS_TARGET =
66 (** primitive axclass and instance commands **)
68 fun axclass_cmd (class, raw_superclasses) raw_specs thy =
70 val ctxt = ProofContext.init thy;
71 val superclasses = map (Sign.read_class thy) raw_superclasses;
72 val name_atts = map ((apsnd o map) (Attrib.attribute thy) o fst)
74 val axiomss = ProofContext.read_propp (ctxt, map (map (rpair []) o snd)
79 AxClass.define_class (class, superclasses) []
80 (name_atts ~~ axiomss) thy
85 fun gen_instance mk_prop add_thm after_qed insts thy =
87 fun after_qed' results =
88 ProofContext.theory ((fold o fold) add_thm results #> after_qed);
92 |> Proof.theorem_i NONE after_qed' ((map (fn t => [(t, [])])
99 gen_instance (Logic.mk_arities oo Sign.cert_arity) AxClass.add_arity;
100 val instance_arity_cmd =
101 gen_instance (Logic.mk_arities oo Sign.read_arity) AxClass.add_arity I;
103 gen_instance (single oo (Logic.mk_classrel oo AxClass.cert_classrel)) AxClass.add_classrel I;
105 gen_instance (single oo (Logic.mk_classrel oo AxClass.read_classrel)) AxClass.add_classrel I;
112 datatype class_data = ClassData of {
115 consts: (string * string) list
116 (*locale parameter ~> constant name*),
119 (*static part of canonical morphism*),
120 assm_intro: thm option,
126 operations: (string * (class * (typ * term))) list
130 fun rep_class_data (ClassData data) = data;
131 fun mk_class_data ((consts, base_sort, base_morph, assm_intro, of_class, axiom),
132 (defs, operations)) =
133 ClassData { consts = consts, base_sort = base_sort,
134 base_morph = base_morph, assm_intro = assm_intro, of_class = of_class, axiom = axiom,
135 defs = defs, operations = operations };
136 fun map_class_data f (ClassData { consts, base_sort, base_morph, assm_intro,
137 of_class, axiom, defs, operations }) =
138 mk_class_data (f ((consts, base_sort, base_morph, assm_intro, of_class, axiom),
139 (defs, operations)));
140 fun merge_class_data _ (ClassData { consts = consts,
141 base_sort = base_sort, base_morph = base_morph, assm_intro = assm_intro,
142 of_class = of_class, axiom = axiom, defs = defs1, operations = operations1 },
143 ClassData { consts = _, base_sort = _, base_morph = _, assm_intro = _,
144 of_class = _, axiom = _, defs = defs2, operations = operations2 }) =
145 mk_class_data ((consts, base_sort, base_morph, assm_intro, of_class, axiom),
146 (Thm.merge_thms (defs1, defs2),
147 AList.merge (op =) (K true) (operations1, operations2)));
149 structure ClassData = TheoryDataFun
151 type T = class_data Graph.T
152 val empty = Graph.empty;
155 fun merge _ = Graph.join merge_class_data;
161 val lookup_class_data = Option.map rep_class_data oo try o Graph.get_node o ClassData.get;
163 fun the_class_data thy class = case lookup_class_data thy class
164 of NONE => error ("Undeclared class " ^ quote class)
167 val is_class = is_some oo lookup_class_data;
169 val ancestry = Graph.all_succs o ClassData.get;
170 val heritage = Graph.all_preds o ClassData.get;
172 fun these_params thy =
176 val const_typs = (#params o AxClass.get_info thy) class;
177 val const_names = (#consts o the_class_data thy) class;
180 (fn c => (class, (c, (the o AList.lookup (op =) const_typs) c))) const_names
182 in maps params o ancestry thy end;
184 val base_sort = #base_sort oo the_class_data;
186 fun rules thy class =
187 let val { axiom, of_class, ... } = the_class_data thy class
188 in (axiom, of_class) end;
190 fun all_assm_intros thy =
191 Graph.fold (fn (_, (data, _)) => fold (insert Thm.eq_thm)
192 ((the_list o #assm_intro o rep_class_data) data)) (ClassData.get thy) [];
194 fun these_defs thy = maps (#defs o the_class_data thy) o ancestry thy;
195 fun these_operations thy = maps (#operations o the_class_data thy) o ancestry thy;
197 val base_morphism = #base_morph oo the_class_data;
198 fun morphism thy class = base_morphism thy class
199 $> Element.eq_morphism thy (these_defs thy [class]);
201 fun print_classes thy =
203 val ctxt = ProofContext.init thy;
204 val algebra = Sign.classes_of thy;
207 |> Symtab.fold (fn (tyco, arities) => fold (fn (class, _) =>
208 Symtab.map_default (class, []) (insert (op =) tyco)) arities)
209 ((#arities o Sorts.rep_algebra) algebra);
210 val the_arities = these o Symtab.lookup arities;
211 fun mk_arity class tyco =
213 val Ss = Sorts.mg_domain algebra tyco [class];
214 in Syntax.pretty_arity ctxt (tyco, Ss, [class]) end;
215 fun mk_param (c, ty) = Pretty.str (Sign.extern_const thy c ^ " :: "
216 ^ setmp show_sorts false (Syntax.string_of_typ ctxt o Type.strip_sorts) ty);
217 fun mk_entry class = (Pretty.block o Pretty.fbreaks o map_filter I) [
218 (SOME o Pretty.str) ("class " ^ Sign.extern_class thy class ^ ":"),
219 (SOME o Pretty.block) [Pretty.str "supersort: ",
220 (Syntax.pretty_sort ctxt o Sign.minimize_sort thy o Sign.super_classes thy) class],
221 ((fn [] => NONE | ps => (SOME o Pretty.block o Pretty.fbreaks)
222 (Pretty.str "parameters:" :: ps)) o map mk_param
223 o these o Option.map #params o try (AxClass.get_info thy)) class,
224 (SOME o Pretty.block o Pretty.breaks) [
225 Pretty.str "instances:",
226 Pretty.list "" "" (map (mk_arity class) (the_arities class))
230 (Pretty.writeln o Pretty.chunks o separate (Pretty.str "")
231 o map mk_entry o Sorts.all_classes) algebra
237 fun register class sups params base_sort morph
238 axiom assm_intro of_class thy =
240 val operations = map (fn (v_ty as (_, ty), (c, _)) =>
241 (c, (class, (ty, Free v_ty)))) params;
242 val add_class = Graph.new_node (class,
243 mk_class_data (((map o pairself) fst params, base_sort,
244 morph, assm_intro, of_class, axiom), ([], operations)))
245 #> fold (curry Graph.add_edge class) sups;
246 in ClassData.map add_class thy end;
248 fun activate_defs class thms thy =
250 val eq_morph = Element.eq_morphism thy thms;
251 fun amend cls thy = Locale.amend_registration eq_morph
252 (cls, morphism thy cls) thy;
253 in fold amend (heritage thy [class]) thy end;
255 fun register_operation class (c, (t, some_def)) thy =
257 val base_sort = base_sort thy class;
258 val prep_typ = map_type_tfree
259 (fn (v, sort) => if Name.aT = v
260 then TFree (v, base_sort) else TVar ((v, 0), sort));
261 val t' = map_types prep_typ t;
262 val ty' = Term.fastype_of t';
265 |> (ClassData.map o Graph.map_node class o map_class_data o apsnd)
266 (fn (defs, operations) =>
267 (fold cons (the_list some_def) defs,
268 (c, (class, (ty', t'))) :: operations))
269 |> is_some some_def ? activate_defs class (the_list some_def)
272 fun register_subclass (sub, sup) some_dep_morph some_wit export thy =
274 val intros = (snd o rules thy) sup :: map_filter I
275 [Option.map (Drule.standard' o Element.conclude_witness) some_wit,
276 (fst o rules thy) sub];
277 val tac = EVERY (map (TRYALL o Tactic.rtac) intros);
278 val classrel = Goal.prove_global thy [] [] (Logic.mk_classrel (sub, sup))
280 val diff_sort = Sign.complete_sort thy [sup]
281 |> subtract (op =) (Sign.complete_sort thy [sub])
282 |> filter (is_class thy);
285 |> AxClass.add_classrel classrel
286 |> ClassData.map (Graph.add_edge (sub, sup))
287 |> activate_defs sub (these_defs thy diff_sort)
288 |> fold (fn dep_morph => Locale.add_dependency sub (sup,
289 dep_morph $> Element.satisfy_morphism (the_list some_wit) $> export))
290 (the_list some_dep_morph)
291 |> (fn thy => fold_rev (Context.theory_map o Locale.activate_facts)
292 (Locale.get_registrations thy) thy)
296 (** classes and class target **)
298 (* class context syntax *)
300 fun these_unchecks thy =
301 map (fn (c, (_, (ty, t))) => (t, Const (c, ty))) o these_operations thy;
303 fun redeclare_const thy c =
304 let val b = Long_Name.base_name c
305 in Sign.intern_const thy b = c ? Variable.declare_const (b, c) end;
307 fun synchronize_class_syntax sort base_sort ctxt =
309 val thy = ProofContext.theory_of ctxt;
310 val algebra = Sign.classes_of thy;
311 val operations = these_operations thy sort;
312 fun subst_class_typ sort = map_type_tfree (K (TVar ((Name.aT, 0), sort)));
313 val primary_constraints =
314 (map o apsnd) (subst_class_typ base_sort o fst o snd) operations;
315 val secondary_constraints =
316 (map o apsnd) (fn (class, (ty, _)) => subst_class_typ [class] ty) operations;
317 fun improve (c, ty) = (case AList.lookup (op =) primary_constraints c
318 of SOME ty' => (case try (Type.raw_match (ty', ty)) Vartab.empty
319 of SOME tyenv => (case Vartab.lookup tyenv (Name.aT, 0)
320 of SOME (_, ty' as TVar (tvar as (vi, sort))) =>
321 if TypeInfer.is_param vi
322 andalso Sorts.sort_le algebra (base_sort, sort)
323 then SOME (ty', TFree (Name.aT, base_sort))
328 fun subst (c, ty) = Option.map snd (AList.lookup (op =) operations c);
329 val unchecks = these_unchecks thy sort;
332 |> fold (redeclare_const thy o fst) primary_constraints
333 |> Overloading.map_improvable_syntax (K (((primary_constraints, secondary_constraints),
334 (((improve, subst), true), unchecks)), false))
335 |> Overloading.set_primary_constraints
338 fun refresh_syntax class ctxt =
340 val thy = ProofContext.theory_of ctxt;
341 val base_sort = base_sort thy class;
342 in synchronize_class_syntax [class] base_sort ctxt end;
344 fun redeclare_operations thy sort =
345 fold (redeclare_const thy o fst) (these_operations thy sort);
347 fun begin sort base_sort ctxt =
349 |> Variable.declare_term
350 (Logic.mk_type (TFree (Name.aT, base_sort)))
351 |> synchronize_class_syntax sort base_sort
352 |> Overloading.add_improvable_syntax;
357 |> begin [class] (base_sort thy class);
362 val class_prefix = Logic.const_of_class o Long_Name.base_name;
364 fun declare class pos ((c, mx), dict) thy =
366 val morph = morphism thy class;
367 val b = Morphism.binding morph c;
368 val b_def = Morphism.binding morph (Binding.suffix_name "_dict" b);
369 val c' = Sign.full_name thy b;
370 val dict' = Morphism.term morph dict;
371 val ty' = Term.fastype_of dict';
372 val def_eq = Logic.mk_equals (Const (c', ty'), dict')
373 |> map_types Type.strip_sorts;
376 |> Sign.declare_const pos ((b, Type.strip_sorts ty'), mx)
378 |> Thm.add_def false false (b_def, def_eq)
380 |-> (fn def_thm => PureThy.store_thm (b_def, def_thm)
382 #> register_operation class (c', (dict', SOME (Thm.symmetric def_thm))))
383 |> Sign.add_const_constraint (c', SOME ty')
386 fun abbrev class prmode pos ((c, mx), rhs) thy =
388 val morph = morphism thy class;
389 val unchecks = these_unchecks thy [class];
390 val b = Morphism.binding morph c;
391 val c' = Sign.full_name thy b;
392 val rhs' = Pattern.rewrite_term thy unchecks [] rhs;
393 val ty' = Term.fastype_of rhs';
394 val rhs'' = map_types ((*Type.strip_sorts o *)Logic.varifyT) rhs';
397 |> Sign.add_abbrev (#1 prmode) pos (b, rhs'')
399 |> Sign.add_const_constraint (c', SOME ty')
400 |> Sign.notation true prmode [(Const (c', ty'), mx)]
401 |> not (#1 prmode = PrintMode.input) ? register_operation class (c', (rhs', NONE))
405 (** instantiation target **)
409 datatype instantiation = Instantiation of {
410 arities: string list * (string * sort) list * sort,
411 params: ((string * string) * (string * typ)) list
412 (*(instantiation parameter, type constructor), (local instantiation parameter, typ)*)
415 structure Instantiation = ProofDataFun
417 type T = instantiation
418 fun init _ = Instantiation { arities = ([], [], []), params = [] };
421 fun mk_instantiation (arities, params) =
422 Instantiation { arities = arities, params = params };
423 fun get_instantiation lthy = case Instantiation.get (LocalTheory.target_of lthy)
424 of Instantiation data => data;
425 fun map_instantiation f = (LocalTheory.target o Instantiation.map)
426 (fn Instantiation { arities, params } => mk_instantiation (f (arities, params)));
428 fun the_instantiation lthy = case get_instantiation lthy
429 of { arities = ([], [], []), ... } => error "No instantiation target"
432 val instantiation_params = #params o get_instantiation;
434 fun instantiation_param lthy b = instantiation_params lthy
435 |> find_first (fn (_, (v, _)) => Binding.name_of b = v)
436 |> Option.map (fst o fst);
441 fun synchronize_inst_syntax ctxt =
443 val Instantiation { arities = (_, _, sort), params = params } = Instantiation.get ctxt;
444 val inst_tyco_of = AxClass.inst_tyco_of (Sign.consts_of (ProofContext.theory_of ctxt));
445 fun subst (c, ty) = case inst_tyco_of (c, ty)
446 of SOME tyco => (case AList.lookup (op =) params (c, tyco)
447 of SOME (v_ty as (_, ty)) => SOME (ty, Free v_ty)
451 map (fn ((c, _), v_ty as (_, ty)) => (Free v_ty, Const (c, ty))) params;
454 |> Overloading.map_improvable_syntax
455 (fn (((primary_constraints, _), (((improve, _), _), _)), _) =>
456 (((primary_constraints, []), (((improve, subst), false), unchecks)), false))
462 val sanatize_name = (*FIXME*)
464 fun is_valid s = Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s
465 orelse s = "'" orelse s = "_";
466 val is_junk = not o is_valid andf Symbol.is_regular;
467 val junk = Scan.many is_junk;
468 val scan_valids = Symbol.scanner "Malformed input"
470 (Scan.optional (Scan.one Symbol.is_ascii_letter) "x" ^^ (Scan.many is_valid >> implode)
472 ::: Scan.repeat ((Scan.many1 is_valid >> implode) --| junk));
474 explode #> scan_valids #> implode
477 fun type_name "*" = "prod"
478 | type_name "+" = "sum"
479 | type_name s = sanatize_name (Long_Name.base_name s);
481 fun resort_terms pp algebra consts constraints ts =
483 fun matchings (Const (c_ty as (c, _))) = (case constraints c
485 | SOME sorts => fold2 (curry (Sorts.meet_sort algebra))
486 (Consts.typargs consts c_ty) sorts)
488 val tvartab = (fold o fold_aterms) matchings ts Vartab.empty
489 handle Sorts.CLASS_ERROR e => error (Sorts.class_error pp e);
490 val inst = map_type_tvar
491 (fn (vi, sort) => TVar (vi, the_default sort (Vartab.lookup tvartab vi)));
492 in if Vartab.is_empty tvartab then NONE else SOME ((map o map_types) inst ts) end;
494 fun init_instantiation (tycos, vs, sort) thy =
496 val _ = if null tycos then error "At least one arity must be given" else ();
497 val params = these_params thy (filter (can (AxClass.get_info thy)) sort);
498 fun get_param tyco (param, (_, (c, ty))) =
499 if can (AxClass.param_of_inst thy) (c, tyco)
500 then NONE else SOME ((c, tyco),
501 (param ^ "_" ^ type_name tyco, map_atyps (K (Type (tyco, map TFree vs))) ty));
502 val inst_params = map_product get_param tycos params |> map_filter I;
503 val primary_constraints = map (apsnd
504 (map_atyps (K (TVar ((Name.aT, 0), [])))) o snd o snd) params;
505 val pp = Syntax.pp_global thy;
506 val algebra = Sign.classes_of thy
507 |> fold (fn tyco => Sorts.add_arities pp
508 (tyco, map (fn class => (class, map snd vs)) sort)) tycos;
509 val consts = Sign.consts_of thy;
510 val improve_constraints = AList.lookup (op =)
511 (map (fn (_, (class, (c, _))) => (c, [[class]])) params);
512 fun resort_check ts ctxt = case resort_terms pp algebra consts improve_constraints ts
514 | SOME ts' => SOME (ts', ctxt);
515 val inst_tyco_of = AxClass.inst_tyco_of (Sign.consts_of thy);
516 val typ_instance = Type.typ_instance (Sign.tsig_of thy);
517 fun improve (c, ty) = case inst_tyco_of (c, ty)
518 of SOME tyco => (case AList.lookup (op =) inst_params (c, tyco)
519 of SOME (_, ty') => if typ_instance (ty', ty)
520 then SOME (ty, ty') else NONE
526 |> Instantiation.put (mk_instantiation ((tycos, vs, sort), inst_params))
527 |> fold (Variable.declare_typ o TFree) vs
528 |> fold (Variable.declare_names o Free o snd) inst_params
529 |> (Overloading.map_improvable_syntax o apfst)
530 (K ((primary_constraints, []), (((improve, K NONE), false), [])))
531 |> Overloading.add_improvable_syntax
532 |> Context.proof_map (Syntax.add_term_check 0 "resorting" resort_check)
533 |> synchronize_inst_syntax
536 fun confirm_declaration b = (map_instantiation o apsnd)
537 (filter_out (fn (_, (c', _)) => c' = Binding.name_of b))
538 #> LocalTheory.target synchronize_inst_syntax
540 fun gen_instantiation_instance do_proof after_qed lthy =
542 val (tycos, vs, sort) = (#arities o the_instantiation) lthy;
543 val arities_proof = maps (fn tyco => Logic.mk_arities (tyco, map snd vs, sort)) tycos;
544 fun after_qed' results =
545 LocalTheory.theory (fold (AxClass.add_arity o Thm.varifyT) results)
549 |> do_proof after_qed' arities_proof
552 val instantiation_instance = gen_instantiation_instance (fn after_qed => fn ts =>
553 Proof.theorem_i NONE (after_qed o map the_single) (map (fn t => [(t, [])]) ts));
555 fun prove_instantiation_instance tac = gen_instantiation_instance (fn after_qed =>
556 fn ts => fn lthy => after_qed (map (fn t => Goal.prove lthy [] [] t
557 (fn {context, ...} => tac context)) ts) lthy) I;
559 fun prove_instantiation_exit tac = prove_instantiation_instance tac
560 #> LocalTheory.exit_global;
562 fun prove_instantiation_exit_result f tac x lthy =
564 val morph = ProofContext.export_morphism lthy
565 (ProofContext.init (ProofContext.theory_of lthy));
569 |> prove_instantiation_exit (fn ctxt => tac ctxt y)
573 fun conclude_instantiation lthy =
575 val { arities, params } = the_instantiation lthy;
576 val (tycos, vs, sort) = arities;
577 val thy = ProofContext.theory_of lthy;
578 val _ = map (fn tyco => if Sign.of_sort thy
579 (Type (tyco, map TFree vs), sort)
580 then () else error ("Missing instance proof for type " ^ quote (Sign.extern_type thy tyco)))
584 fun pretty_instantiation lthy =
586 val { arities, params } = the_instantiation lthy;
587 val (tycos, vs, sort) = arities;
588 val thy = ProofContext.theory_of lthy;
589 fun pr_arity tyco = Syntax.pretty_arity lthy (tyco, map snd vs, sort);
590 fun pr_param ((c, _), (v, ty)) =
591 (Pretty.block o Pretty.breaks) [Pretty.str v, Pretty.str "==",
592 (Pretty.str o Sign.extern_const thy) c, Pretty.str "::", Syntax.pretty_typ_global thy ty];
594 (Pretty.block o Pretty.fbreaks)
595 (Pretty.str "instantiation" :: map pr_arity tycos @ map pr_param params)
599 (** tactics and methods **)
601 fun intro_classes_tac facts st =
603 val thy = Thm.theory_of_thm st;
604 val classes = Sign.all_classes thy;
605 val class_trivs = map (Thm.class_triv thy) classes;
606 val class_intros = map_filter (try (#intro o AxClass.get_info thy)) classes;
607 val assm_intros = all_assm_intros thy;
609 Method.intros_tac (class_trivs @ class_intros @ assm_intros) facts st
612 fun default_intro_tac ctxt [] =
613 intro_classes_tac [] ORELSE Locale.intro_locales_tac true ctxt []
614 | default_intro_tac _ _ = no_tac;
616 fun default_tac rules ctxt facts =
617 HEADGOAL (Method.some_rule_tac rules ctxt facts) ORELSE
618 default_intro_tac ctxt facts;
620 val _ = Context.>> (Context.map_theory
621 (Method.setup (Binding.name "intro_classes") (Scan.succeed (K (METHOD intro_classes_tac)))
622 "back-chain introduction rules of classes" #>
623 Method.setup (Binding.name "default") (Attrib.thms >> (METHOD oo default_tac))
624 "apply some intro/elim rule"));