Isar "default" step needs to fail for solved problems, for clear distinction of '.' and '..' for example -- amending lapse introduced in 9de4d64eee3b (April 2004);
1 (* Title: Pure/Isar/class.ML
2 Author: Florian Haftmann, TU Muenchen
4 Type classes derived from primitive axclasses and locales.
10 val is_class: theory -> class -> bool
11 val these_params: theory -> sort -> (string * (class * (string * typ))) list
12 val base_sort: theory -> class -> sort
13 val rules: theory -> class -> thm option * thm
14 val these_defs: theory -> sort -> thm list
15 val these_operations: theory -> sort
16 -> (string * (class * (typ * term))) list
17 val print_classes: theory -> unit
18 val init: class -> theory -> Proof.context
19 val begin: class list -> sort -> Proof.context -> Proof.context
20 val const: class -> (binding * mixfix) * (term list * term) -> local_theory -> local_theory
21 val abbrev: class -> Syntax.mode -> (binding * mixfix) * term -> local_theory -> local_theory
22 val redeclare_operations: theory -> sort -> Proof.context -> Proof.context
23 val class_prefix: string -> string
24 val register: class -> class list -> ((string * typ) * (string * typ)) list
25 -> sort -> morphism -> morphism -> thm option -> thm option -> thm
29 val instantiation: string list * (string * sort) list * sort -> theory -> local_theory
30 val instantiation_instance: (local_theory -> local_theory)
31 -> local_theory -> Proof.state
32 val prove_instantiation_instance: (Proof.context -> tactic)
33 -> local_theory -> local_theory
34 val prove_instantiation_exit: (Proof.context -> tactic)
35 -> local_theory -> theory
36 val prove_instantiation_exit_result: (morphism -> 'a -> 'b)
37 -> (Proof.context -> 'b -> tactic) -> 'a -> local_theory -> 'b * theory
38 val read_multi_arity: theory -> xstring list * xstring list * xstring
39 -> string list * (string * sort) list * sort
40 val type_name: string -> string
41 val instantiation_cmd: xstring list * xstring list * xstring -> theory -> local_theory
42 val instance_arity_cmd: xstring list * xstring list * xstring -> theory -> Proof.state
45 val classrel: class * class -> theory -> Proof.state
46 val classrel_cmd: xstring * xstring -> theory -> Proof.state
47 val register_subclass: class * class -> morphism option -> Element.witness option
48 -> morphism -> theory -> theory
51 val intro_classes_tac: thm list -> tactic
52 val default_intro_tac: Proof.context -> thm list -> tactic
55 structure Class: CLASS =
60 datatype class_data = ClassData of {
63 consts: (string * string) list
64 (*locale parameter ~> constant name*),
67 (*static part of canonical morphism*),
68 export_morph: morphism,
69 assm_intro: thm option,
75 operations: (string * (class * (typ * term))) list
79 fun make_class_data ((consts, base_sort, base_morph, export_morph, assm_intro, of_class, axiom),
81 ClassData { consts = consts, base_sort = base_sort,
82 base_morph = base_morph, export_morph = export_morph, assm_intro = assm_intro,
83 of_class = of_class, axiom = axiom, defs = defs, operations = operations };
84 fun map_class_data f (ClassData { consts, base_sort, base_morph, export_morph, assm_intro,
85 of_class, axiom, defs, operations }) =
86 make_class_data (f ((consts, base_sort, base_morph, export_morph, assm_intro, of_class, axiom),
88 fun merge_class_data _ (ClassData { consts = consts,
89 base_sort = base_sort, base_morph = base_morph, export_morph = export_morph, assm_intro = assm_intro,
90 of_class = of_class, axiom = axiom, defs = defs1, operations = operations1 },
91 ClassData { consts = _, base_sort = _, base_morph = _, export_morph = _, assm_intro = _,
92 of_class = _, axiom = _, defs = defs2, operations = operations2 }) =
93 make_class_data ((consts, base_sort, base_morph, export_morph, assm_intro, of_class, axiom),
94 (Thm.merge_thms (defs1, defs2),
95 AList.merge (op =) (K true) (operations1, operations2)));
97 structure ClassData = Theory_Data
99 type T = class_data Graph.T
100 val empty = Graph.empty;
102 val merge = Graph.join merge_class_data;
108 fun lookup_class_data thy class = case try (Graph.get_node (ClassData.get thy)) class
109 of SOME (ClassData data) => SOME data
112 fun the_class_data thy class = case lookup_class_data thy class
113 of NONE => error ("Undeclared class " ^ quote class)
116 val is_class = is_some oo lookup_class_data;
118 val ancestry = Graph.all_succs o ClassData.get;
119 val heritage = Graph.all_preds o ClassData.get;
121 fun these_params thy =
125 val const_typs = (#params o AxClass.get_info thy) class;
126 val const_names = (#consts o the_class_data thy) class;
129 (fn c => (class, (c, (the o AList.lookup (op =) const_typs) c))) const_names
131 in maps params o ancestry thy end;
133 val base_sort = #base_sort oo the_class_data;
135 fun rules thy class =
136 let val { axiom, of_class, ... } = the_class_data thy class
137 in (axiom, of_class) end;
139 fun all_assm_intros thy =
140 Graph.fold (fn (_, (ClassData { assm_intro, ... }, _)) => fold (insert Thm.eq_thm)
141 (the_list assm_intro)) (ClassData.get thy) [];
143 fun these_defs thy = maps (#defs o the_class_data thy) o ancestry thy;
144 fun these_operations thy = maps (#operations o the_class_data thy) o ancestry thy;
146 val base_morphism = #base_morph oo the_class_data;
147 fun morphism thy class = case Element.eq_morphism thy (these_defs thy [class])
148 of SOME eq_morph => base_morphism thy class $> eq_morph
149 | NONE => base_morphism thy class;
150 val export_morphism = #export_morph oo the_class_data;
152 fun print_classes thy =
154 val ctxt = ProofContext.init_global thy;
155 val algebra = Sign.classes_of thy;
158 |> Symtab.fold (fn (tyco, arities) => fold (fn (class, _) =>
159 Symtab.map_default (class, []) (insert (op =) tyco)) arities)
160 (Sorts.arities_of algebra);
161 val the_arities = these o Symtab.lookup arities;
162 fun mk_arity class tyco =
164 val Ss = Sorts.mg_domain algebra tyco [class];
165 in Syntax.pretty_arity ctxt (tyco, Ss, [class]) end;
166 fun mk_param (c, ty) = Pretty.str (Sign.extern_const thy c ^ " :: "
167 ^ (Syntax.string_of_typ (Config.put show_sorts false ctxt) o Type.strip_sorts) ty);
168 fun mk_entry class = (Pretty.block o Pretty.fbreaks o map_filter I) [
169 (SOME o Pretty.str) ("class " ^ Sign.extern_class thy class ^ ":"),
170 (SOME o Pretty.block) [Pretty.str "supersort: ",
171 (Syntax.pretty_sort ctxt o Sign.minimize_sort thy o Sign.super_classes thy) class],
172 ((fn [] => NONE | ps => (SOME o Pretty.block o Pretty.fbreaks)
173 (Pretty.str "parameters:" :: ps)) o map mk_param
174 o these o Option.map #params o try (AxClass.get_info thy)) class,
175 (SOME o Pretty.block o Pretty.breaks) [
176 Pretty.str "instances:",
177 Pretty.list "" "" (map (mk_arity class) (the_arities class))
181 (Pretty.writeln o Pretty.chunks o separate (Pretty.str "")
182 o map mk_entry o Sorts.all_classes) algebra
188 fun register class sups params base_sort base_morph export_morph
189 axiom assm_intro of_class thy =
191 val operations = map (fn (v_ty as (_, ty), (c, _)) =>
192 (c, (class, (ty, Free v_ty)))) params;
193 val add_class = Graph.new_node (class,
194 make_class_data (((map o pairself) fst params, base_sort,
195 base_morph, export_morph, assm_intro, of_class, axiom), ([], operations)))
196 #> fold (curry Graph.add_edge class) sups;
197 in ClassData.map add_class thy end;
199 fun activate_defs class thms thy = case Element.eq_morphism thy thms
200 of SOME eq_morph => fold (fn cls => fn thy =>
201 Context.theory_map (Locale.amend_registration (cls, base_morphism thy cls)
202 (eq_morph, true) (export_morphism thy cls)) thy) (heritage thy [class]) thy
205 fun register_operation class (c, (t, some_def)) thy =
207 val base_sort = base_sort thy class;
208 val prep_typ = map_type_tfree
209 (fn (v, sort) => if Name.aT = v
210 then TFree (v, base_sort) else TVar ((v, 0), sort));
211 val t' = map_types prep_typ t;
212 val ty' = Term.fastype_of t';
215 |> (ClassData.map o Graph.map_node class o map_class_data o apsnd)
216 (fn (defs, operations) =>
217 (fold cons (the_list some_def) defs,
218 (c, (class, (ty', t'))) :: operations))
219 |> activate_defs class (the_list some_def)
222 fun register_subclass (sub, sup) some_dep_morph some_wit export thy =
224 val intros = (snd o rules thy) sup :: map_filter I
225 [Option.map (Drule.export_without_context_open o Element.conclude_witness) some_wit,
226 (fst o rules thy) sub];
227 val tac = EVERY (map (TRYALL o Tactic.rtac) intros);
228 val classrel = Skip_Proof.prove_global thy [] [] (Logic.mk_classrel (sub, sup))
230 val diff_sort = Sign.complete_sort thy [sup]
231 |> subtract (op =) (Sign.complete_sort thy [sub])
232 |> filter (is_class thy);
233 val add_dependency = case some_dep_morph
234 of SOME dep_morph => Locale.add_dependency sub
235 (sup, dep_morph $> Element.satisfy_morphism (the_list some_wit)) export
239 |> AxClass.add_classrel classrel
240 |> ClassData.map (Graph.add_edge (sub, sup))
241 |> activate_defs sub (these_defs thy diff_sort)
246 (** classes and class target **)
248 (* class context syntax *)
250 fun these_unchecks thy = map (fn (c, (_, (ty, t))) => (t, Const (c, ty)))
251 o these_operations thy;
253 fun redeclare_const thy c =
254 let val b = Long_Name.base_name c
255 in Sign.intern_const thy b = c ? Variable.declare_const (b, c) end;
257 fun synchronize_class_syntax sort base_sort ctxt =
259 val thy = ProofContext.theory_of ctxt;
260 val algebra = Sign.classes_of thy;
261 val operations = these_operations thy sort;
262 fun subst_class_typ sort = map_type_tfree (K (TVar ((Name.aT, 0), sort)));
263 val primary_constraints =
264 (map o apsnd) (subst_class_typ base_sort o fst o snd) operations;
265 val secondary_constraints =
266 (map o apsnd) (fn (class, (ty, _)) => subst_class_typ [class] ty) operations;
267 fun improve (c, ty) = (case AList.lookup (op =) primary_constraints c
268 of SOME ty' => (case try (Type.raw_match (ty', ty)) Vartab.empty
269 of SOME tyenv => (case Vartab.lookup tyenv (Name.aT, 0)
270 of SOME (_, ty' as TVar (vi, sort)) =>
271 if Type_Infer.is_param vi
272 andalso Sorts.sort_le algebra (base_sort, sort)
273 then SOME (ty', TFree (Name.aT, base_sort))
278 fun subst (c, ty) = Option.map snd (AList.lookup (op =) operations c);
279 val unchecks = these_unchecks thy sort;
282 |> fold (redeclare_const thy o fst) primary_constraints
283 |> Overloading.map_improvable_syntax (K (((primary_constraints, secondary_constraints),
284 (((improve, subst), true), unchecks)), false))
285 |> Overloading.set_primary_constraints
288 fun redeclare_operations thy sort =
289 fold (redeclare_const thy o fst) (these_operations thy sort);
291 fun begin sort base_sort ctxt =
293 |> Variable.declare_term
294 (Logic.mk_type (TFree (Name.aT, base_sort)))
295 |> synchronize_class_syntax sort base_sort
296 |> Overloading.activate_improvable_syntax;
301 |> begin [class] (base_sort thy class);
306 val class_prefix = Logic.const_of_class o Long_Name.base_name;
308 fun target_extension f class lthy =
310 |> Local_Theory.raw_theory f
311 |> Local_Theory.target (synchronize_class_syntax [class]
312 (base_sort (ProofContext.theory_of lthy) class));
316 fun target_const class ((c, mx), (type_params, dict)) thy =
318 val morph = morphism thy class;
319 val b = Morphism.binding morph c;
320 val b_def = Morphism.binding morph (Binding.suffix_name "_dict" b);
321 val c' = Sign.full_name thy b;
322 val dict' = Morphism.term morph dict;
323 val ty' = map Term.fastype_of type_params ---> Term.fastype_of dict';
324 val def_eq = Logic.mk_equals (list_comb (Const (c', ty'), type_params), dict')
325 |> map_types Type.strip_sorts;
328 |> Sign.declare_const ((b, Type.strip_sorts ty'), mx)
330 |> Thm.add_def false false (b_def, def_eq)
331 |>> apsnd Thm.varifyT_global
332 |-> (fn (_, def_thm) => PureThy.store_thm (b_def, def_thm)
334 #> null type_params ? register_operation class (c', (dict', SOME (Thm.symmetric def_thm))))
335 |> Sign.add_const_constraint (c', SOME ty')
338 fun target_abbrev class prmode ((c, mx), rhs) thy =
340 val morph = morphism thy class;
341 val unchecks = these_unchecks thy [class];
342 val b = Morphism.binding morph c;
343 val c' = Sign.full_name thy b;
344 val rhs' = Pattern.rewrite_term thy unchecks [] rhs;
345 val ty' = Term.fastype_of rhs';
346 val rhs'' = map_types Logic.varifyT_global rhs';
349 |> Sign.add_abbrev (#1 prmode) (b, rhs'')
351 |> Sign.add_const_constraint (c', SOME ty')
352 |> Sign.notation true prmode [(Const (c', ty'), mx)]
353 |> not (#1 prmode = Print_Mode.input) ? register_operation class (c', (rhs', NONE))
358 fun const class arg = target_extension (target_const class arg) class;
359 fun abbrev class prmode arg = target_extension (target_abbrev class prmode arg) class;
364 (* simple subclasses *)
368 fun gen_classrel mk_prop classrel thy =
370 fun after_qed results =
371 ProofContext.background_theory ((fold o fold) AxClass.add_classrel results);
374 |> ProofContext.init_global
375 |> Proof.theorem NONE after_qed [[(mk_prop thy classrel, [])]]
381 gen_classrel (Logic.mk_classrel oo AxClass.cert_classrel);
383 gen_classrel (Logic.mk_classrel oo AxClass.read_classrel);
388 (** instantiation target **)
392 datatype instantiation = Instantiation of {
393 arities: string list * (string * sort) list * sort,
394 params: ((string * string) * (string * typ)) list
395 (*(instantiation parameter, type constructor), (local instantiation parameter, typ)*)
398 structure Instantiation = Proof_Data
400 type T = instantiation
401 fun init _ = Instantiation { arities = ([], [], []), params = [] };
404 fun mk_instantiation (arities, params) =
405 Instantiation { arities = arities, params = params };
406 fun get_instantiation lthy = case Instantiation.get (Local_Theory.target_of lthy)
407 of Instantiation data => data;
408 fun map_instantiation f = (Local_Theory.target o Instantiation.map)
409 (fn Instantiation { arities, params } => mk_instantiation (f (arities, params)));
411 fun the_instantiation lthy = case get_instantiation lthy
412 of { arities = ([], [], []), ... } => error "No instantiation target"
415 val instantiation_params = #params o get_instantiation;
417 fun instantiation_param lthy b = instantiation_params lthy
418 |> find_first (fn (_, (v, _)) => Binding.name_of b = v)
419 |> Option.map (fst o fst);
421 fun read_multi_arity thy (raw_tycos, raw_sorts, raw_sort) =
423 val ctxt = ProofContext.init_global thy;
424 val all_arities = map (fn raw_tyco => ProofContext.read_arity ctxt
425 (raw_tyco, raw_sorts, raw_sort)) raw_tycos;
426 val tycos = map #1 all_arities;
427 val (_, sorts, sort) = hd all_arities;
428 val vs = Name.names Name.context Name.aT sorts;
429 in (tycos, vs, sort) end;
434 fun synchronize_inst_syntax ctxt =
436 val Instantiation { params, ... } = Instantiation.get ctxt;
438 val lookup_inst_param = AxClass.lookup_inst_param
439 (Sign.consts_of (ProofContext.theory_of ctxt)) params;
440 fun subst (c, ty) = case lookup_inst_param (c, ty)
441 of SOME (v_ty as (_, ty)) => SOME (ty, Free v_ty)
444 map (fn ((c, _), v_ty as (_, ty)) => (Free v_ty, Const (c, ty))) params;
447 |> Overloading.map_improvable_syntax
448 (fn (((primary_constraints, _), (((improve, _), _), _)), _) =>
449 (((primary_constraints, []), (((improve, subst), false), unchecks)), false))
452 fun resort_terms pp algebra consts constraints ts =
454 fun matchings (Const (c_ty as (c, _))) = (case constraints c
456 | SOME sorts => fold2 (curry (Sorts.meet_sort algebra))
457 (Consts.typargs consts c_ty) sorts)
459 val tvartab = (fold o fold_aterms) matchings ts Vartab.empty
460 handle Sorts.CLASS_ERROR e => error (Sorts.class_error pp e);
461 val inst = map_type_tvar
462 (fn (vi, sort) => TVar (vi, the_default sort (Vartab.lookup tvartab vi)));
463 in if Vartab.is_empty tvartab then NONE else SOME ((map o map_types) inst ts) end;
468 val sanitize_name = (*necessary as long as "dirty" type identifiers are permitted*)
470 fun is_valid s = Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s
471 orelse s = "'" orelse s = "_";
472 val is_junk = not o is_valid andf Symbol.is_regular;
473 val junk = Scan.many is_junk;
474 val scan_valids = Symbol.scanner "Malformed input"
476 (Scan.optional (Scan.one Symbol.is_ascii_letter) "x" ^^ (Scan.many is_valid >> implode)
478 ::: Scan.repeat ((Scan.many1 is_valid >> implode) --| junk));
480 explode #> scan_valids #> implode
483 val type_name = sanitize_name o Long_Name.base_name;
485 fun define_overloaded (c, U) v (b_def, rhs) = Local_Theory.background_theory_result
486 (AxClass.declare_overloaded (c, U) ##>> AxClass.define_overloaded b_def (c, rhs))
487 ##> (map_instantiation o apsnd) (filter_out (fn (_, (v', _)) => v' = v))
488 ##> Local_Theory.target synchronize_inst_syntax;
490 fun foundation (((b, U), mx), (b_def, rhs)) (type_params, term_params) lthy =
491 case instantiation_param lthy b
492 of SOME c => if mx <> NoSyn then error ("Illegal mixfix syntax for overloaded constant " ^ quote c)
493 else lthy |> define_overloaded (c, U) (Binding.name_of b) (b_def, rhs)
495 Generic_Target.theory_foundation (((b, U), mx), (b_def, rhs)) (type_params, term_params);
499 val { arities = (tycos, vs, sort), params } = the_instantiation lthy;
500 val thy = ProofContext.theory_of lthy;
501 fun pr_arity tyco = Syntax.pretty_arity lthy (tyco, map snd vs, sort);
502 fun pr_param ((c, _), (v, ty)) =
503 (Pretty.block o Pretty.breaks) [Pretty.str v, Pretty.str "==",
504 (Pretty.str o Sign.extern_const thy) c, Pretty.str "::", Syntax.pretty_typ_global thy ty];
505 in Pretty.str "instantiation" :: map pr_arity tycos @ map pr_param params end;
509 val (tycos, vs, sort) = (#arities o the_instantiation) lthy;
510 val thy = ProofContext.theory_of lthy;
511 val _ = map (fn tyco => if Sign.of_sort thy
512 (Type (tyco, map TFree vs), sort)
513 then () else error ("Missing instance proof for type " ^ quote (Sign.extern_type thy tyco)))
517 fun instantiation (tycos, vs, sort) thy =
519 val _ = if null tycos then error "At least one arity must be given" else ();
520 val class_params = these_params thy (filter (can (AxClass.get_info thy)) sort);
521 fun get_param tyco (param, (_, (c, ty))) =
522 if can (AxClass.param_of_inst thy) (c, tyco)
523 then NONE else SOME ((c, tyco),
524 (param ^ "_" ^ type_name tyco, map_atyps (K (Type (tyco, map TFree vs))) ty));
525 val params = map_product get_param tycos class_params |> map_filter I;
526 val primary_constraints = map (apsnd
527 (map_atyps (K (TVar ((Name.aT, 0), [])))) o snd o snd) class_params;
528 val algebra = Sign.classes_of thy
529 |> fold (fn tyco => Sorts.add_arities (Syntax.pp_global thy)
530 (tyco, map (fn class => (class, map snd vs)) sort)) tycos;
531 val consts = Sign.consts_of thy;
532 val improve_constraints = AList.lookup (op =)
533 (map (fn (_, (class, (c, _))) => (c, [[class]])) class_params);
534 fun resort_check ts ctxt = case resort_terms (Syntax.pp ctxt) algebra consts improve_constraints ts
536 | SOME ts' => SOME (ts', ctxt);
537 val lookup_inst_param = AxClass.lookup_inst_param consts params;
538 val typ_instance = Type.typ_instance (Sign.tsig_of thy);
539 fun improve (c, ty) = case lookup_inst_param (c, ty)
540 of SOME (_, ty') => if typ_instance (ty', ty) then SOME (ty, ty') else NONE
545 |> ProofContext.init_global
546 |> Instantiation.put (mk_instantiation ((tycos, vs, sort), params))
547 |> fold (Variable.declare_typ o TFree) vs
548 |> fold (Variable.declare_names o Free o snd) params
549 |> (Overloading.map_improvable_syntax o apfst)
550 (K ((primary_constraints, []), (((improve, K NONE), false), [])))
551 |> Overloading.activate_improvable_syntax
552 |> Context.proof_map (Syntax.add_term_check 0 "resorting" resort_check)
553 |> synchronize_inst_syntax
554 |> Local_Theory.init NONE ""
555 {define = Generic_Target.define foundation,
556 notes = Generic_Target.notes
557 (fn kind => fn global_facts => fn _ => Generic_Target.theory_notes kind global_facts),
558 abbrev = Generic_Target.abbrev
559 (fn prmode => fn (b, mx) => fn (t, _) => fn _ => Generic_Target.theory_abbrev prmode ((b, mx), t)),
560 declaration = K Generic_Target.theory_declaration,
561 syntax_declaration = K Generic_Target.theory_declaration,
563 exit = Local_Theory.target_of o conclude}
566 fun instantiation_cmd arities thy =
567 instantiation (read_multi_arity thy arities) thy;
569 fun gen_instantiation_instance do_proof after_qed lthy =
571 val (tycos, vs, sort) = (#arities o the_instantiation) lthy;
572 val arities_proof = maps (fn tyco => Logic.mk_arities (tyco, map snd vs, sort)) tycos;
573 fun after_qed' results =
574 Local_Theory.background_theory (fold (AxClass.add_arity o Thm.varifyT_global) results)
578 |> do_proof after_qed' arities_proof
581 val instantiation_instance = gen_instantiation_instance (fn after_qed => fn ts =>
582 Proof.theorem NONE (after_qed o map the_single) (map (fn t => [(t, [])]) ts));
584 fun prove_instantiation_instance tac = gen_instantiation_instance (fn after_qed =>
585 fn ts => fn lthy => after_qed (map (fn t => Goal.prove lthy [] [] t
586 (fn {context, ...} => tac context)) ts) lthy) I;
588 fun prove_instantiation_exit tac = prove_instantiation_instance tac
589 #> Local_Theory.exit_global;
591 fun prove_instantiation_exit_result f tac x lthy =
593 val morph = ProofContext.export_morphism lthy
594 (ProofContext.init_global (ProofContext.theory_of lthy));
598 |> prove_instantiation_exit (fn ctxt => tac ctxt y)
603 (* simplified instantiation interface with no class parameter *)
605 fun instance_arity_cmd raw_arities thy =
607 val (tycos, vs, sort) = read_multi_arity thy raw_arities;
608 val sorts = map snd vs;
609 val arities = maps (fn tyco => Logic.mk_arities (tyco, sorts, sort)) tycos;
610 fun after_qed results = ProofContext.background_theory
611 ((fold o fold) AxClass.add_arity results);
614 |> ProofContext.init_global
615 |> Proof.theorem NONE after_qed (map (fn t => [(t, [])]) arities)
619 (** tactics and methods **)
621 fun intro_classes_tac facts st =
623 val thy = Thm.theory_of_thm st;
624 val classes = Sign.all_classes thy;
625 val class_trivs = map (Thm.class_triv thy) classes;
626 val class_intros = map_filter (try (#intro o AxClass.get_info thy)) classes;
627 val assm_intros = all_assm_intros thy;
629 Method.intros_tac (class_trivs @ class_intros @ assm_intros) facts st
632 fun default_intro_tac ctxt [] =
633 COND Thm.no_prems no_tac
634 (intro_classes_tac [] ORELSE Locale.intro_locales_tac true ctxt [])
635 | default_intro_tac _ _ = no_tac;
637 fun default_tac rules ctxt facts =
638 HEADGOAL (Method.some_rule_tac rules ctxt facts) ORELSE
639 default_intro_tac ctxt facts;
641 val _ = Context.>> (Context.map_theory
642 (Method.setup (Binding.name "intro_classes") (Scan.succeed (K (METHOD intro_classes_tac)))
643 "back-chain introduction rules of classes" #>
644 Method.setup (Binding.name "default") (Attrib.thms >> (METHOD oo default_tac))
645 "apply some intro/elim rule"));