src/Pure/Isar/class.ML
author wenzelm
Sat, 16 Apr 2011 15:47:52 +0200
changeset 43231 da8817d01e7c
parent 43230 6ca5407863ed
child 43246 774df7c59508
permissions -rw-r--r--
modernized structure Proof_Context;
     1 (*  Title:      Pure/Isar/class.ML
     2     Author:     Florian Haftmann, TU Muenchen
     3 
     4 Type classes derived from primitive axclasses and locales.
     5 *)
     6 
     7 signature CLASS =
     8 sig
     9   (*classes*)
    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: Proof.context -> 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
    26     -> theory -> theory
    27 
    28   (*instances*)
    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
    43 
    44   (*subclasses*)
    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
    49 
    50   (*tactics*)
    51   val intro_classes_tac: thm list -> tactic
    52   val default_intro_tac: Proof.context -> thm list -> tactic
    53 end;
    54 
    55 structure Class: CLASS =
    56 struct
    57 
    58 (** class data **)
    59 
    60 datatype class_data = ClassData of {
    61 
    62   (* static part *)
    63   consts: (string * string) list
    64     (*locale parameter ~> constant name*),
    65   base_sort: sort,
    66   base_morph: morphism
    67     (*static part of canonical morphism*),
    68   export_morph: morphism,
    69   assm_intro: thm option,
    70   of_class: thm,
    71   axiom: thm option,
    72 
    73   (* dynamic part *)
    74   defs: thm list,
    75   operations: (string * (class * (typ * term))) list
    76 
    77 };
    78 
    79 fun make_class_data ((consts, base_sort, base_morph, export_morph, assm_intro, of_class, axiom),
    80     (defs, operations)) =
    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),
    87     (defs, operations)));
    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)));
    96 
    97 structure ClassData = Theory_Data
    98 (
    99   type T = class_data Graph.T
   100   val empty = Graph.empty;
   101   val extend = I;
   102   val merge = Graph.join merge_class_data;
   103 );
   104 
   105 
   106 (* queries *)
   107 
   108 fun lookup_class_data thy class = case try (Graph.get_node (ClassData.get thy)) class
   109  of SOME (ClassData data) => SOME data
   110   | NONE => NONE;
   111 
   112 fun the_class_data thy class = case lookup_class_data thy class
   113  of NONE => error ("Undeclared class " ^ quote class)
   114   | SOME data => data;
   115 
   116 val is_class = is_some oo lookup_class_data;
   117 
   118 val ancestry = Graph.all_succs o ClassData.get;
   119 val heritage = Graph.all_preds o ClassData.get;
   120 
   121 fun these_params thy =
   122   let
   123     fun params class =
   124       let
   125         val const_typs = (#params o AxClass.get_info thy) class;
   126         val const_names = (#consts o the_class_data thy) class;
   127       in
   128         (map o apsnd)
   129           (fn c => (class, (c, (the o AList.lookup (op =) const_typs) c))) const_names
   130       end;
   131   in maps params o ancestry thy end;
   132 
   133 val base_sort = #base_sort oo the_class_data;
   134 
   135 fun rules thy class =
   136   let val { axiom, of_class, ... } = the_class_data thy class
   137   in (axiom, of_class) end;
   138 
   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) [];
   142 
   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;
   145 
   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;
   151 
   152 fun print_classes ctxt =
   153   let
   154     val thy = Proof_Context.theory_of ctxt;
   155     val algebra = Sign.classes_of thy;
   156     val arities =
   157       Symtab.empty
   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 =
   163       let
   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) =
   167       Pretty.str (Proof_Context.extern_const ctxt c ^ " :: " ^
   168         Syntax.string_of_typ (Config.put show_sorts false ctxt) (Type.strip_sorts ty));
   169     fun mk_entry class = (Pretty.block o Pretty.fbreaks o map_filter I) [
   170       (SOME o Pretty.str) ("class " ^ Proof_Context.extern_class ctxt class ^ ":"),
   171       (SOME o Pretty.block) [Pretty.str "supersort: ",
   172         (Syntax.pretty_sort ctxt o Sign.minimize_sort thy o Sign.super_classes thy) class],
   173       ((fn [] => NONE | ps => (SOME o Pretty.block o Pretty.fbreaks)
   174           (Pretty.str "parameters:" :: ps)) o map mk_param
   175         o these o Option.map #params o try (AxClass.get_info thy)) class,
   176       (SOME o Pretty.block o Pretty.breaks) [
   177         Pretty.str "instances:",
   178         Pretty.list "" "" (map (mk_arity class) (the_arities class))
   179       ]
   180     ]
   181   in
   182     (Pretty.writeln o Pretty.chunks o separate (Pretty.str "")
   183       o map mk_entry o Sorts.all_classes) algebra
   184   end;
   185 
   186 
   187 (* updaters *)
   188 
   189 fun register class sups params base_sort base_morph export_morph
   190     axiom assm_intro of_class thy =
   191   let
   192     val operations = map (fn (v_ty as (_, ty), (c, _)) =>
   193       (c, (class, (ty, Free v_ty)))) params;
   194     val add_class = Graph.new_node (class,
   195         make_class_data (((map o pairself) fst params, base_sort,
   196           base_morph, export_morph, assm_intro, of_class, axiom), ([], operations)))
   197       #> fold (curry Graph.add_edge class) sups;
   198   in ClassData.map add_class thy end;
   199 
   200 fun activate_defs class thms thy = case Element.eq_morphism thy thms
   201  of SOME eq_morph => fold (fn cls => fn thy =>
   202       Context.theory_map (Locale.amend_registration (cls, base_morphism thy cls)
   203         (eq_morph, true) (export_morphism thy cls)) thy) (heritage thy [class]) thy
   204   | NONE => thy;
   205 
   206 fun register_operation class (c, (t, some_def)) thy =
   207   let
   208     val base_sort = base_sort thy class;
   209     val prep_typ = map_type_tfree
   210       (fn (v, sort) => if Name.aT = v
   211         then TFree (v, base_sort) else TVar ((v, 0), sort));
   212     val t' = map_types prep_typ t;
   213     val ty' = Term.fastype_of t';
   214   in
   215     thy
   216     |> (ClassData.map o Graph.map_node class o map_class_data o apsnd)
   217       (fn (defs, operations) =>
   218         (fold cons (the_list some_def) defs,
   219           (c, (class, (ty', t'))) :: operations))
   220     |> activate_defs class (the_list some_def)
   221   end;
   222 
   223 fun register_subclass (sub, sup) some_dep_morph some_wit export thy =
   224   let
   225     val intros = (snd o rules thy) sup :: map_filter I
   226       [Option.map (Drule.export_without_context_open o Element.conclude_witness) some_wit,
   227         (fst o rules thy) sub];
   228     val tac = EVERY (map (TRYALL o Tactic.rtac) intros);
   229     val classrel = Skip_Proof.prove_global thy [] [] (Logic.mk_classrel (sub, sup))
   230       (K tac);
   231     val diff_sort = Sign.complete_sort thy [sup]
   232       |> subtract (op =) (Sign.complete_sort thy [sub])
   233       |> filter (is_class thy);
   234     val add_dependency = case some_dep_morph
   235      of SOME dep_morph => Locale.add_dependency sub
   236           (sup, dep_morph $> Element.satisfy_morphism (the_list some_wit)) NONE export
   237       | NONE => I
   238   in
   239     thy
   240     |> AxClass.add_classrel classrel
   241     |> ClassData.map (Graph.add_edge (sub, sup))
   242     |> activate_defs sub (these_defs thy diff_sort)
   243     |> add_dependency
   244   end;
   245 
   246 
   247 (** classes and class target **)
   248 
   249 (* class context syntax *)
   250 
   251 fun these_unchecks thy = map (fn (c, (_, (ty, t))) => (t, Const (c, ty)))
   252   o these_operations thy;
   253 
   254 fun redeclare_const thy c =
   255   let val b = Long_Name.base_name c
   256   in Sign.intern_const thy b = c ? Variable.declare_const (b, c) end;
   257 
   258 fun synchronize_class_syntax sort base_sort ctxt =
   259   let
   260     val thy = Proof_Context.theory_of ctxt;
   261     val algebra = Sign.classes_of thy;
   262     val operations = these_operations thy sort;
   263     fun subst_class_typ sort = map_type_tfree (K (TVar ((Name.aT, 0), sort)));
   264     val primary_constraints =
   265       (map o apsnd) (subst_class_typ base_sort o fst o snd) operations;
   266     val secondary_constraints =
   267       (map o apsnd) (fn (class, (ty, _)) => subst_class_typ [class] ty) operations;
   268     fun improve (c, ty) = (case AList.lookup (op =) primary_constraints c
   269      of SOME ty' => (case try (Type.raw_match (ty', ty)) Vartab.empty
   270          of SOME tyenv => (case Vartab.lookup tyenv (Name.aT, 0)
   271              of SOME (_, ty' as TVar (vi, sort)) =>
   272                   if Type_Infer.is_param vi
   273                     andalso Sorts.sort_le algebra (base_sort, sort)
   274                       then SOME (ty', TFree (Name.aT, base_sort))
   275                       else NONE
   276               | _ => NONE)
   277           | NONE => NONE)
   278       | NONE => NONE)
   279     fun subst (c, ty) = Option.map snd (AList.lookup (op =) operations c);
   280     val unchecks = these_unchecks thy sort;
   281   in
   282     ctxt
   283     |> fold (redeclare_const thy o fst) primary_constraints
   284     |> Overloading.map_improvable_syntax (K (((primary_constraints, secondary_constraints),
   285         (((improve, subst), true), unchecks)), false))
   286     |> Overloading.set_primary_constraints
   287   end;
   288 
   289 fun redeclare_operations thy sort =
   290   fold (redeclare_const thy o fst) (these_operations thy sort);
   291 
   292 fun begin sort base_sort ctxt =
   293   ctxt
   294   |> Variable.declare_term
   295       (Logic.mk_type (TFree (Name.aT, base_sort)))
   296   |> synchronize_class_syntax sort base_sort
   297   |> Overloading.activate_improvable_syntax;
   298 
   299 fun init class thy =
   300   thy
   301   |> Locale.init class
   302   |> begin [class] (base_sort thy class);
   303 
   304 
   305 (* class target *)
   306 
   307 val class_prefix = Logic.const_of_class o Long_Name.base_name;
   308 
   309 fun target_extension f class lthy =
   310   lthy
   311   |> Local_Theory.raw_theory f
   312   |> Local_Theory.target (synchronize_class_syntax [class]
   313       (base_sort (Proof_Context.theory_of lthy) class));
   314 
   315 local
   316 
   317 fun target_const class ((c, mx), (type_params, dict)) thy =
   318   let
   319     val morph = morphism thy class;
   320     val b = Morphism.binding morph c;
   321     val b_def = Morphism.binding morph (Binding.suffix_name "_dict" b);
   322     val c' = Sign.full_name thy b;
   323     val dict' = Morphism.term morph dict;
   324     val ty' = map Term.fastype_of type_params ---> Term.fastype_of dict';
   325     val def_eq = Logic.mk_equals (list_comb (Const (c', ty'), type_params), dict')
   326       |> map_types Type.strip_sorts;
   327   in
   328     thy
   329     |> Sign.declare_const ((b, Type.strip_sorts ty'), mx)
   330     |> snd
   331     |> Thm.add_def false false (b_def, def_eq)
   332     |>> apsnd Thm.varifyT_global
   333     |-> (fn (_, def_thm) => Global_Theory.store_thm (b_def, def_thm)
   334       #> snd
   335       #> null type_params ? register_operation class (c', (dict', SOME (Thm.symmetric def_thm))))
   336     |> Sign.add_const_constraint (c', SOME ty')
   337   end;
   338 
   339 fun target_abbrev class prmode ((c, mx), rhs) thy =
   340   let
   341     val morph = morphism thy class;
   342     val unchecks = these_unchecks thy [class];
   343     val b = Morphism.binding morph c;
   344     val c' = Sign.full_name thy b;
   345     val rhs' = Pattern.rewrite_term thy unchecks [] rhs;
   346     val ty' = Term.fastype_of rhs';
   347     val rhs'' = map_types Logic.varifyT_global rhs';
   348   in
   349     thy
   350     |> Sign.add_abbrev (#1 prmode) (b, rhs'')
   351     |> snd
   352     |> Sign.add_const_constraint (c', SOME ty')
   353     |> Sign.notation true prmode [(Const (c', ty'), mx)]
   354     |> not (#1 prmode = Print_Mode.input) ? register_operation class (c', (rhs', NONE))
   355   end;
   356 
   357 in
   358 
   359 fun const class arg = target_extension (target_const class arg) class;
   360 fun abbrev class prmode arg = target_extension (target_abbrev class prmode arg) class;
   361 
   362 end;
   363 
   364 
   365 (* simple subclasses *)
   366 
   367 local
   368 
   369 fun gen_classrel mk_prop classrel thy =
   370   let
   371     fun after_qed results =
   372       Proof_Context.background_theory ((fold o fold) AxClass.add_classrel results);
   373   in
   374     thy
   375     |> Proof_Context.init_global
   376     |> Proof.theorem NONE after_qed [[(mk_prop thy classrel, [])]]
   377   end;
   378 
   379 in
   380 
   381 val classrel =
   382   gen_classrel (Logic.mk_classrel oo AxClass.cert_classrel);
   383 val classrel_cmd =
   384   gen_classrel (Logic.mk_classrel oo AxClass.read_classrel);
   385 
   386 end; (*local*)
   387 
   388 
   389 (** instantiation target **)
   390 
   391 (* bookkeeping *)
   392 
   393 datatype instantiation = Instantiation of {
   394   arities: string list * (string * sort) list * sort,
   395   params: ((string * string) * (string * typ)) list
   396     (*(instantiation parameter, type constructor), (local instantiation parameter, typ)*)
   397 }
   398 
   399 structure Instantiation = Proof_Data
   400 (
   401   type T = instantiation
   402   fun init _ = Instantiation { arities = ([], [], []), params = [] };
   403 );
   404 
   405 fun mk_instantiation (arities, params) =
   406   Instantiation { arities = arities, params = params };
   407 fun get_instantiation lthy = case Instantiation.get (Local_Theory.target_of lthy)
   408  of Instantiation data => data;
   409 fun map_instantiation f = (Local_Theory.target o Instantiation.map)
   410   (fn Instantiation { arities, params } => mk_instantiation (f (arities, params)));
   411 
   412 fun the_instantiation lthy = case get_instantiation lthy
   413  of { arities = ([], [], []), ... } => error "No instantiation target"
   414   | data => data;
   415 
   416 val instantiation_params = #params o get_instantiation;
   417 
   418 fun instantiation_param lthy b = instantiation_params lthy
   419   |> find_first (fn (_, (v, _)) => Binding.name_of b = v)
   420   |> Option.map (fst o fst);
   421 
   422 fun read_multi_arity thy (raw_tycos, raw_sorts, raw_sort) =
   423   let
   424     val ctxt = Proof_Context.init_global thy;
   425     val all_arities = map (fn raw_tyco => Proof_Context.read_arity ctxt
   426       (raw_tyco, raw_sorts, raw_sort)) raw_tycos;
   427     val tycos = map #1 all_arities;
   428     val (_, sorts, sort) = hd all_arities;
   429     val vs = Name.names Name.context Name.aT sorts;
   430   in (tycos, vs, sort) end;
   431 
   432 
   433 (* syntax *)
   434 
   435 fun synchronize_inst_syntax ctxt =
   436   let
   437     val Instantiation { params, ... } = Instantiation.get ctxt;
   438 
   439     val lookup_inst_param = AxClass.lookup_inst_param
   440       (Sign.consts_of (Proof_Context.theory_of ctxt)) params;
   441     fun subst (c, ty) = case lookup_inst_param (c, ty)
   442      of SOME (v_ty as (_, ty)) => SOME (ty, Free v_ty)
   443       | NONE => NONE;
   444     val unchecks =
   445       map (fn ((c, _), v_ty as (_, ty)) => (Free v_ty, Const (c, ty))) params;
   446   in
   447     ctxt
   448     |> Overloading.map_improvable_syntax
   449          (fn (((primary_constraints, _), (((improve, _), _), _)), _) =>
   450             (((primary_constraints, []), (((improve, subst), false), unchecks)), false))
   451   end;
   452 
   453 fun resort_terms pp algebra consts constraints ts =
   454   let
   455     fun matchings (Const (c_ty as (c, _))) = (case constraints c
   456          of NONE => I
   457           | SOME sorts => fold2 (curry (Sorts.meet_sort algebra))
   458               (Consts.typargs consts c_ty) sorts)
   459       | matchings _ = I
   460     val tvartab = (fold o fold_aterms) matchings ts Vartab.empty
   461       handle Sorts.CLASS_ERROR e => error (Sorts.class_error pp e);
   462     val inst = map_type_tvar
   463       (fn (vi, sort) => TVar (vi, the_default sort (Vartab.lookup tvartab vi)));
   464   in if Vartab.is_empty tvartab then NONE else SOME ((map o map_types) inst ts) end;
   465 
   466 
   467 (* target *)
   468 
   469 val sanitize_name = (*necessary as long as "dirty" type identifiers are permitted*)
   470   let
   471     fun is_valid s = Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s
   472       orelse s = "'" orelse s = "_";
   473     val is_junk = not o is_valid andf Symbol.is_regular;
   474     val junk = Scan.many is_junk;
   475     val scan_valids = Symbol.scanner "Malformed input"
   476       ((junk |--
   477         (Scan.optional (Scan.one Symbol.is_ascii_letter) "x" ^^ (Scan.many is_valid >> implode)
   478         --| junk))
   479       ::: Scan.repeat ((Scan.many1 is_valid >> implode) --| junk));
   480   in
   481     raw_explode #> scan_valids #> implode
   482   end;
   483 
   484 val type_name = sanitize_name o Long_Name.base_name;
   485 
   486 fun define_overloaded (c, U) v (b_def, rhs) = Local_Theory.background_theory_result
   487     (AxClass.declare_overloaded (c, U) ##>> AxClass.define_overloaded b_def (c, rhs))
   488   ##> (map_instantiation o apsnd) (filter_out (fn (_, (v', _)) => v' = v))
   489   ##> Local_Theory.target synchronize_inst_syntax;
   490 
   491 fun foundation (((b, U), mx), (b_def, rhs)) (type_params, term_params) lthy =
   492   case instantiation_param lthy b
   493    of SOME c => if mx <> NoSyn then error ("Illegal mixfix syntax for overloaded constant " ^ quote c)
   494         else lthy |> define_overloaded (c, U) (Binding.name_of b) (b_def, rhs)
   495     | NONE => lthy |>
   496         Generic_Target.theory_foundation (((b, U), mx), (b_def, rhs)) (type_params, term_params);
   497 
   498 fun pretty lthy =
   499   let
   500     val { arities = (tycos, vs, sort), params } = the_instantiation lthy;
   501     val thy = Proof_Context.theory_of lthy;
   502     fun pr_arity tyco = Syntax.pretty_arity lthy (tyco, map snd vs, sort);
   503     fun pr_param ((c, _), (v, ty)) =
   504       Pretty.block (Pretty.breaks
   505         [Pretty.str v, Pretty.str "==", Pretty.str (Proof_Context.extern_const lthy c),
   506           Pretty.str "::", Syntax.pretty_typ lthy ty]);
   507   in Pretty.str "instantiation" :: map pr_arity tycos @ map pr_param params end;
   508 
   509 fun conclude lthy =
   510   let
   511     val (tycos, vs, sort) = #arities (the_instantiation lthy);
   512     val thy = Proof_Context.theory_of lthy;
   513     val _ = tycos |> List.app (fn tyco =>
   514       if Sign.of_sort thy
   515         (Type (tyco, map TFree vs), sort)
   516       then ()
   517       else error ("Missing instance proof for type " ^ quote (Proof_Context.extern_type lthy tyco)));
   518   in lthy end;
   519 
   520 fun instantiation (tycos, vs, sort) thy =
   521   let
   522     val _ = if null tycos then error "At least one arity must be given" else ();
   523     val class_params = these_params thy (filter (can (AxClass.get_info thy)) sort);
   524     fun get_param tyco (param, (_, (c, ty))) =
   525       if can (AxClass.param_of_inst thy) (c, tyco)
   526       then NONE else SOME ((c, tyco),
   527         (param ^ "_" ^ type_name tyco, map_atyps (K (Type (tyco, map TFree vs))) ty));
   528     val params = map_product get_param tycos class_params |> map_filter I;
   529     val primary_constraints = map (apsnd
   530       (map_atyps (K (TVar ((Name.aT, 0), [])))) o snd o snd) class_params;
   531     val algebra = Sign.classes_of thy
   532       |> fold (fn tyco => Sorts.add_arities (Syntax.pp_global thy)
   533             (tyco, map (fn class => (class, map snd vs)) sort)) tycos;
   534     val consts = Sign.consts_of thy;
   535     val improve_constraints = AList.lookup (op =)
   536       (map (fn (_, (class, (c, _))) => (c, [[class]])) class_params);
   537     fun resort_check ts ctxt = case resort_terms (Syntax.pp ctxt) algebra consts improve_constraints ts
   538      of NONE => NONE
   539       | SOME ts' => SOME (ts', ctxt);
   540     val lookup_inst_param = AxClass.lookup_inst_param consts params;
   541     val typ_instance = Type.typ_instance (Sign.tsig_of thy);
   542     fun improve (c, ty) = case lookup_inst_param (c, ty)
   543      of SOME (_, ty') => if typ_instance (ty', ty) then SOME (ty, ty') else NONE
   544       | NONE => NONE;
   545   in
   546     thy
   547     |> Theory.checkpoint
   548     |> Proof_Context.init_global
   549     |> Instantiation.put (mk_instantiation ((tycos, vs, sort), params))
   550     |> fold (Variable.declare_typ o TFree) vs
   551     |> fold (Variable.declare_names o Free o snd) params
   552     |> (Overloading.map_improvable_syntax o apfst)
   553          (K ((primary_constraints, []), (((improve, K NONE), false), [])))
   554     |> Overloading.activate_improvable_syntax
   555     |> Context.proof_map (Syntax.add_term_check 0 "resorting" resort_check)
   556     |> synchronize_inst_syntax
   557     |> Local_Theory.init NONE ""
   558        {define = Generic_Target.define foundation,
   559         notes = Generic_Target.notes
   560           (fn kind => fn global_facts => fn _ => Generic_Target.theory_notes kind global_facts),
   561         abbrev = Generic_Target.abbrev
   562           (fn prmode => fn (b, mx) => fn (t, _) => fn _ => Generic_Target.theory_abbrev prmode ((b, mx), t)),
   563         declaration = K Generic_Target.theory_declaration,
   564         syntax_declaration = K Generic_Target.theory_declaration,
   565         pretty = pretty,
   566         exit = Local_Theory.target_of o conclude}
   567   end;
   568 
   569 fun instantiation_cmd arities thy =
   570   instantiation (read_multi_arity thy arities) thy;
   571 
   572 fun gen_instantiation_instance do_proof after_qed lthy =
   573   let
   574     val (tycos, vs, sort) = (#arities o the_instantiation) lthy;
   575     val arities_proof = maps (fn tyco => Logic.mk_arities (tyco, map snd vs, sort)) tycos;
   576     fun after_qed' results =
   577       Local_Theory.background_theory (fold (AxClass.add_arity o Thm.varifyT_global) results)
   578       #> after_qed;
   579   in
   580     lthy
   581     |> do_proof after_qed' arities_proof
   582   end;
   583 
   584 val instantiation_instance = gen_instantiation_instance (fn after_qed => fn ts =>
   585   Proof.theorem NONE (after_qed o map the_single) (map (fn t => [(t, [])]) ts));
   586 
   587 fun prove_instantiation_instance tac = gen_instantiation_instance (fn after_qed =>
   588   fn ts => fn lthy => after_qed (map (fn t => Goal.prove lthy [] [] t
   589     (fn {context, ...} => tac context)) ts) lthy) I;
   590 
   591 fun prove_instantiation_exit tac = prove_instantiation_instance tac
   592   #> Local_Theory.exit_global;
   593 
   594 fun prove_instantiation_exit_result f tac x lthy =
   595   let
   596     val morph = Proof_Context.export_morphism lthy
   597       (Proof_Context.init_global (Proof_Context.theory_of lthy));
   598     val y = f morph x;
   599   in
   600     lthy
   601     |> prove_instantiation_exit (fn ctxt => tac ctxt y)
   602     |> pair y
   603   end;
   604 
   605 
   606 (* simplified instantiation interface with no class parameter *)
   607 
   608 fun instance_arity_cmd raw_arities thy =
   609   let
   610     val (tycos, vs, sort) = read_multi_arity thy raw_arities;
   611     val sorts = map snd vs;
   612     val arities = maps (fn tyco => Logic.mk_arities (tyco, sorts, sort)) tycos;
   613     fun after_qed results = Proof_Context.background_theory
   614       ((fold o fold) AxClass.add_arity results);
   615   in
   616     thy
   617     |> Proof_Context.init_global
   618     |> Proof.theorem NONE after_qed (map (fn t => [(t, [])]) arities)
   619   end;
   620 
   621 
   622 (** tactics and methods **)
   623 
   624 fun intro_classes_tac facts st =
   625   let
   626     val thy = Thm.theory_of_thm st;
   627     val classes = Sign.all_classes thy;
   628     val class_trivs = map (Thm.class_triv thy) classes;
   629     val class_intros = map_filter (try (#intro o AxClass.get_info thy)) classes;
   630     val assm_intros = all_assm_intros thy;
   631   in
   632     Method.intros_tac (class_trivs @ class_intros @ assm_intros) facts st
   633   end;
   634 
   635 fun default_intro_tac ctxt [] =
   636       COND Thm.no_prems no_tac
   637         (intro_classes_tac [] ORELSE Locale.intro_locales_tac true ctxt [])
   638   | default_intro_tac _ _ = no_tac;
   639 
   640 fun default_tac rules ctxt facts =
   641   HEADGOAL (Method.some_rule_tac rules ctxt facts) ORELSE
   642     default_intro_tac ctxt facts;
   643 
   644 val _ = Context.>> (Context.map_theory
   645  (Method.setup (Binding.name "intro_classes") (Scan.succeed (K (METHOD intro_classes_tac)))
   646     "back-chain introduction rules of classes" #>
   647   Method.setup (Binding.name "default") (Attrib.thms >> (METHOD oo default_tac))
   648     "apply some intro/elim rule"));
   649 
   650 end;
   651