src/Pure/Isar/class.ML
author wenzelm
Tue, 09 Oct 2007 00:20:13 +0200
changeset 24920 2a45e400fdad
parent 24914 95cda5dd58d5
child 24930 cc2e0e8c81af
permissions -rw-r--r--
generic Syntax.pretty/string_of operations;
     1 (*  Title:      Pure/Isar/class.ML
     2     ID:         $Id$
     3     Author:     Florian Haftmann, TU Muenchen
     4 
     5 Type classes derived from primitive axclasses and locales.
     6 *)
     7 
     8 signature CLASS =
     9 sig
    10   val fork_mixfix: bool -> string option -> mixfix -> mixfix * mixfix
    11 
    12   val axclass_cmd: bstring * xstring list
    13     -> ((bstring * Attrib.src list) * string list) list
    14     -> theory -> class * theory
    15   val classrel_cmd: xstring * xstring -> theory -> Proof.state
    16 
    17   val class: bool -> bstring -> class list -> Element.context_i Locale.element list
    18     -> string list -> theory -> string * Proof.context
    19   val class_cmd: bool -> bstring -> xstring list -> Element.context Locale.element list
    20     -> xstring list -> theory -> string * Proof.context
    21   val init: class -> Proof.context -> Proof.context;
    22   val add_const_in_class: string -> (string * term) * Syntax.mixfix
    23     -> theory -> string * theory
    24   val add_abbrev_in_class: string -> Syntax.mode -> (string * term) * mixfix
    25     -> theory -> string * theory
    26   val remove_constraint: sort -> string -> Proof.context -> Proof.context
    27   val intro_classes_tac: thm list -> tactic
    28   val default_intro_classes_tac: thm list -> tactic
    29   val class_of_locale: theory -> string -> class option
    30   val locale_of_class: theory -> class -> string
    31   val local_syntax: theory -> class -> bool
    32   val print_classes: theory -> unit
    33 
    34   val instance_arity: (theory -> theory) -> arity list -> theory -> Proof.state
    35   val instance: arity list -> ((bstring * Attrib.src list) * term) list
    36     -> (thm list -> theory -> theory)
    37     -> theory -> Proof.state
    38   val instance_cmd: (bstring * xstring list * xstring) list
    39     -> ((bstring * Attrib.src list) * xstring) list
    40     -> (thm list -> theory -> theory)
    41     -> theory -> Proof.state
    42   val prove_instance: tactic -> arity list
    43     -> ((bstring * Attrib.src list) * term) list
    44     -> theory -> thm list * theory
    45   val unoverload: theory -> conv
    46   val overload: theory -> conv
    47   val unoverload_const: theory -> string * typ -> string
    48   val inst_const: theory -> string * string -> string
    49   val param_const: theory -> string -> (string * string) option
    50   val params_of_sort: theory -> sort -> (string * (string * typ)) list
    51 end;
    52 
    53 structure Class : CLASS =
    54 struct
    55 
    56 (** auxiliary **)
    57 
    58 fun fork_mixfix is_loc some_class mx =
    59   let
    60     val mx' = Syntax.unlocalize_mixfix mx;
    61     val mx_global = if not is_loc orelse (is_some some_class andalso not (mx = mx'))
    62       then mx' else NoSyn;
    63     val mx_local = if is_loc then mx else NoSyn;
    64   in (mx_global, mx_local) end;
    65 
    66 fun prove_interpretation tac prfx_atts expr insts =
    67   Locale.interpretation_i I prfx_atts expr insts
    68   #> Proof.global_terminal_proof
    69       (Method.Basic (K (Method.SIMPLE_METHOD tac), Position.none), NONE)
    70   #> ProofContext.theory_of;
    71 
    72 fun OF_LAST thm1 thm2 =
    73   let
    74     val n = (length o Logic.strip_imp_prems o prop_of) thm2;
    75   in (thm1 RSN (n, thm2)) end;
    76 
    77 fun strip_all_ofclass thy sort =
    78   let
    79     val typ = TVar ((Name.aT, 0), sort);
    80     fun prem_inclass t =
    81       case Logic.strip_imp_prems t
    82        of ofcls :: _ => try Logic.dest_inclass ofcls
    83         | [] => NONE;
    84     fun strip_ofclass class thm =
    85       thm OF (fst o AxClass.of_sort thy (typ, [class])) AxClass.cache;
    86     fun strip thm = case (prem_inclass o Thm.prop_of) thm
    87      of SOME (_, class) => thm |> strip_ofclass class |> strip
    88       | NONE => thm;
    89   in strip end;
    90 
    91 
    92 (** axclass command **)
    93 
    94 fun axclass_cmd (class, raw_superclasses) raw_specs thy =
    95   let
    96     val ctxt = ProofContext.init thy;
    97     val superclasses = map (Sign.read_class thy) raw_superclasses;
    98     val name_atts = map ((apsnd o map) (Attrib.attribute thy) o fst)
    99       raw_specs;
   100     val axiomss = ProofContext.read_propp (ctxt, map (map (rpair []) o snd)
   101           raw_specs)
   102       |> snd
   103       |> (map o map) fst;
   104   in
   105     AxClass.define_class (class, superclasses) []
   106       (name_atts ~~ axiomss) thy
   107   end;
   108 
   109 local
   110 
   111 fun gen_instance mk_prop add_thm after_qed insts thy =
   112   let
   113     fun after_qed' results =
   114       ProofContext.theory ((fold o fold) add_thm results #> after_qed);
   115   in
   116     thy
   117     |> ProofContext.init
   118     |> Proof.theorem_i NONE after_qed' ((map (fn t => [(t, [])])
   119         o maps (mk_prop thy)) insts)
   120   end;
   121 
   122 in
   123 
   124 val instance_arity =
   125   gen_instance (Logic.mk_arities oo Sign.cert_arity) AxClass.add_arity;
   126 val classrel =
   127   gen_instance (single oo (Logic.mk_classrel oo AxClass.cert_classrel))
   128     AxClass.add_classrel I o single;
   129 val classrel_cmd =
   130   gen_instance (single oo (Logic.mk_classrel oo AxClass.read_classrel))
   131     AxClass.add_classrel I o single;
   132 
   133 end; (*local*)
   134 
   135 
   136 (** explicit constants for overloaded definitions **)
   137 
   138 structure InstData = TheoryDataFun
   139 (
   140   type T = (string * thm) Symtab.table Symtab.table * (string * string) Symtab.table;
   141     (*constant name ~> type constructor ~> (constant name, equation),
   142         constant name ~> (constant name, type constructor)*)
   143   val empty = (Symtab.empty, Symtab.empty);
   144   val copy = I;
   145   val extend = I;
   146   fun merge _ ((taba1, tabb1), (taba2, tabb2)) =
   147     (Symtab.join (K (Symtab.merge (K true))) (taba1, taba2),
   148       Symtab.merge (K true) (tabb1, tabb2));
   149 );
   150 
   151 fun inst_thms f thy = (Symtab.fold (Symtab.fold (cons o f o snd o snd) o snd) o fst)
   152     (InstData.get thy) [];
   153 fun add_inst (c, tyco) inst = (InstData.map o apfst
   154       o Symtab.map_default (c, Symtab.empty)) (Symtab.update_new (tyco, inst))
   155   #> (InstData.map o apsnd) (Symtab.update_new (fst inst, (c, tyco)));
   156 
   157 fun unoverload thy = MetaSimplifier.rewrite false (inst_thms I thy);
   158 fun overload thy = MetaSimplifier.rewrite false (inst_thms symmetric thy);
   159 
   160 fun inst_const thy (c, tyco) =
   161   (fst o the o Symtab.lookup ((the o Symtab.lookup (fst (InstData.get thy))) c)) tyco;
   162 fun unoverload_const thy (c_ty as (c, _)) =
   163   case AxClass.class_of_param thy c
   164    of SOME class => (case Sign.const_typargs thy c_ty
   165        of [Type (tyco, _)] => (case Symtab.lookup
   166            ((the o Symtab.lookup (fst (InstData.get thy))) c) tyco
   167              of SOME (c, _) => c
   168               | NONE => c)
   169         | [_] => c)
   170     | NONE => c;
   171 
   172 val param_const = Symtab.lookup o snd o InstData.get;
   173 
   174 fun add_inst_def (class, tyco) (c, ty) thy =
   175   let
   176     val tyco_base = NameSpace.base tyco;
   177     val name_inst = NameSpace.base class ^ "_" ^ tyco_base ^ "_inst";
   178     val c_inst_base = NameSpace.base c ^ "_" ^ tyco_base;
   179   in
   180     thy
   181     |> Sign.sticky_prefix name_inst
   182     |> PureThy.simple_def ("", [])
   183          (((c_inst_base, ty, Syntax.NoSyn), []), Const (c, ty))
   184     |-> (fn (c_inst, thm) => add_inst (c, tyco) (c_inst, symmetric thm))
   185     |> Sign.restore_naming thy
   186   end;
   187 
   188 fun add_inst_def' (class, tyco) (c, ty) thy =
   189   if case Symtab.lookup (fst (InstData.get thy)) c
   190    of NONE => true
   191     | SOME tab => is_none (Symtab.lookup tab tyco)
   192   then add_inst_def (class, tyco) (c, Logic.unvarifyT ty) thy
   193   else thy;
   194 
   195 fun add_def ((class, tyco), ((name, prop), atts)) thy =
   196   let
   197     val ((lhs as Const (c, ty), args), rhs) =
   198       (apfst Term.strip_comb o Logic.dest_equals) prop;
   199     fun (*add_inst' def ([], (Const (c_inst, ty))) =
   200           if forall (fn TFree _ => true | _ => false) (Sign.const_typargs thy (c_inst, ty))
   201           then add_inst (c, tyco) (c_inst, def)
   202           else add_inst_def (class, tyco) (c, ty)
   203       |*) add_inst' _ t = add_inst_def (class, tyco) (c, ty);
   204   in
   205     thy
   206     |> PureThy.add_defs_i true [((name, prop), map (Attrib.attribute thy) atts)]
   207     |-> (fn [def] => add_inst' def (args, rhs) #> pair def)
   208   end;
   209 
   210 
   211 (** instances with implicit parameter handling **)
   212 
   213 local
   214 
   215 fun gen_read_def thy prep_att parse_prop ((raw_name, raw_atts), raw_t) =
   216   let
   217     val ctxt = ProofContext.init thy;
   218     val t = parse_prop ctxt raw_t |> Syntax.check_prop ctxt;
   219     val ((c, ty), _) = Sign.cert_def (Sign.pp thy) t;
   220     val atts = map (prep_att thy) raw_atts;
   221     val insts = Consts.typargs (Sign.consts_of thy) (c, ty);
   222     val name = case raw_name
   223      of "" => NONE
   224       | _ => SOME raw_name;
   225   in (c, (insts, ((name, t), atts))) end;
   226 
   227 fun read_def_cmd thy = gen_read_def thy Attrib.intern_src Syntax.parse_prop;
   228 fun read_def thy = gen_read_def thy (K I) (K I);
   229 
   230 fun gen_instance prep_arity read_def do_proof raw_arities raw_defs after_qed theory =
   231   let
   232     val arities = map (prep_arity theory) raw_arities;
   233     val _ = if null arities then error "at least one arity must be given" else ();
   234     val _ = case (duplicates (op =) o map #1) arities
   235      of [] => ()
   236       | dupl_tycos => error ("type constructors occur more than once in arities: "
   237           ^ (commas o map quote) dupl_tycos);
   238     fun get_remove_constraint c thy =
   239       let
   240         val ty = Sign.the_const_constraint thy c;
   241       in
   242         thy
   243         |> Sign.add_const_constraint (c, NONE)
   244         |> pair (c, Logic.unvarifyT ty)
   245       end;
   246     fun get_consts_class tyco ty class =
   247       let
   248         val cs = (these o Option.map snd o try (AxClass.params_of_class theory)) class;
   249         val subst_ty = map_type_tfree (K ty);
   250       in
   251         map (fn (c, ty) => (c, ((class, tyco), subst_ty ty))) cs
   252       end;
   253     fun get_consts_sort (tyco, asorts, sort) =
   254       let
   255         val ty = Type (tyco, map (fn (v, sort) => TVar ((v, 0), sort))
   256           (Name.names Name.context Name.aT asorts))
   257       in maps (get_consts_class tyco ty) (Sign.complete_sort theory sort) end;
   258     val cs = maps get_consts_sort arities;
   259     fun mk_typnorm thy (ty, ty_sc) =
   260       case try (Sign.typ_match thy (Logic.varifyT ty_sc, ty)) Vartab.empty
   261        of SOME env => SOME (Logic.varifyT #> Envir.typ_subst_TVars env #> Logic.unvarifyT)
   262         | NONE => NONE;
   263     fun read_defs defs cs thy_read =
   264       let
   265         fun read raw_def cs =
   266           let
   267             val (c, (inst, ((name_opt, t), atts))) = read_def thy_read raw_def;
   268             val ty = Consts.instance (Sign.consts_of thy_read) (c, inst);
   269             val ((class, tyco), ty') = case AList.lookup (op =) cs c
   270              of NONE => error ("illegal definition for constant " ^ quote c)
   271               | SOME class_ty => class_ty;
   272             val name = case name_opt
   273              of NONE => Thm.def_name (Logic.name_arity (tyco, [], c))
   274               | SOME name => name;
   275             val t' = case mk_typnorm thy_read (ty', ty)
   276              of NONE => error ("illegal definition for constant " ^
   277                   quote (c ^ "::" ^ setmp show_sorts true
   278                     (Sign.string_of_typ thy_read) ty))
   279               | SOME norm => map_types norm t
   280           in (((class, tyco), ((name, t'), atts)), AList.delete (op =) c cs) end;
   281       in fold_map read defs cs end;
   282     val (defs, other_cs) = read_defs raw_defs cs
   283       (fold Sign.primitive_arity arities (Theory.copy theory));
   284     fun after_qed' cs defs =
   285       fold Sign.add_const_constraint (map (apsnd SOME) cs)
   286       #> after_qed defs;
   287   in
   288     theory
   289     |> fold_map get_remove_constraint (map fst cs |> distinct (op =))
   290     ||>> fold_map add_def defs
   291     ||> fold (fn (c, ((class, tyco), ty)) => add_inst_def' (class, tyco) (c, ty)) other_cs
   292     |-> (fn (cs, defs) => do_proof (after_qed' cs defs) arities defs)
   293   end;
   294 
   295 fun tactic_proof tac f arities defs =
   296   fold (fn arity => AxClass.prove_arity arity tac) arities
   297   #> f
   298   #> pair defs;
   299 
   300 in
   301 
   302 val instance = gen_instance Sign.cert_arity read_def
   303   (fn f => fn arities => fn defs => instance_arity f arities);
   304 val instance_cmd = gen_instance Sign.read_arity read_def_cmd
   305   (fn f => fn arities => fn defs => instance_arity f arities);
   306 fun prove_instance tac arities defs =
   307   gen_instance Sign.cert_arity read_def
   308     (tactic_proof tac) arities defs (K I);
   309 
   310 end; (*local*)
   311 
   312 
   313 
   314 (** class data **)
   315 
   316 datatype class_data = ClassData of {
   317   locale: string,
   318   consts: (string * string) list
   319     (*locale parameter ~> constant name*),
   320   local_sort: sort,
   321   inst: typ Symtab.table * term Symtab.table
   322     (*canonical interpretation*),
   323   intro: thm,
   324   local_syntax: bool,
   325   defs: thm list,
   326   operations: (string * (term * int) option) list,
   327     (*constant name ~> (locale term, instantiaton index of class typ)*)
   328   constraints: (string * typ) list
   329 };
   330 
   331 fun rep_class_data (ClassData d) = d;
   332 fun mk_class_data ((locale, consts, local_sort, inst, intro, local_syntax),
   333     (defs, (operations, constraints))) =
   334   ClassData { locale = locale, consts = consts, local_sort = local_sort, inst = inst,
   335     intro = intro, local_syntax = local_syntax, defs = defs, operations = operations,
   336     constraints = constraints };
   337 fun map_class_data f (ClassData { locale, consts, local_sort, inst, intro,
   338     local_syntax, defs, operations, constraints }) =
   339   mk_class_data (f ((locale, consts, local_sort, inst, intro, local_syntax),
   340     (defs, (operations, constraints))));
   341 fun merge_class_data _ (ClassData { locale = locale, consts = consts,
   342     local_sort = local_sort, inst = inst, intro = intro, local_syntax = local_syntax,
   343     defs = defs1, operations = operations1, constraints = constraints1 },
   344   ClassData { locale = _, consts = _, local_sort = _, inst = _, intro = _, local_syntax = _,
   345     defs = defs2, operations = operations2, constraints = constraints2 }) =
   346   mk_class_data ((locale, consts, local_sort, inst, intro, local_syntax),
   347     (Thm.merge_thms (defs1, defs2),
   348       (AList.merge (op =) (K true) (operations1, operations2),
   349       AList.merge (op =) (K true) (constraints1, constraints2))));
   350 
   351 fun merge_pair f1 f2 ((x1, y1), (x2, y2)) = (f1 (x1, x2), f2 (y1, y2));
   352 
   353 structure ClassData = TheoryDataFun
   354 (
   355   type T = class_data Graph.T * class Symtab.table
   356     (*locale name ~> class name*);
   357   val empty = (Graph.empty, Symtab.empty);
   358   val copy = I;
   359   val extend = I;
   360   fun merge _ = merge_pair (Graph.join merge_class_data) (Symtab.merge (K true));
   361 );
   362 
   363 
   364 (* queries *)
   365 
   366 val lookup_class_data = Option.map rep_class_data oo try o Graph.get_node
   367   o fst o ClassData.get;
   368 fun class_of_locale thy = Symtab.lookup ((snd o ClassData.get) thy);
   369 
   370 fun the_class_data thy class = case lookup_class_data thy class
   371  of NONE => error ("undeclared class " ^ quote class)
   372   | SOME data => data;
   373 
   374 val locale_of_class = #locale oo the_class_data;
   375 
   376 val ancestry = Graph.all_succs o fst o ClassData.get;
   377 
   378 fun params_of_sort thy =
   379   let
   380     fun params class =
   381       let
   382         val const_typs = (#params o AxClass.get_definition thy) class;
   383         val const_names = (#consts o the_class_data thy) class;
   384       in
   385         (map o apsnd) (fn c => (c, (the o AList.lookup (op =) const_typs) c)) const_names
   386       end;
   387   in maps params o ancestry thy end;
   388 
   389 fun these_defs thy = maps (these o Option.map #defs o lookup_class_data thy) o ancestry thy;
   390 
   391 fun these_intros thy =
   392   Graph.fold (fn (_, (data, _)) => insert Thm.eq_thm ((#intro o rep_class_data) data))
   393     ((fst o ClassData.get) thy) [];
   394 
   395 fun these_operations thy =
   396   maps (#operations o the_class_data thy) o ancestry thy;
   397 
   398 fun these_constraints thy =
   399   maps (#constraints o the_class_data thy) o ancestry thy;
   400 
   401 fun local_operation thy = Option.join oo AList.lookup (op =) o these_operations thy;
   402 
   403 fun sups_local_sort thy sort =
   404   let
   405     val sups = filter (is_some o lookup_class_data thy) sort
   406       |> Sign.minimize_sort thy;
   407     val local_sort = case sups
   408      of sup :: _ => #local_sort (the_class_data thy sup)
   409       | [] => sort;
   410   in (sups, local_sort) end;
   411 
   412 fun local_syntax thy = #local_syntax o the_class_data thy;
   413 
   414 fun print_classes thy =
   415   let
   416     val ctxt = ProofContext.init thy;
   417 
   418     val algebra = Sign.classes_of thy;
   419     val arities =
   420       Symtab.empty
   421       |> Symtab.fold (fn (tyco, arities) => fold (fn (class, _) =>
   422            Symtab.map_default (class, []) (insert (op =) tyco)) arities)
   423              ((#arities o Sorts.rep_algebra) algebra);
   424     val the_arities = these o Symtab.lookup arities;
   425     fun mk_arity class tyco =
   426       let
   427         val Ss = Sorts.mg_domain algebra tyco [class];
   428       in Syntax.pretty_arity ctxt (tyco, Ss, [class]) end;
   429     fun mk_param (c, ty) = Pretty.str (Sign.extern_const thy c ^ " :: "
   430       ^ setmp show_sorts false (Syntax.string_of_typ ctxt o Type.strip_sorts) ty);
   431     fun mk_entry class = (Pretty.block o Pretty.fbreaks o map_filter I) [
   432       (SOME o Pretty.str) ("class " ^ class ^ ":"),
   433       (SOME o Pretty.block) [Pretty.str "supersort: ",
   434         (Syntax.pretty_sort ctxt o Sign.minimize_sort thy o Sign.super_classes thy) class],
   435       Option.map (Pretty.str o prefix "locale: " o #locale) (lookup_class_data thy class),
   436       ((fn [] => NONE | ps => (SOME o Pretty.block o Pretty.fbreaks) (Pretty.str "parameters:" :: ps)) o map mk_param
   437         o these o Option.map #params o try (AxClass.get_definition thy)) class,
   438       (SOME o Pretty.block o Pretty.breaks) [
   439         Pretty.str "instances:",
   440         Pretty.list "" "" (map (mk_arity class) (the_arities class))
   441       ]
   442     ]
   443   in
   444     (Pretty.writeln o Pretty.chunks o separate (Pretty.str "")
   445       o map mk_entry o Sorts.all_classes) algebra
   446   end;
   447 
   448 
   449 (* updaters *)
   450 
   451 fun add_class_data ((class, superclasses), (locale, consts, local_sort, inst, intro, local_syntax)) =
   452   ClassData.map (fn (gr, tab) => (
   453     gr
   454     |> Graph.new_node (class, mk_class_data ((locale, (map o pairself) fst consts,
   455         local_sort, inst, intro, local_syntax),
   456           ([], (map (apfst fst o apsnd (SOME o rpair 0 o Free) o swap) consts, map snd consts))))
   457     |> fold (curry Graph.add_edge class) superclasses,
   458     tab
   459     |> Symtab.update (locale, class)
   460   ));
   461 
   462 fun register_const class (entry, def) =
   463   (ClassData.map o apfst o Graph.map_node class o map_class_data o apsnd)
   464     (fn (defs, (operations, constraints)) =>
   465       (def :: defs, (apsnd (SOME o snd) entry :: operations, apsnd fst entry :: constraints)));
   466 
   467 fun register_abbrev class (abbrev, ty) =
   468   (ClassData.map o apfst o Graph.map_node class o map_class_data o apsnd o apsnd)
   469     (fn (operations, constraints) => ((abbrev, NONE) :: operations, (abbrev, ty) :: constraints));
   470 
   471 
   472 (** rule calculation, tactics and methods **)
   473 
   474 fun class_intro thy locale class sups =
   475   let
   476     fun class_elim class =
   477       case (map Drule.unconstrainTs o #axioms o AxClass.get_definition thy) class
   478        of [thm] => SOME thm
   479         | [] => NONE;
   480     val pred_intro = case Locale.intros thy locale
   481      of ([ax_intro], [intro]) => intro |> OF_LAST ax_intro |> SOME
   482       | ([intro], []) => SOME intro
   483       | ([], [intro]) => SOME intro
   484       | _ => NONE;
   485     val pred_intro' = pred_intro
   486       |> Option.map (fn intro => intro OF map_filter class_elim sups);
   487     val class_intro = (#intro o AxClass.get_definition thy) class;
   488     val raw_intro = case pred_intro'
   489      of SOME pred_intro => class_intro |> OF_LAST pred_intro
   490       | NONE => class_intro;
   491     val sort = Sign.super_classes thy class;
   492     val typ = TVar ((Name.aT, 0), sort);
   493     val defs = these_defs thy sups;
   494   in
   495     raw_intro
   496     |> Drule.instantiate' [SOME (Thm.ctyp_of thy typ)] []
   497     |> strip_all_ofclass thy sort
   498     |> Thm.strip_shyps
   499     |> MetaSimplifier.rewrite_rule defs
   500     |> Drule.unconstrainTs
   501   end;
   502 
   503 fun class_interpretation class facts defs thy =
   504   let
   505     val { locale, inst, ... } = the_class_data thy class;
   506     val tac = (ALLGOALS o ProofContext.fact_tac) facts;
   507     val prfx = Logic.const_of_class (NameSpace.base class);
   508   in
   509     prove_interpretation tac ((false, prfx), []) (Locale.Locale locale)
   510       (inst, defs) thy
   511   end;
   512 
   513 fun interpretation_in_rule thy (class1, class2) =
   514   let
   515     fun mk_axioms class =
   516       let
   517         val { locale, inst = (_, insttab), ... } = the_class_data thy class;
   518       in
   519         Locale.global_asms_of thy locale
   520         |> maps snd
   521         |> (map o map_aterms) (fn Free (s, _) => (the o Symtab.lookup insttab) s | t => t)
   522         |> (map o map_types o map_atyps) (fn TFree _ => TFree (Name.aT, [class1]) | T => T)
   523         |> map (ObjectLogic.ensure_propT thy)
   524       end;
   525     val (prems, concls) = pairself mk_axioms (class1, class2);
   526   in
   527     Goal.prove_global thy [] prems (Logic.mk_conjunction_list concls)
   528       (Locale.intro_locales_tac true (ProofContext.init thy))
   529   end;
   530 
   531 fun intro_classes_tac facts st =
   532   let
   533     val thy = Thm.theory_of_thm st;
   534     val classes = Sign.all_classes thy;
   535     val class_trivs = map (Thm.class_triv thy) classes;
   536     val class_intros = these_intros thy;
   537     fun add_axclass_intro class =
   538       case try (AxClass.get_definition thy) class of SOME {intro, ...} => cons intro | _ => I;
   539     val axclass_intros = fold add_axclass_intro classes [];
   540   in
   541     st
   542     |> ((ALLGOALS (Method.insert_tac facts THEN'
   543           REPEAT_ALL_NEW (resolve_tac (class_trivs @ class_intros @ axclass_intros))))
   544             THEN Tactic.distinct_subgoals_tac)
   545   end;
   546 
   547 fun default_intro_classes_tac [] = intro_classes_tac []
   548   | default_intro_classes_tac _ = Tactical.no_tac;    (*no error message!*)
   549 
   550 fun default_tac rules ctxt facts =
   551   HEADGOAL (Method.some_rule_tac rules ctxt facts) ORELSE
   552     default_intro_classes_tac facts;
   553 
   554 val _ = Context.add_setup (Method.add_methods
   555  [("intro_classes", Method.no_args (Method.METHOD intro_classes_tac),
   556     "back-chain introduction rules of classes"),
   557   ("default", Method.thms_ctxt_args (Method.METHOD oo default_tac),
   558     "apply some intro/elim rule")]);
   559 
   560 
   561 (** classes and class target **)
   562 
   563 (* class context initialization *)
   564 
   565 fun remove_constraint local_sort c ctxt =
   566   let
   567     val ty = ProofContext.the_const_constraint ctxt c;
   568     val ty' = map_atyps (fn ty as TFree (v, _) => if v = Name.aT
   569       then TFree (v, local_sort) else ty | ty => ty) ty;
   570   in
   571     ctxt
   572     |> ProofContext.add_const_constraint (c, SOME ty')
   573   end;
   574 
   575 fun sort_term_check sups local_sort ts ctxt =
   576   let
   577     val thy = ProofContext.theory_of ctxt;
   578     val local_operation = local_operation thy sups;
   579     val constraints = these_constraints thy sups;
   580     val consts = ProofContext.consts_of ctxt;
   581     fun check_typ (c, ty) (t', idx) = case nth (Consts.typargs consts (c, ty)) idx
   582      of TFree (v, _) => if v = Name.aT
   583           then apfst (AList.update (op =) ((c, ty), t')) else I
   584       | TVar (vi, _) => if TypeInfer.is_param vi
   585           then apfst (AList.update (op =) ((c, ty), t'))
   586             #> apsnd (insert (op =) vi) else I
   587       | _ => I;
   588     fun add_const (Const (c, ty)) = (case local_operation c
   589          of SOME (t', idx) => check_typ (c, ty) (t', idx)
   590           | NONE => I)
   591       | add_const _ = I;
   592     val (cs, typarams) = (fold o fold_aterms) add_const ts ([], []);
   593     val ts' = map (map_aterms
   594         (fn t as Const (c, ty) => the_default t (AList.lookup (op =) cs (c, ty)) | t => t)
   595       #> (map_types o map_type_tvar) (fn var as (vi, _) => if member (op =) typarams vi
   596            then TFree (Name.aT, local_sort) else TVar var)) ts;
   597     val ctxt' = fold (ProofContext.add_const_constraint o apsnd SOME) constraints ctxt;
   598   in (ts', ctxt') end;
   599 
   600 fun init_class_ctxt sups local_sort ctxt =
   601   let
   602     val operations = these_operations (ProofContext.theory_of ctxt) sups;
   603   in
   604     ctxt
   605     |> Variable.declare_term
   606         (Logic.mk_type (TFree (Name.aT, local_sort)))
   607     |> fold (remove_constraint local_sort o fst) operations
   608     |> Context.proof_map (Syntax.add_term_check 50 "class"
   609         (sort_term_check sups local_sort))
   610   end;
   611 
   612 fun init class ctxt =
   613   init_class_ctxt [class]
   614     ((#local_sort o the_class_data (ProofContext.theory_of ctxt)) class) ctxt;
   615 
   616 
   617 (* class definition *)
   618 
   619 local
   620 
   621 fun read_param thy raw_t =  (* FIXME ProofContext.read_const (!?) *)
   622   let
   623     val t = Syntax.read_term_global thy raw_t
   624   in case try dest_Const t
   625    of SOME (c, _) => c
   626     | NONE => error ("Not a constant: " ^ Sign.string_of_term thy t)
   627   end;
   628 
   629 fun gen_class_spec prep_class prep_expr process_expr thy raw_supclasses raw_includes_elems =
   630   let
   631     val supclasses = map (prep_class thy) raw_supclasses;
   632     val (sups, local_sort) = sups_local_sort thy supclasses;
   633     val supsort = Sign.minimize_sort thy supclasses;
   634     val suplocales = map (Locale.Locale o #locale o the_class_data thy) sups;
   635     val (raw_elems, includes) = fold_rev (fn Locale.Elem e => apfst (cons e)
   636       | Locale.Expr i => apsnd (cons (prep_expr thy i))) raw_includes_elems ([], []);
   637     val supexpr = Locale.Merge suplocales;
   638     val supparams = (map fst o Locale.parameters_of_expr thy) supexpr;
   639     val supconsts = AList.make (the o AList.lookup (op =) (params_of_sort thy sups))
   640       (map fst supparams);
   641     val mergeexpr = Locale.Merge (suplocales @ includes);
   642     val constrain = Element.Constrains ((map o apsnd o map_atyps)
   643       (fn TFree (_, sort) => TFree (Name.aT, sort)) supparams);
   644   in
   645     ProofContext.init thy
   646     |> Locale.cert_expr supexpr [constrain]
   647     |> snd
   648     |> init_class_ctxt sups local_sort
   649     |> process_expr Locale.empty raw_elems
   650     |> fst
   651     |> (fn elems => ((((sups, supconsts), (supsort, local_sort, mergeexpr)),
   652           (*FIXME*) if null includes then constrain :: elems else elems)))
   653   end;
   654 
   655 val read_class_spec = gen_class_spec Sign.intern_class Locale.intern_expr Locale.read_expr;
   656 val check_class_spec = gen_class_spec (K I) (K I) Locale.cert_expr;
   657 
   658 fun gen_class prep_spec prep_param local_syntax bname
   659     raw_supclasses raw_includes_elems raw_other_consts thy =
   660   let
   661     val (((sups, supconsts), (supsort, local_sort, mergeexpr)), elems) =
   662       prep_spec thy raw_supclasses raw_includes_elems;
   663     val other_consts = map (prep_param thy) raw_other_consts;
   664     fun mk_instT class = Symtab.empty
   665       |> Symtab.update (Name.aT, TFree (Name.aT, [class]));
   666     fun mk_inst class param_names cs =
   667       Symtab.empty
   668       |> fold2 (fn v => fn (c, ty) => Symtab.update (v, Const
   669            (c, Term.map_type_tfree (fn (v, _) => TFree (v, [class])) ty))) param_names cs;
   670     fun extract_params thy name_locale =
   671       let
   672         val params = Locale.parameters_of thy name_locale;
   673         val _ = if Sign.subsort thy (supsort, local_sort) then () else error
   674           ("Sort " ^ Sign.string_of_sort thy local_sort
   675             ^ " is less general than permitted least general sort "
   676             ^ Sign.string_of_sort thy supsort);
   677       in
   678         (map fst params, params
   679         |> (map o apfst o apsnd o Term.map_type_tfree) (K (TFree (Name.aT, local_sort)))
   680         |> (map o apsnd) (fork_mixfix true (SOME "") #> fst)
   681         |> chop (length supconsts)
   682         |> snd)
   683       end;
   684     fun extract_assumes name_locale params thy cs =
   685       let
   686         val consts = supconsts @ (map (fst o fst) params ~~ cs);
   687         fun subst (Free (c, ty)) =
   688               Const ((fst o the o AList.lookup (op =) consts) c, ty)
   689           | subst t = t;
   690         fun prep_asm ((name, atts), ts) =
   691           ((NameSpace.base name, map (Attrib.attribute thy) atts),
   692             (map o map_aterms) subst ts);
   693       in
   694         Locale.global_asms_of thy name_locale
   695         |> map prep_asm
   696       end;
   697     fun note_intro name_axclass class_intro =
   698       PureThy.note_thmss_qualified "" ((Logic.const_of_class o NameSpace.base) name_axclass)
   699         [(("intro", []), [([class_intro], [])])]
   700       #> snd
   701   in
   702     thy
   703     |> Locale.add_locale_i (SOME "") bname mergeexpr elems
   704     |-> (fn name_locale => ProofContext.theory_result (
   705       `(fn thy => extract_params thy name_locale)
   706       #-> (fn (globals, params) =>
   707         AxClass.define_class_params (bname, supsort) params
   708           (extract_assumes name_locale params) other_consts
   709       #-> (fn (name_axclass, (consts, axioms)) =>
   710         `(fn thy => class_intro thy name_locale name_axclass sups)
   711       #-> (fn class_intro =>
   712         add_class_data ((name_axclass, sups),
   713           (name_locale, map fst params ~~ consts, local_sort,
   714             (mk_instT name_axclass, mk_inst name_axclass (map fst globals)
   715               (map snd supconsts @ consts)), class_intro, local_syntax))
   716       #> note_intro name_axclass class_intro
   717       #> class_interpretation name_axclass axioms []
   718       #> pair name_axclass
   719       )))))
   720   end;
   721 
   722 in
   723 
   724 val class_cmd = gen_class read_class_spec read_param;
   725 val class = gen_class check_class_spec (K I);
   726 
   727 end; (*local*)
   728 
   729 
   730 (* definition in class target *)
   731 
   732 fun export_fixes thy class =
   733   let
   734     val consts = params_of_sort thy [class];
   735     fun subst_aterm (t as Free (v, ty)) = (case AList.lookup (op =) consts v
   736          of SOME (c, _) => Const (c, ty)
   737           | NONE => t)
   738       | subst_aterm t = t;
   739   in Term.map_aterms subst_aterm end;
   740 
   741 fun mk_operation_entry thy (c, rhs) =
   742   let
   743     val ty = fastype_of rhs;
   744     val typargs = Sign.const_typargs thy (c, ty);
   745     val typidx = find_index (fn TFree (w, _) => Name.aT = w | _ => false) typargs;
   746   in (c, (ty, (rhs, typidx))) end;
   747 
   748 fun add_const_in_class class ((c, rhs), syn) thy =
   749   let
   750     val prfx = (Logic.const_of_class o NameSpace.base) class;
   751     fun mk_name c =
   752       let
   753         val n1 = Sign.full_name thy c;
   754         val n2 = NameSpace.qualifier n1;
   755         val n3 = NameSpace.base n1;
   756       in NameSpace.implode [n2, prfx, n3] end;
   757     val constrain_sort = curry (Sorts.inter_sort (Sign.classes_of thy)) [class];
   758     val rhs' = export_fixes thy class rhs;
   759     val subst_typ = Term.map_type_tfree (fn var as (w, sort) =>
   760       if w = Name.aT then TFree (w, constrain_sort sort) else TFree var);
   761     val ty' = Term.fastype_of rhs';
   762     val ty'' = subst_typ ty';
   763     val c' = mk_name c;
   764     val def = (c, Logic.mk_equals (Const (c', ty'), rhs'));
   765     val (syn', _) = fork_mixfix true (SOME class) syn;
   766     fun interpret def thy =
   767       let
   768         val def' = symmetric def;
   769         val def_eq = Thm.prop_of def';
   770         val entry = mk_operation_entry thy (c', rhs);
   771       in
   772         thy
   773         |> class_interpretation class [def'] [def_eq]
   774         |> register_const class (entry, def')
   775       end;
   776   in
   777     thy
   778     |> Sign.add_path prfx
   779     |> Sign.add_consts_authentic [] [(c, ty', syn')]
   780     |> Sign.parent_path
   781     |> Sign.sticky_prefix prfx
   782     |> PureThy.add_defs_i false [(def, [])]
   783     |-> (fn [def] => interpret def)
   784     |> Sign.add_const_constraint (c', SOME ty'')
   785     |> Sign.restore_naming thy
   786     |> pair c'
   787   end;
   788 
   789 
   790 (* abbreviation in class target *)
   791 
   792 fun add_abbrev_in_class class prmode ((c, rhs), syn) thy =
   793   let
   794     val prfx = (Logic.const_of_class o NameSpace.base) class;
   795     fun mk_name c =
   796       let
   797         val n1 = Sign.full_name thy c;
   798         val n2 = NameSpace.qualifier n1;
   799         val n3 = NameSpace.base n1;
   800       in NameSpace.implode [n2, prfx, prfx, n3] end;
   801     val c' = mk_name c;
   802     val rhs' = export_fixes thy class rhs;
   803     val ty' = fastype_of rhs';
   804   in
   805     thy
   806     |> Sign.add_notation prmode [(Const (c', ty'), syn)]
   807     |> register_abbrev class (c', ty')
   808     |> pair c'
   809   end;
   810 
   811 end;