rearranged target theories
authorhaftmann
Mon, 05 Jan 2009 15:36:24 +0100
changeset 29358 efdfe5dfe008
parent 29357 11956fa598b7
child 29359 f831192b9366
rearranged target theories
src/Pure/IsaMakefile
src/Pure/Isar/ROOT.ML
src/Pure/Isar/class.ML
src/Pure/Isar/class_target.ML
src/Pure/Isar/expression.ML
src/Pure/Isar/instance.ML
src/Pure/Isar/isar_syn.ML
src/Pure/Isar/subclass.ML
src/Pure/Isar/theory_target.ML
--- a/src/Pure/IsaMakefile	Mon Jan 05 15:35:42 2009 +0100
+++ b/src/Pure/IsaMakefile	Mon Jan 05 15:36:24 2009 +0100
@@ -36,10 +36,10 @@
   General/stack.ML General/symbol.ML General/symbol_pos.ML		\
   General/table.ML General/url.ML General/xml.ML General/yxml.ML	\
   Isar/ROOT.ML Isar/antiquote.ML Isar/args.ML Isar/attrib.ML		\
-  Isar/auto_bind.ML Isar/calculation.ML Isar/class.ML Isar/code.ML	\
+  Isar/auto_bind.ML Isar/calculation.ML Isar/class.ML Isar/class_target.ML Isar/code.ML	\
   Isar/code_unit.ML Isar/constdefs.ML Isar/context_rules.ML		\
   Isar/element.ML Isar/expression.ML Isar/find_theorems.ML		\
-  Isar/instance.ML Isar/isar.ML Isar/isar_cmd.ML Isar/isar_syn.ML	\
+  Isar/isar.ML Isar/isar_cmd.ML Isar/isar_syn.ML	\
   Isar/local_defs.ML Isar/local_syntax.ML Isar/local_theory.ML		\
   Isar/locale.ML Isar/method.ML Isar/net_rules.ML Isar/new_locale.ML	\
   Isar/object_logic.ML Isar/obtain.ML Isar/outer_keyword.ML		\
@@ -47,7 +47,7 @@
   Isar/overloading.ML Isar/proof.ML Isar/proof_context.ML		\
   Isar/proof_display.ML Isar/proof_node.ML Isar/rule_cases.ML		\
   Isar/rule_insts.ML Isar/session.ML Isar/skip_proof.ML			\
-  Isar/spec_parse.ML Isar/specification.ML Isar/subclass.ML		\
+  Isar/spec_parse.ML Isar/specification.ML		\
   Isar/theory_target.ML Isar/toplevel.ML Isar/value_parse.ML		\
   ML-Systems/alice.ML ML-Systems/exn.ML					\
   ML-Systems/install_pp_polyml.ML ML-Systems/ml_name_space.ML		\
--- a/src/Pure/Isar/ROOT.ML	Mon Jan 05 15:35:42 2009 +0100
+++ b/src/Pure/Isar/ROOT.ML	Mon Jan 05 15:36:24 2009 +0100
@@ -55,11 +55,10 @@
 use "overloading.ML";
 use "locale.ML";
 use "new_locale.ML";
-use "class.ML";
+use "class_target.ML";
 use "theory_target.ML";
 use "expression.ML";
-use "instance.ML";
-use "subclass.ML";
+use "class.ML";
 
 (*complex proof machineries*)
 use "../simplifier.ML";
--- a/src/Pure/Isar/class.ML	Mon Jan 05 15:35:42 2009 +0100
+++ b/src/Pure/Isar/class.ML	Mon Jan 05 15:36:24 2009 +0100
@@ -1,361 +1,30 @@
-(*  Title:      Pure/Isar/class.ML
-    ID:         $Id$
+(*  Title:      Pure/Isar/ML
     Author:     Florian Haftmann, TU Muenchen
 
-Type classes derived from primitive axclasses and locales.
+Type classes derived from primitive axclasses and locales - interfaces
 *)
 
 signature CLASS =
 sig
-  (*classes*)
+  include CLASS_TARGET
+    (*FIXME the split in class_target.ML, theory_target.ML and
+      ML is artificial*)
+
   val class: bstring -> class list -> Element.context_i list
     -> theory -> string * Proof.context
   val class_cmd: bstring -> xstring list -> Element.context list
     -> theory -> string * Proof.context
-
-  val init: class -> theory -> Proof.context
-  val declare: class -> Properties.T
-    -> (string * mixfix) * term -> theory -> theory
-  val abbrev: class -> Syntax.mode -> Properties.T
-    -> (string * mixfix) * term -> theory -> theory
-  val refresh_syntax: class -> Proof.context -> Proof.context
-
-  val intro_classes_tac: thm list -> tactic
-  val default_intro_tac: Proof.context -> thm list -> tactic
-  val prove_subclass: class * class -> thm option -> theory -> theory
-
-  val class_prefix: string -> string
-  val is_class: theory -> class -> bool
-  val these_params: theory -> sort -> (string * (class * (string * typ))) list
-  val print_classes: theory -> unit
-
-  (*instances*)
-  val init_instantiation: string list * (string * sort) list * sort
-    -> theory -> local_theory
-  val instantiation_instance: (local_theory -> local_theory)
-    -> local_theory -> Proof.state
-  val prove_instantiation_instance: (Proof.context -> tactic)
-    -> local_theory -> local_theory
-  val prove_instantiation_exit: (Proof.context -> tactic)
-    -> local_theory -> theory
-  val prove_instantiation_exit_result: (morphism -> 'a -> 'b)
-    -> (Proof.context -> 'b -> tactic) -> 'a -> local_theory -> 'b * theory
-  val conclude_instantiation: local_theory -> local_theory
-  val instantiation_param: local_theory -> string -> string option
-  val confirm_declaration: string -> local_theory -> local_theory
-  val pretty_instantiation: local_theory -> Pretty.T
-  val type_name: string -> string
-
-  (*old axclass layer*)
-  val axclass_cmd: bstring * xstring list
-    -> (Attrib.binding * string list) list
-    -> theory -> class * theory
-  val classrel_cmd: xstring * xstring -> theory -> Proof.state
-
-  (*old instance layer*)
-  val instance_arity: (theory -> theory) -> arity -> theory -> Proof.state
-  val instance_arity_cmd: bstring * xstring list * xstring -> theory -> Proof.state
+  val prove_subclass: tactic -> class -> local_theory -> local_theory
+  val subclass: class -> local_theory -> Proof.state
+  val subclass_cmd: xstring -> local_theory -> Proof.state
 end;
 
 structure Class : CLASS =
 struct
 
-(*temporary adaption code to mediate between old and new locale code*)
-
-structure Old_Locale =
-struct
-
-val intro_locales_tac = Locale.intro_locales_tac; (*already forked!*)
-
-val interpretation = Locale.interpretation;
-val interpretation_in_locale = Locale.interpretation_in_locale;
-val get_interpret_morph = Locale.get_interpret_morph;
-val Locale = Locale.Locale;
-val extern = Locale.extern;
-val intros = Locale.intros;
-val dests = Locale.dests;
-val init = Locale.init;
-val Merge = Locale.Merge;
-val parameters_of_expr = Locale.parameters_of_expr;
-val empty = Locale.empty;
-val cert_expr = Locale.cert_expr;
-val read_expr = Locale.read_expr;
-val parameters_of = Locale.parameters_of;
-val add_locale = Locale.add_locale;
-
-end;
-
-(*structure New_Locale =
-struct
-
-val intro_locales_tac = Locale.intro_locales_tac; (*already forked!*)
-
-val interpretation = Locale.interpretation; (*!*)
-val interpretation_in_locale = Locale.interpretation_in_locale; (*!*)
-val get_interpret_morph = Locale.get_interpret_morph; (*!*)
-fun Locale loc = ([(loc, ("", Expression.Positional []))], []);
-val extern = NewLocale.extern;
-val intros = Locale.intros; (*!*)
-val dests = Locale.dests; (*!*)
-val init = NewLocale.init;
-fun Merge locs = (map (fn loc => (loc, ("", Expression.Positional []))) locs, []);
-val parameters_of_expr = Locale.parameters_of_expr; (*!*)
-val empty = ([], []);
-val cert_expr = Locale.cert_expr; (*!"*)
-val read_expr = Locale.read_expr; (*!"*)
-val parameters_of = NewLocale.params_of; (*why typ option?*)
-val add_locale = Expression.add_locale;
-
-end;*)
-
-structure Locale = Old_Locale;
-
-(*proper code again*)
-
-
-(** auxiliary **)
-
-fun prove_interpretation tac prfx_atts expr inst =
-  Locale.interpretation I prfx_atts expr inst
-  ##> Proof.global_terminal_proof
-      (Method.Basic (fn ctxt => Method.SIMPLE_METHOD (tac ctxt), Position.none), NONE)
-  ##> ProofContext.theory_of;
-
-fun prove_interpretation_in tac after_qed (name, expr) =
-  Locale.interpretation_in_locale
-      (ProofContext.theory after_qed) (name, expr)
-  #> Proof.global_terminal_proof
-      (Method.Basic (K (Method.SIMPLE_METHOD tac), Position.none), NONE)
-  #> ProofContext.theory_of;
-
-val class_prefix = Logic.const_of_class o Sign.base_name;
-
-fun class_name_morphism class =
-  Binding.map_prefix (K (Binding.add_prefix false (class_prefix class)));
-
-fun activate_class_morphism thy class inst morphism =
-  Locale.get_interpret_morph thy (class_name_morphism class) (class, "") morphism class inst;
-
-fun prove_class_interpretation class inst consts hyp_facts def_facts thy =
-  let
-    val constraints = map (fn (class, c) => map_atyps (K (TFree (Name.aT,
-      [class]))) (Sign.the_const_type thy c)) consts;
-    val no_constraints = map (map_atyps (K (TFree (Name.aT, [])))) constraints;
-    fun add_constraint c T = Sign.add_const_constraint (c, SOME T);
-    fun tac ctxt = ALLGOALS (ProofContext.fact_tac (hyp_facts @ def_facts)
-      ORELSE' (fn n => SELECT_GOAL (Locale.intro_locales_tac false ctxt []) n));
-    val prfx = class_prefix class;
-  in
-    thy
-    |> fold2 add_constraint (map snd consts) no_constraints
-    |> prove_interpretation tac (class_name_morphism class) (Locale.Locale class)
-          (map SOME inst, map (pair (Attrib.empty_binding) o Thm.prop_of) def_facts)
-    ||> fold2 add_constraint (map snd consts) constraints
-  end;
-
-
-(** primitive axclass and instance commands **)
-
-fun axclass_cmd (class, raw_superclasses) raw_specs thy =
-  let
-    val ctxt = ProofContext.init thy;
-    val superclasses = map (Sign.read_class thy) raw_superclasses;
-    val name_atts = map ((apsnd o map) (Attrib.attribute thy) o fst)
-      raw_specs;
-    val axiomss = ProofContext.read_propp (ctxt, map (map (rpair []) o snd)
-          raw_specs)
-      |> snd
-      |> (map o map) fst;
-  in
-    AxClass.define_class (class, superclasses) []
-      (name_atts ~~ axiomss) thy
-  end;
-
-local
-
-fun gen_instance mk_prop add_thm after_qed insts thy =
-  let
-    fun after_qed' results =
-      ProofContext.theory ((fold o fold) add_thm results #> after_qed);
-  in
-    thy
-    |> ProofContext.init
-    |> Proof.theorem_i NONE after_qed' ((map (fn t => [(t, [])])
-        o mk_prop thy) insts)
-  end;
-
-in
-
-val instance_arity =
-  gen_instance (Logic.mk_arities oo Sign.cert_arity) AxClass.add_arity;
-val instance_arity_cmd =
-  gen_instance (Logic.mk_arities oo Sign.read_arity) AxClass.add_arity I;
-val classrel =
-  gen_instance (single oo (Logic.mk_classrel oo AxClass.cert_classrel)) AxClass.add_classrel I;
-val classrel_cmd =
-  gen_instance (single oo (Logic.mk_classrel oo AxClass.read_classrel)) AxClass.add_classrel I;
-
-end; (*local*)
-
-
-(** class data **)
-
-datatype class_data = ClassData of {
+open Class_Target;
 
-  (* static part *)
-  consts: (string * string) list
-    (*locale parameter ~> constant name*),
-  base_sort: sort,
-  inst: term list
-    (*canonical interpretation*),
-  morphism: Morphism.morphism * ((typ Vartab.table * typ list) * (term Vartab.table * term list))
-    (*morphism cookie of canonical interpretation*),
-  assm_intro: thm option,
-  of_class: thm,
-  axiom: thm option,
-  
-  (* dynamic part *)
-  defs: thm list,
-  operations: (string * (class * (typ * term))) list
-
-};
-
-fun rep_class_data (ClassData d) = d;
-fun mk_class_data ((consts, base_sort, inst, morphism, assm_intro, of_class, axiom),
-    (defs, operations)) =
-  ClassData { consts = consts, base_sort = base_sort, inst = inst,
-    morphism = morphism, assm_intro = assm_intro, of_class = of_class, axiom = axiom,
-    defs = defs, operations = operations };
-fun map_class_data f (ClassData { consts, base_sort, inst, morphism, assm_intro,
-    of_class, axiom, defs, operations }) =
-  mk_class_data (f ((consts, base_sort, inst, morphism, assm_intro, of_class, axiom),
-    (defs, operations)));
-fun merge_class_data _ (ClassData { consts = consts,
-    base_sort = base_sort, inst = inst, morphism = morphism, assm_intro = assm_intro,
-    of_class = of_class, axiom = axiom, defs = defs1, operations = operations1 },
-  ClassData { consts = _, base_sort = _, inst = _, morphism = _, assm_intro = _,
-    of_class = _, axiom = _, defs = defs2, operations = operations2 }) =
-  mk_class_data ((consts, base_sort, inst, morphism, assm_intro, of_class, axiom),
-    (Thm.merge_thms (defs1, defs2),
-      AList.merge (op =) (K true) (operations1, operations2)));
-
-structure ClassData = TheoryDataFun
-(
-  type T = class_data Graph.T
-  val empty = Graph.empty;
-  val copy = I;
-  val extend = I;
-  fun merge _ = Graph.join merge_class_data;
-);
-
-
-(* queries *)
-
-val lookup_class_data = Option.map rep_class_data oo try o Graph.get_node o ClassData.get;
-
-fun the_class_data thy class = case lookup_class_data thy class
- of NONE => error ("Undeclared class " ^ quote class)
-  | SOME data => data;
-
-val is_class = is_some oo lookup_class_data;
-
-val ancestry = Graph.all_succs o ClassData.get;
-
-fun the_inst thy = #inst o the_class_data thy;
-
-fun these_params thy =
-  let
-    fun params class =
-      let
-        val const_typs = (#params o AxClass.get_info thy) class;
-        val const_names = (#consts o the_class_data thy) class;
-      in
-        (map o apsnd)
-          (fn c => (class, (c, (the o AList.lookup (op =) const_typs) c))) const_names
-      end;
-  in maps params o ancestry thy end;
-
-fun these_assm_intros thy =
-  Graph.fold (fn (_, (data, _)) => fold (insert Thm.eq_thm)
-    ((the_list o #assm_intro o rep_class_data) data)) (ClassData.get thy) [];
-
-fun these_defs thy = maps (these o Option.map #defs o lookup_class_data thy) o ancestry thy;
-
-fun these_operations thy =
-  maps (#operations o the_class_data thy) o ancestry thy;
-
-fun morphism thy class =
-  let
-    val { inst, morphism, ... } = the_class_data thy class;
-  in activate_class_morphism thy class inst morphism end;
-
-fun print_classes thy =
-  let
-    val ctxt = ProofContext.init thy;
-    val algebra = Sign.classes_of thy;
-    val arities =
-      Symtab.empty
-      |> Symtab.fold (fn (tyco, arities) => fold (fn (class, _) =>
-           Symtab.map_default (class, []) (insert (op =) tyco)) arities)
-             ((#arities o Sorts.rep_algebra) algebra);
-    val the_arities = these o Symtab.lookup arities;
-    fun mk_arity class tyco =
-      let
-        val Ss = Sorts.mg_domain algebra tyco [class];
-      in Syntax.pretty_arity ctxt (tyco, Ss, [class]) end;
-    fun mk_param (c, ty) = Pretty.str (Sign.extern_const thy c ^ " :: "
-      ^ setmp show_sorts false (Syntax.string_of_typ ctxt o Type.strip_sorts) ty);
-    fun mk_entry class = (Pretty.block o Pretty.fbreaks o map_filter I) [
-      (SOME o Pretty.str) ("class " ^ Sign.extern_class thy class ^ ":"),
-      (SOME o Pretty.block) [Pretty.str "supersort: ",
-        (Syntax.pretty_sort ctxt o Sign.minimize_sort thy o Sign.super_classes thy) class],
-      if is_class thy class then (SOME o Pretty.str)
-        ("locale: " ^ Locale.extern thy class) else NONE,
-      ((fn [] => NONE | ps => (SOME o Pretty.block o Pretty.fbreaks)
-          (Pretty.str "parameters:" :: ps)) o map mk_param
-        o these o Option.map #params o try (AxClass.get_info thy)) class,
-      (SOME o Pretty.block o Pretty.breaks) [
-        Pretty.str "instances:",
-        Pretty.list "" "" (map (mk_arity class) (the_arities class))
-      ]
-    ]
-  in
-    (Pretty.writeln o Pretty.chunks o separate (Pretty.str "")
-      o map mk_entry o Sorts.all_classes) algebra
-  end;
-
-
-(* updaters *)
-
-fun add_class_data ((class, superclasses),
-    (params, base_sort, inst, phi, axiom, assm_intro, of_class)) thy =
-  let
-    val operations = map (fn (v_ty as (_, ty), (c, _)) =>
-      (c, (class, (ty, Free v_ty)))) params;
-    val add_class = Graph.new_node (class,
-        mk_class_data (((map o pairself) fst params, base_sort,
-          inst, phi, assm_intro, of_class, axiom), ([], operations)))
-      #> fold (curry Graph.add_edge class) superclasses;
-  in ClassData.map add_class thy end;
-
-fun register_operation class (c, (t, some_def)) thy =
-  let
-    val base_sort = (#base_sort o the_class_data thy) class;
-    val prep_typ = map_type_tvar
-      (fn (vi as (v, _), sort) => if Name.aT = v
-        then TFree (v, base_sort) else TVar (vi, sort));
-    val t' = map_types prep_typ t;
-    val ty' = Term.fastype_of t';
-  in
-    thy
-    |> (ClassData.map o Graph.map_node class o map_class_data o apsnd)
-      (fn (defs, operations) =>
-        (fold cons (the_list some_def) defs,
-          (c, (class, (ty', t'))) :: operations))
-  end;
-
-
-(** rule calculation, tactics and methods **)
+(** rule calculation **)
 
 fun calculate_axiom thy sups base_sort assm_axiom param_map class =
   case Locale.intros thy class
@@ -366,7 +35,7 @@
           (base_sort, sort)], map (fn (v, (c, ty)) => pairself (Thm.cterm_of thy)
             (Var ((v, 0), map_atyps (fn _ => TVar ((Name.aT, 0), sort)) ty),
               Const (c, map_atyps (fn _ => TVar ((Name.aT, 0), sort)) ty))) param_map);
-        val axiom_premises = map_filter (#axiom o the_class_data thy) sups
+        val axiom_premises = map_filter (fst o rules thy) sups
           @ the_list assm_axiom;
       in intro
         |> instantiate thy [class]
@@ -394,7 +63,7 @@
         (TVar ((Name.aT, 0), base_sort), TFree (Name.aT, base_sort))], [])
     val of_class_sups = if null sups
       then map (fixate o Thm.class_triv thy) base_sort
-      else map (fixate o #of_class o the_class_data thy) sups;
+      else map (fixate o snd o rules thy) sups;
     val locale_dests = map Drule.standard' (Locale.dests thy class);
     val num_trivs = case length locale_dests
      of 0 => if is_none axiom then 0 else 1
@@ -411,23 +80,6 @@
       |> Thm.close_derivation;
   in (assm_intro, of_class) end;
 
-fun prove_subclass (sub, sup) some_thm thy =
-  let
-    val of_class = (#of_class o the_class_data thy) sup;
-    val intro = case some_thm
-     of SOME thm => Drule.standard' (of_class OF [Drule.standard' thm])
-      | NONE => Thm.instantiate ([pairself (Thm.ctyp_of thy o TVar o pair (Name.aT, 0))
-          ([], [sub])], []) of_class;
-    val classrel = (intro OF (the_list o #axiom o the_class_data thy) sub)
-      |> Thm.close_derivation;
-  in
-    thy
-    |> AxClass.add_classrel classrel
-    |> prove_interpretation_in (ALLGOALS (ProofContext.fact_tac (the_list some_thm)))
-         I (sub, Locale.Locale sup)
-    |> ClassData.map (Graph.add_edge (sub, sup))
-  end;
-
 fun note_assm_intro class assm_intro thy =
   thy
   |> Sign.sticky_prefix (class_prefix class ^ "_" ^ AxClass.axiomsN)
@@ -435,147 +87,8 @@
   |> snd
   |> Sign.restore_naming thy;
 
-fun intro_classes_tac facts st =
-  let
-    val thy = Thm.theory_of_thm st;
-    val classes = Sign.all_classes thy;
-    val class_trivs = map (Thm.class_triv thy) classes;
-    val class_intros = map_filter (try (#intro o AxClass.get_info thy)) classes;
-    val assm_intros = these_assm_intros thy;
-  in
-    Method.intros_tac (class_trivs @ class_intros @ assm_intros) facts st
-  end;
 
-fun default_intro_tac ctxt [] =
-      intro_classes_tac [] ORELSE NewLocale.intro_locales_tac true ctxt [] ORELSE
-      Locale.intro_locales_tac true ctxt []
-  | default_intro_tac _ _ = no_tac;
-
-fun default_tac rules ctxt facts =
-  HEADGOAL (Method.some_rule_tac rules ctxt facts) ORELSE
-    default_intro_tac ctxt facts;
-
-val _ = Context.>> (Context.map_theory
-  (Method.add_methods
-   [("intro_classes", Method.no_args (Method.METHOD intro_classes_tac),
-      "back-chain introduction rules of classes"),
-    ("default", Method.thms_ctxt_args (Method.METHOD oo default_tac),
-      "apply some intro/elim rule")]));
-
-
-(** classes and class target **)
-
-(* class context syntax *)
-
-fun synchronize_class_syntax sups base_sort ctxt =
-  let
-    val thy = ProofContext.theory_of ctxt;
-    val algebra = Sign.classes_of thy;
-    val operations = these_operations thy sups;
-    fun subst_class_typ sort = map_type_tfree (K (TVar ((Name.aT, 0), sort)));
-    val primary_constraints =
-      (map o apsnd) (subst_class_typ base_sort o fst o snd) operations;
-    val secondary_constraints =
-      (map o apsnd) (fn (class, (ty, _)) => subst_class_typ [class] ty) operations;
-    fun declare_const (c, _) =
-      let val b = Sign.base_name c
-      in Sign.intern_const thy b = c ? Variable.declare_const (b, c) end;
-    fun improve (c, ty) = (case AList.lookup (op =) primary_constraints c
-     of SOME ty' => (case try (Type.raw_match (ty', ty)) Vartab.empty
-         of SOME tyenv => (case Vartab.lookup tyenv (Name.aT, 0)
-             of SOME (_, ty' as TVar (tvar as (vi, sort))) =>
-                  if TypeInfer.is_param vi
-                    andalso Sorts.sort_le algebra (base_sort, sort)
-                      then SOME (ty', TFree (Name.aT, base_sort))
-                      else NONE
-              | _ => NONE)
-          | NONE => NONE)
-      | NONE => NONE)
-    fun subst (c, ty) = Option.map snd (AList.lookup (op =) operations c);
-    val unchecks = map (fn (c, (_, (ty, t))) => (t, Const (c, ty))) operations;
-  in
-    ctxt
-    |> fold declare_const primary_constraints
-    |> Overloading.map_improvable_syntax (K (((primary_constraints, secondary_constraints),
-        (((improve, subst), true), unchecks)), false))
-    |> Overloading.set_primary_constraints
-  end;
-
-fun refresh_syntax class ctxt =
-  let
-    val thy = ProofContext.theory_of ctxt;
-    val base_sort = (#base_sort o the_class_data thy) class;
-  in synchronize_class_syntax [class] base_sort ctxt end;
-
-fun init_ctxt sups base_sort ctxt =
-  ctxt
-  |> Variable.declare_term
-      (Logic.mk_type (TFree (Name.aT, base_sort)))
-  |> synchronize_class_syntax sups base_sort
-  |> Overloading.add_improvable_syntax;
-
-fun init class thy =
-  thy
-  |> Locale.init class
-  |> init_ctxt [class] ((#base_sort o the_class_data thy) class);
-
-
-(* class target *)
-
-fun declare class pos ((c, mx), dict) thy =
-  let
-    val prfx = class_prefix class;
-    val thy' = thy |> Sign.add_path prfx;
-    val phi = morphism thy' class;
-    val inst = the_inst thy' class;
-    val params = map (apsnd fst o snd) (these_params thy' [class]);
-
-    val c' = Sign.full_bname thy' c;
-    val dict' = Morphism.term phi dict;
-    val dict_def = map_types Logic.unvarifyT dict';
-    val ty' = Term.fastype_of dict_def;
-    val ty'' = Type.strip_sorts ty';
-    val def_eq = Logic.mk_equals (Const (c', ty'), dict_def);
-    fun get_axiom thy = ((Thm.varifyT o Thm.axiom thy) c', thy);
-  in
-    thy'
-    |> Sign.declare_const pos ((Binding.name c, ty''), mx) |> snd
-    |> Thm.add_def false false (c, def_eq)
-    |>> Thm.symmetric
-    ||>> get_axiom
-    |-> (fn (def, def') => prove_class_interpretation class inst params [] [def]
-      #> snd
-        (*assumption: interpretation cookie does not change
-          by adding equations to interpretation*)
-      #> register_operation class (c', (dict', SOME (Thm.symmetric def')))
-      #> PureThy.store_thm (c ^ "_raw", def')
-      #> snd)
-    |> Sign.restore_naming thy
-    |> Sign.add_const_constraint (c', SOME ty')
-  end;
-
-fun abbrev class prmode pos ((c, mx), rhs) thy =
-  let
-    val prfx = class_prefix class;
-    val thy' = thy |> Sign.add_path prfx;
-
-    val unchecks = map (fn (c, (_, (ty, t))) => (t, Const (c, ty)))
-      (these_operations thy [class]);
-    val c' = Sign.full_bname thy' c;
-    val rhs' = Pattern.rewrite_term thy unchecks [] rhs;
-    val rhs'' = map_types Logic.varifyT rhs';
-    val ty' = Term.fastype_of rhs';
-  in
-    thy'
-    |> Sign.add_abbrev (#1 prmode) pos (Binding.name c, map_types Type.strip_sorts rhs'') |> snd
-    |> Sign.add_const_constraint (c', SOME ty')
-    |> Sign.notation true prmode [(Const (c', ty'), mx)]
-    |> not (#1 prmode = PrintMode.input) ? register_operation class (c', (rhs', NONE))
-    |> Sign.restore_naming thy
-  end;
-
-
-(* class definition *)
+(** define classes **)
 
 local
 
@@ -591,7 +104,7 @@
       else ();
     val base_sort = if null sups then supsort else
       foldr1 (Sorts.inter_sort (Sign.classes_of thy))
-        (map (#base_sort o the_class_data thy) sups);
+        (map (base_sort thy) sups);
     val suplocales = map Locale.Locale sups;
     val supexpr = Locale.Merge suplocales;
     val supparams = (map fst o Locale.parameters_of_expr thy) supexpr;
@@ -610,7 +123,7 @@
       ProofContext.init thy
       |> Locale.cert_expr supexpr [constrain]
       |> snd
-      |> init_ctxt sups base_sort
+      |> begin sups base_sort
       |> process_expr Locale.empty raw_elems
       |> fst
       |> fork_syntax
@@ -692,8 +205,8 @@
     #-> (fn phi =>
         `(fn thy => calculate_rules thy phi sups base_sort param_map axiom class)
     #-> (fn (assm_intro, of_class) =>
-        add_class_data ((class, sups), (params, base_sort, inst,
-          morphism, axiom, assm_intro, of_class))
+        register class sups params base_sort inst
+          morphism axiom assm_intro of_class
     #> fold (note_assm_intro class) (the_list assm_intro))))))
     |> init class
     |> pair class
@@ -707,196 +220,58 @@
 end; (*local*)
 
 
-
-(** instantiation target **)
-
-(* bookkeeping *)
-
-datatype instantiation = Instantiation of {
-  arities: string list * (string * sort) list * sort,
-  params: ((string * string) * (string * typ)) list
-    (*(instantiation parameter, type constructor), (local instantiation parameter, typ)*)
-}
-
-structure Instantiation = ProofDataFun
-(
-  type T = instantiation
-  fun init _ = Instantiation { arities = ([], [], []), params = [] };
-);
+(** subclass relations **)
 
-fun mk_instantiation (arities, params) =
-  Instantiation { arities = arities, params = params };
-fun get_instantiation lthy = case Instantiation.get (LocalTheory.target_of lthy)
- of Instantiation data => data;
-fun map_instantiation f = (LocalTheory.target o Instantiation.map)
-  (fn Instantiation { arities, params } => mk_instantiation (f (arities, params)));
-
-fun the_instantiation lthy = case get_instantiation lthy
- of { arities = ([], [], []), ... } => error "No instantiation target"
-  | data => data;
-
-val instantiation_params = #params o get_instantiation;
-
-fun instantiation_param lthy v = instantiation_params lthy
-  |> find_first (fn (_, (v', _)) => v = v')
-  |> Option.map (fst o fst);
-
+local
 
-(* syntax *)
-
-fun synchronize_inst_syntax ctxt =
+fun gen_subclass prep_class do_proof raw_sup lthy =
   let
-    val Instantiation { arities = (_, _, sort), params = params } = Instantiation.get ctxt;
-    val thy = ProofContext.theory_of ctxt;
-    fun subst (c, ty) = case AxClass.inst_tyco_of thy (c, ty)
-         of SOME tyco => (case AList.lookup (op =) params (c, tyco)
-             of SOME (v_ty as (_, ty)) => SOME (ty, Free v_ty)
-              | NONE => NONE)
-          | NONE => NONE;
-    val unchecks =
-      map (fn ((c, _), v_ty as (_, ty)) => (Free v_ty, Const (c, ty))) params;
+    val thy = ProofContext.theory_of lthy;
+    val sup = prep_class thy raw_sup;
+    val sub = case TheoryTarget.peek lthy
+     of {is_class = false, ...} => error "Not a class context"
+      | {target, ...} => target;
+    val _ = if Sign.subsort thy ([sup], [sub])
+      then error ("Class " ^ Syntax.string_of_sort lthy [sup]
+        ^ " is subclass of class " ^ Syntax.string_of_sort lthy [sub])
+      else ();
+    val sub_params = map fst (these_params thy [sub]);
+    val sup_params = map fst (these_params thy [sup]);
+    val err_params = subtract (op =) sub_params sup_params;
+    val _ = if null err_params then [] else
+      error ("Class " ^ Syntax.string_of_sort lthy [sub] ^ " lacks parameter(s) " ^
+        commas_quote err_params ^ " of " ^ Syntax.string_of_sort lthy [sup]);
+    val sublocale_prop =
+      Locale.global_asms_of thy sup
+      |> maps snd
+      |> try the_single
+      |> Option.map (ObjectLogic.ensure_propT thy);
+    fun after_qed some_thm =
+      LocalTheory.theory (prove_subclass_relation (sub, sup) some_thm)
+      #> (TheoryTarget.init (SOME sub) o ProofContext.theory_of);
   in
-    ctxt
-    |> Overloading.map_improvable_syntax
-         (fn (((primary_constraints, _), (((improve, _), _), _)), _) =>
-            (((primary_constraints, []), (((improve, subst), false), unchecks)), false))
-  end;
-
-
-(* target *)
-
-val sanatize_name = (*FIXME*)
-  let
-    fun is_valid s = Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s
-      orelse s = "'" orelse s = "_";
-    val is_junk = not o is_valid andf Symbol.is_regular;
-    val junk = Scan.many is_junk;
-    val scan_valids = Symbol.scanner "Malformed input"
-      ((junk |--
-        (Scan.optional (Scan.one Symbol.is_ascii_letter) "x" ^^ (Scan.many is_valid >> implode)
-        --| junk))
-      ::: Scan.repeat ((Scan.many1 is_valid >> implode) --| junk));
-  in
-    explode #> scan_valids #> implode
+    do_proof after_qed sublocale_prop lthy
   end;
 
-fun type_name "*" = "prod"
-  | type_name "+" = "sum"
-  | type_name s = sanatize_name (NameSpace.base s); (*FIXME*)
-
-fun resort_terms pp algebra consts constraints ts =
-  let
-    fun matchings (Const (c_ty as (c, _))) = (case constraints c
-         of NONE => I
-          | SOME sorts => fold2 (curry (Sorts.meet_sort algebra))
-              (Consts.typargs consts c_ty) sorts)
-      | matchings _ = I
-    val tvartab = (fold o fold_aterms) matchings ts Vartab.empty
-      handle Sorts.CLASS_ERROR e => error (Sorts.class_error pp e);
-    val inst = map_type_tvar
-      (fn (vi, sort) => TVar (vi, the_default sort (Vartab.lookup tvartab vi)));
-  in if Vartab.is_empty tvartab then NONE else SOME ((map o map_types) inst ts) end;
-
-fun init_instantiation (tycos, vs, sort) thy =
-  let
-    val _ = if null tycos then error "At least one arity must be given" else ();
-    val params = these_params thy sort;
-    fun get_param tyco (param, (_, (c, ty))) = if can (AxClass.param_of_inst thy) (c, tyco)
-      then NONE else SOME ((c, tyco),
-        (param ^ "_" ^ type_name tyco, map_atyps (K (Type (tyco, map TFree vs))) ty));
-    val inst_params = map_product get_param tycos params |> map_filter I;
-    val primary_constraints = map (apsnd
-      (map_atyps (K (TVar ((Name.aT, 0), [])))) o snd o snd) params;
-    val pp = Syntax.pp_global thy;
-    val algebra = Sign.classes_of thy
-      |> fold (fn tyco => Sorts.add_arities pp
-            (tyco, map (fn class => (class, map snd vs)) sort)) tycos;
-    val consts = Sign.consts_of thy;
-    val improve_constraints = AList.lookup (op =)
-      (map (fn (_, (class, (c, _))) => (c, [[class]])) params);
-    fun resort_check ts ctxt = case resort_terms pp algebra consts improve_constraints ts
-     of NONE => NONE
-      | SOME ts' => SOME (ts', ctxt);
-    fun improve (c, ty) = case AxClass.inst_tyco_of thy (c, ty)
-     of SOME tyco => (case AList.lookup (op =) inst_params (c, tyco)
-         of SOME (_, ty') => if Type.raw_instance (ty', ty) then SOME (ty, ty') else NONE
-          | NONE => NONE)
-      | NONE => NONE;
-  in
-    thy
-    |> ProofContext.init
-    |> Instantiation.put (mk_instantiation ((tycos, vs, sort), inst_params))
-    |> fold (Variable.declare_typ o TFree) vs
-    |> fold (Variable.declare_names o Free o snd) inst_params
-    |> (Overloading.map_improvable_syntax o apfst)
-         (fn ((_, _), ((_, subst), unchecks)) =>
-            ((primary_constraints, []), (((improve, K NONE), false), [])))
-    |> Overloading.add_improvable_syntax
-    |> Context.proof_map (Syntax.add_term_check 0 "resorting" resort_check)
-    |> synchronize_inst_syntax
-  end;
+fun user_proof after_qed NONE =
+      Proof.theorem_i NONE (K (after_qed NONE)) [[]]
+  | user_proof after_qed (SOME prop) =
+      Proof.theorem_i NONE (after_qed o try the_single o the_single) [[(prop, [])]];
 
-fun confirm_declaration c = (map_instantiation o apsnd)
-  (filter_out (fn (_, (c', _)) => c' = c))
-  #> LocalTheory.target synchronize_inst_syntax
-
-fun gen_instantiation_instance do_proof after_qed lthy =
-  let
-    val (tycos, vs, sort) = (#arities o the_instantiation) lthy;
-    val arities_proof = maps (fn tyco => Logic.mk_arities (tyco, map snd vs, sort)) tycos;
-    fun after_qed' results =
-      LocalTheory.theory (fold (AxClass.add_arity o Thm.varifyT) results)
-      #> after_qed;
-  in
-    lthy
-    |> do_proof after_qed' arities_proof
-  end;
-
-val instantiation_instance = gen_instantiation_instance (fn after_qed => fn ts =>
-  Proof.theorem_i NONE (after_qed o map the_single) (map (fn t => [(t, [])]) ts));
-
-fun prove_instantiation_instance tac = gen_instantiation_instance (fn after_qed =>
-  fn ts => fn lthy => after_qed (map (fn t => Goal.prove lthy [] [] t
-    (fn {context, ...} => tac context)) ts) lthy) I;
-
-fun prove_instantiation_exit tac = prove_instantiation_instance tac
-  #> LocalTheory.exit_global;
+fun tactic_proof tac after_qed NONE lthy =
+      after_qed NONE lthy
+  | tactic_proof tac after_qed (SOME prop) lthy =
+      after_qed (SOME (Goal.prove (LocalTheory.target_of lthy) [] [] prop
+        (K tac))) lthy;
 
-fun prove_instantiation_exit_result f tac x lthy =
-  let
-    val phi = ProofContext.export_morphism lthy
-      (ProofContext.init (ProofContext.theory_of lthy));
-    val y = f phi x;
-  in
-    lthy
-    |> prove_instantiation_exit (fn ctxt => tac ctxt y)
-    |> pair y
-  end;
+in
 
-fun conclude_instantiation lthy =
-  let
-    val { arities, params } = the_instantiation lthy;
-    val (tycos, vs, sort) = arities;
-    val thy = ProofContext.theory_of lthy;
-    val _ = map (fn tyco => if Sign.of_sort thy
-        (Type (tyco, map TFree vs), sort)
-      then () else error ("Missing instance proof for type " ^ quote (Sign.extern_type thy tyco)))
-        tycos;
-  in lthy end;
+val subclass = gen_subclass (K I) user_proof;
+fun prove_subclass tac = gen_subclass (K I) (tactic_proof tac);
+val subclass_cmd = gen_subclass Sign.read_class user_proof;
 
-fun pretty_instantiation lthy =
-  let
-    val { arities, params } = the_instantiation lthy;
-    val (tycos, vs, sort) = arities;
-    val thy = ProofContext.theory_of lthy;
-    fun pr_arity tyco = Syntax.pretty_arity lthy (tyco, map snd vs, sort);
-    fun pr_param ((c, _), (v, ty)) =
-      (Pretty.block o Pretty.breaks) [Pretty.str v, Pretty.str "==",
-        (Pretty.str o Sign.extern_const thy) c, Pretty.str "::", Syntax.pretty_typ_global thy ty];
-  in
-    (Pretty.block o Pretty.fbreaks)
-      (Pretty.str "instantiation" :: map pr_arity tycos @ map pr_param params)
-  end;
+end; (*local*)
+
 
 end;
 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Isar/class_target.ML	Mon Jan 05 15:36:24 2009 +0100
@@ -0,0 +1,723 @@
+(*  Title:      Pure/Isar/class_target.ML
+    Author:     Florian Haftmann, TU Muenchen
+
+Type classes derived from primitive axclasses and locales - mechanisms.
+*)
+
+signature CLASS_TARGET =
+sig
+  (*classes*)
+  type raw_morphism = morphism * ((typ Vartab.table * typ list) *  (term Vartab.table * term list));
+  val register: class -> class list -> ((string * typ) * (string * typ)) list
+    -> sort -> term list -> raw_morphism
+    -> thm option -> thm option -> thm -> theory -> theory
+
+  val begin: class list -> sort -> Proof.context -> Proof.context
+  val init: class -> theory -> Proof.context
+  val declare: class -> Properties.T
+    -> (string * mixfix) * term -> theory -> theory
+  val abbrev: class -> Syntax.mode -> Properties.T
+    -> (string * mixfix) * term -> theory -> theory
+  val refresh_syntax: class -> Proof.context -> Proof.context
+
+  val intro_classes_tac: thm list -> tactic
+  val default_intro_tac: Proof.context -> thm list -> tactic
+  val activate_class_morphism: theory -> class -> term list
+    -> raw_morphism -> morphism
+  val prove_class_interpretation: class -> term list -> (class * string) list
+    -> thm list -> thm list -> theory -> raw_morphism * theory
+  val prove_subclass_relation: class * class -> thm option -> theory -> theory
+
+  val class_prefix: string -> string
+  val is_class: theory -> class -> bool
+  val these_params: theory -> sort -> (string * (class * (string * typ))) list
+  val these_defs: theory -> sort -> thm list
+  val base_sort: theory -> class -> sort
+  val rules: theory -> class -> thm option * thm
+  val print_classes: theory -> unit
+
+  (*instances*)
+  val init_instantiation: string list * (string * sort) list * sort
+    -> theory -> local_theory
+  val instantiation_instance: (local_theory -> local_theory)
+    -> local_theory -> Proof.state
+  val prove_instantiation_instance: (Proof.context -> tactic)
+    -> local_theory -> local_theory
+  val prove_instantiation_exit: (Proof.context -> tactic)
+    -> local_theory -> theory
+  val prove_instantiation_exit_result: (morphism -> 'a -> 'b)
+    -> (Proof.context -> 'b -> tactic) -> 'a -> local_theory -> 'b * theory
+  val conclude_instantiation: local_theory -> local_theory
+  val instantiation_param: local_theory -> string -> string option
+  val confirm_declaration: string -> local_theory -> local_theory
+  val pretty_instantiation: local_theory -> Pretty.T
+  val type_name: string -> string
+
+  (*old axclass layer*)
+  val axclass_cmd: bstring * xstring list
+    -> (Attrib.binding * string list) list
+    -> theory -> class * theory
+  val classrel_cmd: xstring * xstring -> theory -> Proof.state
+
+  (*old instance layer*)
+  val instance_arity: (theory -> theory) -> arity -> theory -> Proof.state
+  val instance_arity_cmd: bstring * xstring list * xstring -> theory -> Proof.state
+end;
+
+structure Class_Target : CLASS_TARGET =
+struct
+
+(*temporary adaption code to mediate between old and new locale code*)
+
+structure Old_Locale =
+struct
+
+val intro_locales_tac = Locale.intro_locales_tac; (*already forked!*)
+
+val interpretation = Locale.interpretation;
+val interpretation_in_locale = Locale.interpretation_in_locale;
+val get_interpret_morph = Locale.get_interpret_morph;
+val Locale = Locale.Locale;
+val extern = Locale.extern;
+val intros = Locale.intros;
+val dests = Locale.dests;
+val init = Locale.init;
+val Merge = Locale.Merge;
+val parameters_of_expr = Locale.parameters_of_expr;
+val empty = Locale.empty;
+val cert_expr = Locale.cert_expr;
+val read_expr = Locale.read_expr;
+val parameters_of = Locale.parameters_of;
+val add_locale = Locale.add_locale;
+
+end;
+
+(*structure New_Locale =
+struct
+
+val intro_locales_tac = Locale.intro_locales_tac; (*already forked!*)
+
+val interpretation = Locale.interpretation; (*!*)
+val interpretation_in_locale = Locale.interpretation_in_locale; (*!*)
+val get_interpret_morph = Locale.get_interpret_morph; (*!*)
+fun Locale loc = ([(loc, ("", Expression.Positional []))], []);
+val extern = NewLocale.extern;
+val intros = Locale.intros; (*!*)
+val dests = Locale.dests; (*!*)
+val init = NewLocale.init;
+fun Merge locs = (map (fn loc => (loc, ("", Expression.Positional []))) locs, []);
+val parameters_of_expr = Locale.parameters_of_expr; (*!*)
+val empty = ([], []);
+val cert_expr = Locale.cert_expr; (*!"*)
+val read_expr = Locale.read_expr; (*!"*)
+val parameters_of = NewLocale.params_of; (*why typ option?*)
+val add_locale = Expression.add_locale;
+
+end;*)
+
+structure Locale = Old_Locale;
+
+(*proper code again*)
+
+
+(** auxiliary **)
+
+fun prove_interpretation tac prfx_atts expr inst =
+  Locale.interpretation I prfx_atts expr inst
+  ##> Proof.global_terminal_proof
+      (Method.Basic (fn ctxt => Method.SIMPLE_METHOD (tac ctxt), Position.none), NONE)
+  ##> ProofContext.theory_of;
+
+fun prove_interpretation_in tac after_qed (name, expr) =
+  Locale.interpretation_in_locale
+      (ProofContext.theory after_qed) (name, expr)
+  #> Proof.global_terminal_proof
+      (Method.Basic (K (Method.SIMPLE_METHOD tac), Position.none), NONE)
+  #> ProofContext.theory_of;
+
+val class_prefix = Logic.const_of_class o Sign.base_name;
+
+fun class_name_morphism class =
+  Binding.map_prefix (K (Binding.add_prefix false (class_prefix class)));
+
+type raw_morphism = morphism * ((typ Vartab.table * typ list) *  (term Vartab.table * term list));
+
+fun activate_class_morphism thy class inst morphism =
+  Locale.get_interpret_morph thy (class_name_morphism class) (class, "") morphism class inst;
+
+
+(** primitive axclass and instance commands **)
+
+fun axclass_cmd (class, raw_superclasses) raw_specs thy =
+  let
+    val ctxt = ProofContext.init thy;
+    val superclasses = map (Sign.read_class thy) raw_superclasses;
+    val name_atts = map ((apsnd o map) (Attrib.attribute thy) o fst)
+      raw_specs;
+    val axiomss = ProofContext.read_propp (ctxt, map (map (rpair []) o snd)
+          raw_specs)
+      |> snd
+      |> (map o map) fst;
+  in
+    AxClass.define_class (class, superclasses) []
+      (name_atts ~~ axiomss) thy
+  end;
+
+local
+
+fun gen_instance mk_prop add_thm after_qed insts thy =
+  let
+    fun after_qed' results =
+      ProofContext.theory ((fold o fold) add_thm results #> after_qed);
+  in
+    thy
+    |> ProofContext.init
+    |> Proof.theorem_i NONE after_qed' ((map (fn t => [(t, [])])
+        o mk_prop thy) insts)
+  end;
+
+in
+
+val instance_arity =
+  gen_instance (Logic.mk_arities oo Sign.cert_arity) AxClass.add_arity;
+val instance_arity_cmd =
+  gen_instance (Logic.mk_arities oo Sign.read_arity) AxClass.add_arity I;
+val classrel =
+  gen_instance (single oo (Logic.mk_classrel oo AxClass.cert_classrel)) AxClass.add_classrel I;
+val classrel_cmd =
+  gen_instance (single oo (Logic.mk_classrel oo AxClass.read_classrel)) AxClass.add_classrel I;
+
+end; (*local*)
+
+
+(** class data **)
+
+datatype class_data = ClassData of {
+
+  (* static part *)
+  consts: (string * string) list
+    (*locale parameter ~> constant name*),
+  base_sort: sort,
+  inst: term list
+    (*canonical interpretation*),
+  morphism: Morphism.morphism * ((typ Vartab.table * typ list) * (term Vartab.table * term list))
+    (*morphism cookie of canonical interpretation*),
+  assm_intro: thm option,
+  of_class: thm,
+  axiom: thm option,
+  
+  (* dynamic part *)
+  defs: thm list,
+  operations: (string * (class * (typ * term))) list
+
+};
+
+fun rep_class_data (ClassData d) = d;
+fun mk_class_data ((consts, base_sort, inst, morphism, assm_intro, of_class, axiom),
+    (defs, operations)) =
+  ClassData { consts = consts, base_sort = base_sort, inst = inst,
+    morphism = morphism, assm_intro = assm_intro, of_class = of_class, axiom = axiom,
+    defs = defs, operations = operations };
+fun map_class_data f (ClassData { consts, base_sort, inst, morphism, assm_intro,
+    of_class, axiom, defs, operations }) =
+  mk_class_data (f ((consts, base_sort, inst, morphism, assm_intro, of_class, axiom),
+    (defs, operations)));
+fun merge_class_data _ (ClassData { consts = consts,
+    base_sort = base_sort, inst = inst, morphism = morphism, assm_intro = assm_intro,
+    of_class = of_class, axiom = axiom, defs = defs1, operations = operations1 },
+  ClassData { consts = _, base_sort = _, inst = _, morphism = _, assm_intro = _,
+    of_class = _, axiom = _, defs = defs2, operations = operations2 }) =
+  mk_class_data ((consts, base_sort, inst, morphism, assm_intro, of_class, axiom),
+    (Thm.merge_thms (defs1, defs2),
+      AList.merge (op =) (K true) (operations1, operations2)));
+
+structure ClassData = TheoryDataFun
+(
+  type T = class_data Graph.T
+  val empty = Graph.empty;
+  val copy = I;
+  val extend = I;
+  fun merge _ = Graph.join merge_class_data;
+);
+
+
+(* queries *)
+
+val lookup_class_data = Option.map rep_class_data oo try o Graph.get_node o ClassData.get;
+
+fun the_class_data thy class = case lookup_class_data thy class
+ of NONE => error ("Undeclared class " ^ quote class)
+  | SOME data => data;
+
+val is_class = is_some oo lookup_class_data;
+
+val ancestry = Graph.all_succs o ClassData.get;
+
+fun the_inst thy = #inst o the_class_data thy;
+
+fun these_params thy =
+  let
+    fun params class =
+      let
+        val const_typs = (#params o AxClass.get_info thy) class;
+        val const_names = (#consts o the_class_data thy) class;
+      in
+        (map o apsnd)
+          (fn c => (class, (c, (the o AList.lookup (op =) const_typs) c))) const_names
+      end;
+  in maps params o ancestry thy end;
+
+fun these_assm_intros thy =
+  Graph.fold (fn (_, (data, _)) => fold (insert Thm.eq_thm)
+    ((the_list o #assm_intro o rep_class_data) data)) (ClassData.get thy) [];
+
+fun these_defs thy = maps (these o Option.map #defs o lookup_class_data thy) o ancestry thy;
+
+fun these_operations thy =
+  maps (#operations o the_class_data thy) o ancestry thy;
+
+val base_sort = #base_sort oo the_class_data;
+
+fun rules thy class = let
+    val { axiom, of_class, ... } = the_class_data thy class
+  in (axiom, of_class) end;
+
+fun morphism thy class =
+  let
+    val { inst, morphism, ... } = the_class_data thy class;
+  in activate_class_morphism thy class inst morphism end;
+
+fun print_classes thy =
+  let
+    val ctxt = ProofContext.init thy;
+    val algebra = Sign.classes_of thy;
+    val arities =
+      Symtab.empty
+      |> Symtab.fold (fn (tyco, arities) => fold (fn (class, _) =>
+           Symtab.map_default (class, []) (insert (op =) tyco)) arities)
+             ((#arities o Sorts.rep_algebra) algebra);
+    val the_arities = these o Symtab.lookup arities;
+    fun mk_arity class tyco =
+      let
+        val Ss = Sorts.mg_domain algebra tyco [class];
+      in Syntax.pretty_arity ctxt (tyco, Ss, [class]) end;
+    fun mk_param (c, ty) = Pretty.str (Sign.extern_const thy c ^ " :: "
+      ^ setmp show_sorts false (Syntax.string_of_typ ctxt o Type.strip_sorts) ty);
+    fun mk_entry class = (Pretty.block o Pretty.fbreaks o map_filter I) [
+      (SOME o Pretty.str) ("class " ^ Sign.extern_class thy class ^ ":"),
+      (SOME o Pretty.block) [Pretty.str "supersort: ",
+        (Syntax.pretty_sort ctxt o Sign.minimize_sort thy o Sign.super_classes thy) class],
+      if is_class thy class then (SOME o Pretty.str)
+        ("locale: " ^ Locale.extern thy class) else NONE,
+      ((fn [] => NONE | ps => (SOME o Pretty.block o Pretty.fbreaks)
+          (Pretty.str "parameters:" :: ps)) o map mk_param
+        o these o Option.map #params o try (AxClass.get_info thy)) class,
+      (SOME o Pretty.block o Pretty.breaks) [
+        Pretty.str "instances:",
+        Pretty.list "" "" (map (mk_arity class) (the_arities class))
+      ]
+    ]
+  in
+    (Pretty.writeln o Pretty.chunks o separate (Pretty.str "")
+      o map mk_entry o Sorts.all_classes) algebra
+  end;
+
+
+(* updaters *)
+
+fun register class superclasses params base_sort inst phi
+    axiom assm_intro of_class thy =
+  let
+    val operations = map (fn (v_ty as (_, ty), (c, _)) =>
+      (c, (class, (ty, Free v_ty)))) params;
+    val add_class = Graph.new_node (class,
+        mk_class_data (((map o pairself) fst params, base_sort,
+          inst, phi, assm_intro, of_class, axiom), ([], operations)))
+      #> fold (curry Graph.add_edge class) superclasses;
+  in ClassData.map add_class thy end;
+
+fun register_operation class (c, (t, some_def)) thy =
+  let
+    val base_sort = base_sort thy class;
+    val prep_typ = map_type_tvar
+      (fn (vi as (v, _), sort) => if Name.aT = v
+        then TFree (v, base_sort) else TVar (vi, sort));
+    val t' = map_types prep_typ t;
+    val ty' = Term.fastype_of t';
+  in
+    thy
+    |> (ClassData.map o Graph.map_node class o map_class_data o apsnd)
+      (fn (defs, operations) =>
+        (fold cons (the_list some_def) defs,
+          (c, (class, (ty', t'))) :: operations))
+  end;
+
+
+(** tactics and methods **)
+
+fun prove_class_interpretation class inst consts hyp_facts def_facts thy =
+  let
+    val constraints = map (fn (class, c) => map_atyps (K (TFree (Name.aT,
+      [class]))) (Sign.the_const_type thy c)) consts;
+    val no_constraints = map (map_atyps (K (TFree (Name.aT, [])))) constraints;
+    fun add_constraint c T = Sign.add_const_constraint (c, SOME T);
+    fun tac ctxt = ALLGOALS (ProofContext.fact_tac (hyp_facts @ def_facts)
+      ORELSE' (fn n => SELECT_GOAL (Locale.intro_locales_tac false ctxt []) n));
+    val prfx = class_prefix class;
+  in
+    thy
+    |> fold2 add_constraint (map snd consts) no_constraints
+    |> prove_interpretation tac (class_name_morphism class) (Locale.Locale class)
+          (map SOME inst, map (pair (Attrib.empty_binding) o Thm.prop_of) def_facts)
+    ||> fold2 add_constraint (map snd consts) constraints
+  end;
+
+fun prove_subclass_relation (sub, sup) some_thm thy =
+  let
+    val of_class = (snd o rules thy) sup;
+    val intro = case some_thm
+     of SOME thm => Drule.standard' (of_class OF [Drule.standard' thm])
+      | NONE => Thm.instantiate ([pairself (Thm.ctyp_of thy o TVar o pair (Name.aT, 0))
+          ([], [sub])], []) of_class;
+    val classrel = (intro OF (the_list o fst o rules thy) sub)
+      |> Thm.close_derivation;
+  in
+    thy
+    |> AxClass.add_classrel classrel
+    |> prove_interpretation_in (ALLGOALS (ProofContext.fact_tac (the_list some_thm)))
+         I (sub, Locale.Locale sup)
+    |> ClassData.map (Graph.add_edge (sub, sup))
+  end;
+
+fun intro_classes_tac facts st =
+  let
+    val thy = Thm.theory_of_thm st;
+    val classes = Sign.all_classes thy;
+    val class_trivs = map (Thm.class_triv thy) classes;
+    val class_intros = map_filter (try (#intro o AxClass.get_info thy)) classes;
+    val assm_intros = these_assm_intros thy;
+  in
+    Method.intros_tac (class_trivs @ class_intros @ assm_intros) facts st
+  end;
+
+fun default_intro_tac ctxt [] =
+      intro_classes_tac [] ORELSE NewLocale.intro_locales_tac true ctxt [] ORELSE
+      Locale.intro_locales_tac true ctxt []
+  | default_intro_tac _ _ = no_tac;
+
+fun default_tac rules ctxt facts =
+  HEADGOAL (Method.some_rule_tac rules ctxt facts) ORELSE
+    default_intro_tac ctxt facts;
+
+val _ = Context.>> (Context.map_theory
+  (Method.add_methods
+   [("intro_classes", Method.no_args (Method.METHOD intro_classes_tac),
+      "back-chain introduction rules of classes"),
+    ("default", Method.thms_ctxt_args (Method.METHOD oo default_tac),
+      "apply some intro/elim rule")]));
+
+
+(** classes and class target **)
+
+(* class context syntax *)
+
+fun synchronize_class_syntax sups base_sort ctxt =
+  let
+    val thy = ProofContext.theory_of ctxt;
+    val algebra = Sign.classes_of thy;
+    val operations = these_operations thy sups;
+    fun subst_class_typ sort = map_type_tfree (K (TVar ((Name.aT, 0), sort)));
+    val primary_constraints =
+      (map o apsnd) (subst_class_typ base_sort o fst o snd) operations;
+    val secondary_constraints =
+      (map o apsnd) (fn (class, (ty, _)) => subst_class_typ [class] ty) operations;
+    fun declare_const (c, _) =
+      let val b = Sign.base_name c
+      in Sign.intern_const thy b = c ? Variable.declare_const (b, c) end;
+    fun improve (c, ty) = (case AList.lookup (op =) primary_constraints c
+     of SOME ty' => (case try (Type.raw_match (ty', ty)) Vartab.empty
+         of SOME tyenv => (case Vartab.lookup tyenv (Name.aT, 0)
+             of SOME (_, ty' as TVar (tvar as (vi, sort))) =>
+                  if TypeInfer.is_param vi
+                    andalso Sorts.sort_le algebra (base_sort, sort)
+                      then SOME (ty', TFree (Name.aT, base_sort))
+                      else NONE
+              | _ => NONE)
+          | NONE => NONE)
+      | NONE => NONE)
+    fun subst (c, ty) = Option.map snd (AList.lookup (op =) operations c);
+    val unchecks = map (fn (c, (_, (ty, t))) => (t, Const (c, ty))) operations;
+  in
+    ctxt
+    |> fold declare_const primary_constraints
+    |> Overloading.map_improvable_syntax (K (((primary_constraints, secondary_constraints),
+        (((improve, subst), true), unchecks)), false))
+    |> Overloading.set_primary_constraints
+  end;
+
+fun refresh_syntax class ctxt =
+  let
+    val thy = ProofContext.theory_of ctxt;
+    val base_sort = base_sort thy class;
+  in synchronize_class_syntax [class] base_sort ctxt end;
+
+fun begin sups base_sort ctxt =
+  ctxt
+  |> Variable.declare_term
+      (Logic.mk_type (TFree (Name.aT, base_sort)))
+  |> synchronize_class_syntax sups base_sort
+  |> Overloading.add_improvable_syntax;
+
+fun init class thy =
+  thy
+  |> Locale.init class
+  |> begin [class] (base_sort thy class);
+
+
+(* class target *)
+
+fun declare class pos ((c, mx), dict) thy =
+  let
+    val prfx = class_prefix class;
+    val thy' = thy |> Sign.add_path prfx;
+    val phi = morphism thy' class;
+    val inst = the_inst thy' class;
+    val params = map (apsnd fst o snd) (these_params thy' [class]);
+
+    val c' = Sign.full_bname thy' c;
+    val dict' = Morphism.term phi dict;
+    val dict_def = map_types Logic.unvarifyT dict';
+    val ty' = Term.fastype_of dict_def;
+    val ty'' = Type.strip_sorts ty';
+    val def_eq = Logic.mk_equals (Const (c', ty'), dict_def);
+    fun get_axiom thy = ((Thm.varifyT o Thm.axiom thy) c', thy);
+  in
+    thy'
+    |> Sign.declare_const pos ((Binding.name c, ty''), mx) |> snd
+    |> Thm.add_def false false (c, def_eq)
+    |>> Thm.symmetric
+    ||>> get_axiom
+    |-> (fn (def, def') => prove_class_interpretation class inst params [] [def]
+      #> snd
+        (*assumption: interpretation cookie does not change
+          by adding equations to interpretation*)
+      #> register_operation class (c', (dict', SOME (Thm.symmetric def')))
+      #> PureThy.store_thm (c ^ "_raw", def')
+      #> snd)
+    |> Sign.restore_naming thy
+    |> Sign.add_const_constraint (c', SOME ty')
+  end;
+
+fun abbrev class prmode pos ((c, mx), rhs) thy =
+  let
+    val prfx = class_prefix class;
+    val thy' = thy |> Sign.add_path prfx;
+
+    val unchecks = map (fn (c, (_, (ty, t))) => (t, Const (c, ty)))
+      (these_operations thy [class]);
+    val c' = Sign.full_bname thy' c;
+    val rhs' = Pattern.rewrite_term thy unchecks [] rhs;
+    val rhs'' = map_types Logic.varifyT rhs';
+    val ty' = Term.fastype_of rhs';
+  in
+    thy'
+    |> Sign.add_abbrev (#1 prmode) pos (Binding.name c, map_types Type.strip_sorts rhs'') |> snd
+    |> Sign.add_const_constraint (c', SOME ty')
+    |> Sign.notation true prmode [(Const (c', ty'), mx)]
+    |> not (#1 prmode = PrintMode.input) ? register_operation class (c', (rhs', NONE))
+    |> Sign.restore_naming thy
+  end;
+
+
+(** instantiation target **)
+
+(* bookkeeping *)
+
+datatype instantiation = Instantiation of {
+  arities: string list * (string * sort) list * sort,
+  params: ((string * string) * (string * typ)) list
+    (*(instantiation parameter, type constructor), (local instantiation parameter, typ)*)
+}
+
+structure Instantiation = ProofDataFun
+(
+  type T = instantiation
+  fun init _ = Instantiation { arities = ([], [], []), params = [] };
+);
+
+fun mk_instantiation (arities, params) =
+  Instantiation { arities = arities, params = params };
+fun get_instantiation lthy = case Instantiation.get (LocalTheory.target_of lthy)
+ of Instantiation data => data;
+fun map_instantiation f = (LocalTheory.target o Instantiation.map)
+  (fn Instantiation { arities, params } => mk_instantiation (f (arities, params)));
+
+fun the_instantiation lthy = case get_instantiation lthy
+ of { arities = ([], [], []), ... } => error "No instantiation target"
+  | data => data;
+
+val instantiation_params = #params o get_instantiation;
+
+fun instantiation_param lthy v = instantiation_params lthy
+  |> find_first (fn (_, (v', _)) => v = v')
+  |> Option.map (fst o fst);
+
+
+(* syntax *)
+
+fun synchronize_inst_syntax ctxt =
+  let
+    val Instantiation { arities = (_, _, sort), params = params } = Instantiation.get ctxt;
+    val thy = ProofContext.theory_of ctxt;
+    fun subst (c, ty) = case AxClass.inst_tyco_of thy (c, ty)
+         of SOME tyco => (case AList.lookup (op =) params (c, tyco)
+             of SOME (v_ty as (_, ty)) => SOME (ty, Free v_ty)
+              | NONE => NONE)
+          | NONE => NONE;
+    val unchecks =
+      map (fn ((c, _), v_ty as (_, ty)) => (Free v_ty, Const (c, ty))) params;
+  in
+    ctxt
+    |> Overloading.map_improvable_syntax
+         (fn (((primary_constraints, _), (((improve, _), _), _)), _) =>
+            (((primary_constraints, []), (((improve, subst), false), unchecks)), false))
+  end;
+
+
+(* target *)
+
+val sanatize_name = (*FIXME*)
+  let
+    fun is_valid s = Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s
+      orelse s = "'" orelse s = "_";
+    val is_junk = not o is_valid andf Symbol.is_regular;
+    val junk = Scan.many is_junk;
+    val scan_valids = Symbol.scanner "Malformed input"
+      ((junk |--
+        (Scan.optional (Scan.one Symbol.is_ascii_letter) "x" ^^ (Scan.many is_valid >> implode)
+        --| junk))
+      ::: Scan.repeat ((Scan.many1 is_valid >> implode) --| junk));
+  in
+    explode #> scan_valids #> implode
+  end;
+
+fun type_name "*" = "prod"
+  | type_name "+" = "sum"
+  | type_name s = sanatize_name (NameSpace.base s); (*FIXME*)
+
+fun resort_terms pp algebra consts constraints ts =
+  let
+    fun matchings (Const (c_ty as (c, _))) = (case constraints c
+         of NONE => I
+          | SOME sorts => fold2 (curry (Sorts.meet_sort algebra))
+              (Consts.typargs consts c_ty) sorts)
+      | matchings _ = I
+    val tvartab = (fold o fold_aterms) matchings ts Vartab.empty
+      handle Sorts.CLASS_ERROR e => error (Sorts.class_error pp e);
+    val inst = map_type_tvar
+      (fn (vi, sort) => TVar (vi, the_default sort (Vartab.lookup tvartab vi)));
+  in if Vartab.is_empty tvartab then NONE else SOME ((map o map_types) inst ts) end;
+
+fun init_instantiation (tycos, vs, sort) thy =
+  let
+    val _ = if null tycos then error "At least one arity must be given" else ();
+    val params = these_params thy sort;
+    fun get_param tyco (param, (_, (c, ty))) = if can (AxClass.param_of_inst thy) (c, tyco)
+      then NONE else SOME ((c, tyco),
+        (param ^ "_" ^ type_name tyco, map_atyps (K (Type (tyco, map TFree vs))) ty));
+    val inst_params = map_product get_param tycos params |> map_filter I;
+    val primary_constraints = map (apsnd
+      (map_atyps (K (TVar ((Name.aT, 0), [])))) o snd o snd) params;
+    val pp = Syntax.pp_global thy;
+    val algebra = Sign.classes_of thy
+      |> fold (fn tyco => Sorts.add_arities pp
+            (tyco, map (fn class => (class, map snd vs)) sort)) tycos;
+    val consts = Sign.consts_of thy;
+    val improve_constraints = AList.lookup (op =)
+      (map (fn (_, (class, (c, _))) => (c, [[class]])) params);
+    fun resort_check ts ctxt = case resort_terms pp algebra consts improve_constraints ts
+     of NONE => NONE
+      | SOME ts' => SOME (ts', ctxt);
+    fun improve (c, ty) = case AxClass.inst_tyco_of thy (c, ty)
+     of SOME tyco => (case AList.lookup (op =) inst_params (c, tyco)
+         of SOME (_, ty') => if Type.raw_instance (ty', ty) then SOME (ty, ty') else NONE
+          | NONE => NONE)
+      | NONE => NONE;
+  in
+    thy
+    |> ProofContext.init
+    |> Instantiation.put (mk_instantiation ((tycos, vs, sort), inst_params))
+    |> fold (Variable.declare_typ o TFree) vs
+    |> fold (Variable.declare_names o Free o snd) inst_params
+    |> (Overloading.map_improvable_syntax o apfst)
+         (fn ((_, _), ((_, subst), unchecks)) =>
+            ((primary_constraints, []), (((improve, K NONE), false), [])))
+    |> Overloading.add_improvable_syntax
+    |> Context.proof_map (Syntax.add_term_check 0 "resorting" resort_check)
+    |> synchronize_inst_syntax
+  end;
+
+fun confirm_declaration c = (map_instantiation o apsnd)
+  (filter_out (fn (_, (c', _)) => c' = c))
+  #> LocalTheory.target synchronize_inst_syntax
+
+fun gen_instantiation_instance do_proof after_qed lthy =
+  let
+    val (tycos, vs, sort) = (#arities o the_instantiation) lthy;
+    val arities_proof = maps (fn tyco => Logic.mk_arities (tyco, map snd vs, sort)) tycos;
+    fun after_qed' results =
+      LocalTheory.theory (fold (AxClass.add_arity o Thm.varifyT) results)
+      #> after_qed;
+  in
+    lthy
+    |> do_proof after_qed' arities_proof
+  end;
+
+val instantiation_instance = gen_instantiation_instance (fn after_qed => fn ts =>
+  Proof.theorem_i NONE (after_qed o map the_single) (map (fn t => [(t, [])]) ts));
+
+fun prove_instantiation_instance tac = gen_instantiation_instance (fn after_qed =>
+  fn ts => fn lthy => after_qed (map (fn t => Goal.prove lthy [] [] t
+    (fn {context, ...} => tac context)) ts) lthy) I;
+
+fun prove_instantiation_exit tac = prove_instantiation_instance tac
+  #> LocalTheory.exit_global;
+
+fun prove_instantiation_exit_result f tac x lthy =
+  let
+    val phi = ProofContext.export_morphism lthy
+      (ProofContext.init (ProofContext.theory_of lthy));
+    val y = f phi x;
+  in
+    lthy
+    |> prove_instantiation_exit (fn ctxt => tac ctxt y)
+    |> pair y
+  end;
+
+fun conclude_instantiation lthy =
+  let
+    val { arities, params } = the_instantiation lthy;
+    val (tycos, vs, sort) = arities;
+    val thy = ProofContext.theory_of lthy;
+    val _ = map (fn tyco => if Sign.of_sort thy
+        (Type (tyco, map TFree vs), sort)
+      then () else error ("Missing instance proof for type " ^ quote (Sign.extern_type thy tyco)))
+        tycos;
+  in lthy end;
+
+fun pretty_instantiation lthy =
+  let
+    val { arities, params } = the_instantiation lthy;
+    val (tycos, vs, sort) = arities;
+    val thy = ProofContext.theory_of lthy;
+    fun pr_arity tyco = Syntax.pretty_arity lthy (tyco, map snd vs, sort);
+    fun pr_param ((c, _), (v, ty)) =
+      (Pretty.block o Pretty.breaks) [Pretty.str v, Pretty.str "==",
+        (Pretty.str o Sign.extern_const thy) c, Pretty.str "::", Syntax.pretty_typ_global thy ty];
+  in
+    (Pretty.block o Pretty.fbreaks)
+      (Pretty.str "instantiation" :: map pr_arity tycos @ map pr_param params)
+  end;
+
+end;
+
--- a/src/Pure/Isar/expression.ML	Mon Jan 05 15:35:42 2009 +0100
+++ b/src/Pure/Isar/expression.ML	Mon Jan 05 15:36:24 2009 +0100
@@ -18,18 +18,18 @@
     Proof.context -> (term * term list) list list * Proof.context;
 
   (* Declaring locales *)
-  val add_locale_cmd: string -> bstring -> expression -> Element.context list -> theory ->
-    (string * (string * (Attrib.binding * (thm list * Attrib.src list) list) list) list) *
-    Proof.context
-  val add_locale: string -> bstring -> expression_i -> Element.context_i list -> theory ->
-    (string * (string * (Attrib.binding * (thm list * Attrib.src list) list) list) list) *
-    Proof.context
+  val add_locale_cmd: string -> bstring -> expression -> Element.context list ->
+    theory -> string * local_theory;
+  val add_locale: string -> bstring -> expression_i -> Element.context_i list ->
+    theory -> string * local_theory;
 
   (* Interpretation *)
   val sublocale_cmd: string -> expression -> theory -> Proof.state;
   val sublocale: string -> expression_i -> theory -> Proof.state;
-  val interpretation_cmd: expression -> (Attrib.binding * string) list -> theory -> Proof.state;
-  val interpretation: expression_i -> (Attrib.binding * term) list -> theory -> Proof.state;
+  val interpretation_cmd: expression -> (Attrib.binding * string) list ->
+    theory -> Proof.state;
+  val interpretation: expression_i -> (Attrib.binding * term) list ->
+    theory -> Proof.state;
   val interpret_cmd: expression -> bool -> Proof.state -> Proof.state;
   val interpret: expression_i -> bool -> Proof.state -> Proof.state;
 end;
@@ -89,11 +89,11 @@
     fun params_inst (expr as (loc, (prfx, Positional insts))) =
           let
             val (ps, loc') = params_loc loc;
-	    val d = length ps - length insts;
-	    val insts' =
-	      if d < 0 then error ("More arguments than parameters in instantiation of locale " ^
+            val d = length ps - length insts;
+            val insts' =
+              if d < 0 then error ("More arguments than parameters in instantiation of locale " ^
                 quote (NewLocale.extern thy loc))
-	      else insts @ replicate d NONE;
+              else insts @ replicate d NONE;
             val ps' = (ps ~~ insts') |>
               map_filter (fn (p, NONE) => SOME p | (_, SOME _) => NONE);
           in (ps', (loc', (prfx, Positional insts'))) end
@@ -333,9 +333,9 @@
 local
 
 fun prep_full_context_statement parse_typ parse_prop parse_inst prep_vars prep_expr
-    strict do_close context raw_import raw_elems raw_concl =
+    strict do_close raw_import raw_elems raw_concl ctxt1 =
   let
-    val thy = ProofContext.theory_of context;
+    val thy = ProofContext.theory_of ctxt1;
 
     val (raw_insts, fixed) = parameters_of thy strict (apfst (prep_expr thy) raw_import);
 
@@ -366,22 +366,23 @@
         val concl = parse_concl parse_prop ctxt raw_concl;
       in check_autofix insts elems concl ctxt end;
 
-    val fors = prep_vars fixed context |> fst;
-    val ctxt = context |> ProofContext.add_fixes_i fors |> snd;
-    val (_, insts', ctxt') = fold prep_inst raw_insts (0, [], ctxt);
-    val (_, elems'', ctxt'') = fold prep_elem raw_elems (insts', [], ctxt');
-    val (insts, elems, concl, ctxt) = prep_concl raw_concl (insts', elems'', ctxt'');
+    val fors = prep_vars fixed ctxt1 |> fst;
+    val ctxt2 = ctxt1 |> ProofContext.add_fixes_i fors |> snd;
+    val (_, insts', ctxt3) = fold prep_inst raw_insts (0, [], ctxt2);
+    val (_, elems'', ctxt4) = fold prep_elem raw_elems (insts', [], ctxt3);
+    val (insts, elems, concl, ctxt5) =
+      prep_concl raw_concl (insts', elems'', ctxt4);
 
     (* Retrieve parameter types *)
     val xs = fold (fn Fixes fixes => (fn ps => ps @ map (Binding.base_name o #1) fixes) |
       _ => fn ps => ps) (Fixes fors :: elems) [];
-    val (Ts, ctxt') = fold_map ProofContext.inferred_param xs ctxt; 
+    val (Ts, ctxt6) = fold_map ProofContext.inferred_param xs ctxt5; 
     val parms = xs ~~ Ts;  (* params from expression and elements *)
 
     val Fixes fors' = finish_primitive parms I (Fixes fors);
-    val (deps, elems') = finish ctxt' parms do_close insts elems;
+    val (deps, elems') = finish ctxt6 parms do_close insts elems;
 
-  in ((fors', deps, elems', concl), (parms, ctxt')) end
+  in ((fors', deps, elems', concl), (parms, ctxt6)) end
 
 in
 
@@ -400,7 +401,8 @@
 
 fun prep_statement prep activate raw_elems raw_concl context =
   let
-     val ((_, _, elems, concl), _) = prep true false context ([], []) raw_elems raw_concl;
+     val ((_, _, elems, concl), _) =
+       prep true false ([], []) raw_elems raw_concl context ;
      val (_, context') = activate elems (ProofContext.set_stmt true context);
   in (concl, context') end;
 
@@ -418,7 +420,8 @@
 
 fun prep_declaration prep activate raw_import raw_elems context =
   let
-    val ((fixed, deps, elems, _), (parms, ctxt')) = prep false true context raw_import raw_elems [];
+    val ((fixed, deps, elems, _), (parms, ctxt')) =
+      prep false true raw_import raw_elems [] context ;
     (* Declare parameters and imported facts *)
     val context' = context |>
       ProofContext.add_fixes_i fixed |> snd |>
@@ -449,7 +452,8 @@
   let
     val thy = ProofContext.theory_of context;
 
-    val ((fixed, deps, _, _), _) = prep true true context expression [] [];
+    val ((fixed, deps, _, _), _) =
+      prep true true expression [] [] context;
     (* proof obligations *)
     val propss = map (props_of thy) deps;
 
@@ -718,7 +722,6 @@
       (body_elems |> map_filter (fn Fixes fixes => SOME fixes | _ => NONE) |> flat);
     val asm = if is_some b_statement then b_statement else a_statement;
 
-    (* These are added immediately. *)
     val notes =
         if is_some asm
         then [(Thm.internalK, [((Binding.name (bname ^ "_" ^ axiomsN), []),
@@ -726,7 +729,6 @@
             [(Attrib.internal o K) NewLocale.witness_attrib])])])]
         else [];
 
-    (* These will be added in the local theory. *)
     val notes' = body_elems |>
       map (defines_to_notes thy') |>
       map (Element.morph_ctxt a_satisfy) |>
@@ -737,13 +739,14 @@
 
     val deps' = map (fn (l, morph) => (l, morph $> b_satisfy)) deps;
 
-    val loc_ctxt = thy' |>
-      NewLocale.register_locale bname (extraTs, params)
-        (asm, rev defs) ([], [])
-        (map (fn n => (n, stamp ())) notes |> rev) (map (fn d => (d, stamp ())) deps' |> rev) |>
-      NewLocale.init name;
+    val loc_ctxt = thy'
+      |> NewLocale.register_locale bname (extraTs, params)
+          (asm, rev defs) ([], [])
+          (map (fn n => (n, stamp ())) notes |> rev) (map (fn d => (d, stamp ())) deps' |> rev)
+      |> TheoryTarget.init (SOME name)
+      |> fold (fn (kind, facts) => LocalTheory.notes kind facts #> snd) notes';
 
-  in ((name, notes'), loc_ctxt) end;
+  in (name, loc_ctxt) end;
 
 in
 
--- a/src/Pure/Isar/instance.ML	Mon Jan 05 15:35:42 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,28 +0,0 @@
-(*  Title:      Pure/Isar/instance.ML
-    ID:         $Id$
-    Author:     Florian Haftmann, TU Muenchen
-
-Wrapper for instantiation command.
-*)
-
-signature INSTANCE =
-sig
-  val instantiation_cmd: xstring list * sort * xstring -> theory -> local_theory
-end;
-
-structure Instance : INSTANCE =
-struct
-
-fun read_multi_arity thy (raw_tycos, raw_sorts, raw_sort) =
-  let
-    val all_arities = map (fn raw_tyco => Sign.read_arity thy
-      (raw_tyco, raw_sorts, raw_sort)) raw_tycos;
-    val tycos = map #1 all_arities;
-    val (_, sorts, sort) = hd all_arities;
-    val vs = Name.names Name.context Name.aT sorts;
-  in (tycos, vs, sort) end;
-
-fun instantiation_cmd raw_arities thy =
-  TheoryTarget.instantiation (read_multi_arity thy raw_arities) thy;
-
-end;
--- a/src/Pure/Isar/isar_syn.ML	Mon Jan 05 15:35:42 2009 +0100
+++ b/src/Pure/Isar/isar_syn.ML	Mon Jan 05 15:36:24 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      Pure/Isar/isar_syn.ML
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 
 Isar/Pure outer syntax.
@@ -394,9 +393,7 @@
     (P.name -- Scan.optional (P.$$$ "=" |-- P.!!! locale_val) (([], []), []) -- P.opt_begin
       >> (fn ((name, (expr, elems)), begin) =>
           (begin ? Toplevel.print) o Toplevel.begin_local_theory begin
-            (Expression.add_locale_cmd name name expr elems #>
-              (fn ((target, notes), ctxt) => TheoryTarget.begin target ctxt |>
-                fold (fn (kind, facts) => LocalTheory.notes kind facts #> snd) notes))));
+            (Expression.add_locale_cmd name name expr elems #> snd)));
 
 val _ =
   OuterSyntax.command "sublocale"
@@ -477,13 +474,13 @@
 
 val _ =
   OuterSyntax.local_theory_to_proof "subclass" "prove a subclass relation" K.thy_goal
-    (P.xname >> Subclass.subclass_cmd);
+    (P.xname >> Class.subclass_cmd);
 
 val _ =
   OuterSyntax.command "instantiation" "instantiate and prove type arity" K.thy_decl
    (P.multi_arity --| P.begin
      >> (fn arities => Toplevel.print o
-         Toplevel.begin_local_theory true (Instance.instantiation_cmd arities)));
+         Toplevel.begin_local_theory true (TheoryTarget.instantiation_cmd arities)));
 
 val _ =
   OuterSyntax.command "instance" "prove type arity or subclass relation" K.thy_goal
--- a/src/Pure/Isar/subclass.ML	Mon Jan 05 15:35:42 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,67 +0,0 @@
-(*  Title:      Pure/Isar/subclass.ML
-    Author:     Florian Haftmann, TU Muenchen
-
-User interface for proving subclass relationship between type classes.
-*)
-
-signature SUBCLASS =
-sig
-  val subclass: class -> local_theory -> Proof.state
-  val subclass_cmd: xstring -> local_theory -> Proof.state
-  val prove_subclass: tactic -> class -> local_theory -> local_theory
-end;
-
-structure Subclass : SUBCLASS =
-struct
-
-local
-
-fun gen_subclass prep_class do_proof raw_sup lthy =
-  let
-    val thy = ProofContext.theory_of lthy;
-    val sup = prep_class thy raw_sup;
-    val sub = case TheoryTarget.peek lthy
-     of {is_class = false, ...} => error "Not a class context"
-      | {target, ...} => target;
-    val _ = if Sign.subsort thy ([sup], [sub])
-      then error ("Class " ^ Syntax.string_of_sort lthy [sup]
-        ^ " is subclass of class " ^ Syntax.string_of_sort lthy [sub])
-      else ();
-    val sub_params = map fst (Class.these_params thy [sub]);
-    val sup_params = map fst (Class.these_params thy [sup]);
-    val err_params = subtract (op =) sub_params sup_params;
-    val _ = if null err_params then [] else
-      error ("Class " ^ Syntax.string_of_sort lthy [sub] ^ " lacks parameter(s) " ^
-        commas_quote err_params ^ " of " ^ Syntax.string_of_sort lthy [sup]);
-    val sublocale_prop =
-      Locale.global_asms_of thy sup
-      |> maps snd
-      |> try the_single
-      |> Option.map (ObjectLogic.ensure_propT thy);
-    fun after_qed some_thm =
-      LocalTheory.theory (Class.prove_subclass (sub, sup) some_thm)
-      #> (TheoryTarget.init (SOME sub) o ProofContext.theory_of);
-  in
-    do_proof after_qed sublocale_prop lthy
-  end;
-
-fun user_proof after_qed NONE =
-      Proof.theorem_i NONE (K (after_qed NONE)) [[]]
-  | user_proof after_qed (SOME prop) =
-      Proof.theorem_i NONE (after_qed o try the_single o the_single) [[(prop, [])]];
-
-fun tactic_proof tac after_qed NONE lthy =
-      after_qed NONE lthy
-  | tactic_proof tac after_qed (SOME prop) lthy =
-      after_qed (SOME (Goal.prove (LocalTheory.target_of lthy) [] [] prop
-        (K tac))) lthy;
-
-in
-
-val subclass = gen_subclass (K I) user_proof;
-val subclass_cmd = gen_subclass Sign.read_class user_proof;
-fun prove_subclass tac = gen_subclass (K I) (tactic_proof tac);
-
-end; (*local*)
-
-end;
--- a/src/Pure/Isar/theory_target.ML	Mon Jan 05 15:35:42 2009 +0100
+++ b/src/Pure/Isar/theory_target.ML	Mon Jan 05 15:36:24 2009 +0100
@@ -13,6 +13,7 @@
   val begin: string -> Proof.context -> local_theory
   val context: xstring -> theory -> local_theory
   val instantiation: string list * (string * sort) list * sort -> theory -> local_theory
+  val instantiation_cmd: xstring list * sort * xstring -> theory -> local_theory
   val overloading: (string * (string * typ) * bool) list -> theory -> local_theory
   val overloading_cmd: (string * string * bool) list -> theory -> local_theory
 end;
@@ -82,7 +83,7 @@
   Pretty.block [Pretty.str "theory", Pretty.brk 1,
       Pretty.str (Context.theory_name (ProofContext.theory_of ctxt))] ::
     (if not (null overloading) then [Overloading.pretty ctxt]
-     else if not (null (#1 instantiation)) then [Class.pretty_instantiation ctxt]
+     else if not (null (#1 instantiation)) then [Class_Target.pretty_instantiation ctxt]
      else pretty_thy ctxt target is_locale is_class);
 
 
@@ -108,7 +109,7 @@
 
 fun class_target (Target {target, ...}) f =
   LocalTheory.raw_theory f #>
-  LocalTheory.target (Class.refresh_syntax target);
+  LocalTheory.target (Class_Target.refresh_syntax target);
 
 
 (* notes *)
@@ -207,7 +208,7 @@
     val (prefix', _) = Binding.dest b';
     val class_global = Binding.base_name b = Binding.base_name b'
       andalso not (null prefix')
-      andalso (fst o snd o split_last) prefix' = Class.class_prefix target;
+      andalso (fst o snd o split_last) prefix' = Class_Target.class_prefix target;
   in
     not (is_class andalso (similar_body orelse class_global)) ?
       (Context.mapping_result
@@ -231,11 +232,11 @@
     val (mx1, mx2, mx3) = fork_mixfix ta mx;
     fun syntax_error c = error ("Illegal mixfix syntax for overloaded constant " ^ quote c);
     val declare_const =
-      (case Class.instantiation_param lthy c of
+      (case Class_Target.instantiation_param lthy c of
         SOME c' =>
           if mx3 <> NoSyn then syntax_error c'
           else LocalTheory.theory_result (AxClass.declare_overloaded (c', U))
-            ##> Class.confirm_declaration c
+            ##> Class_Target.confirm_declaration c
         | NONE =>
             (case Overloading.operation lthy c of
               SOME (c', _) =>
@@ -248,7 +249,7 @@
   in
     lthy'
     |> is_locale ? term_syntax ta (locale_const ta Syntax.mode_default tags ((b, mx2), t))
-    |> is_class ? class_target ta (Class.declare target tags ((c, mx1), t))
+    |> is_class ? class_target ta (Class_Target.declare target tags ((c, mx1), t))
     |> LocalDefs.add_def ((b, NoSyn), t)
   end;
 
@@ -275,7 +276,7 @@
         #-> (fn (lhs, _) =>
           let val lhs' = Term.list_comb (Logic.unvarify lhs, xs) in
             term_syntax ta (locale_const ta prmode tags ((b, mx2), lhs')) #>
-            is_class ? class_target ta (Class.abbrev target prmode tags ((c, mx1), t'))
+            is_class ? class_target ta (Class_Target.abbrev target prmode tags ((c, mx1), t'))
           end)
       else
         LocalTheory.theory
@@ -311,7 +312,7 @@
         SOME (_, checked) =>
           (fn name => fn (Const (c, _), rhs) => Overloading.define checked name (c, rhs))
       | NONE =>
-          if is_none (Class.instantiation_param lthy c)
+          if is_none (Class_Target.instantiation_param lthy c)
           then (fn name => fn eq => Thm.add_def false false (name, Logic.mk_equals eq))
           else (fn name => fn (Const (c, _), rhs) => AxClass.define_overloaded name (c, rhs)));
     val (global_def, lthy3) = lthy2
@@ -334,14 +335,14 @@
 fun init_target _ NONE = global_target
   | init_target thy (SOME target) =
       make_target target (NewLocale.test_locale thy (NewLocale.intern thy target))
-      true (Class.is_class thy target) ([], [], []) [];
+      true (Class_Target.is_class thy target) ([], [], []) [];
 
 fun init_ctxt (Target {target, new_locale, is_locale, is_class, instantiation, overloading}) =
-  if not (null (#1 instantiation)) then Class.init_instantiation instantiation
+  if not (null (#1 instantiation)) then Class_Target.init_instantiation instantiation
   else if not (null overloading) then Overloading.init overloading
   else if not is_locale then ProofContext.init
   else if not is_class then locale_init new_locale target
-  else Class.init target;
+  else Class_Target.init target;
 
 fun init_lthy (ta as Target {target, instantiation, overloading, ...}) =
   Data.put ta #>
@@ -355,11 +356,20 @@
     declaration = declaration ta,
     reinit = fn lthy => init_lthy_ctxt ta (ProofContext.theory_of lthy),
     exit = LocalTheory.target_of o
-      (if not (null (#1 instantiation)) then Class.conclude_instantiation
+      (if not (null (#1 instantiation)) then Class_Target.conclude_instantiation
        else if not (null overloading) then Overloading.conclude
        else I)}
 and init_lthy_ctxt ta = init_lthy ta o init_ctxt ta;
 
+fun read_multi_arity thy (raw_tycos, raw_sorts, raw_sort) =
+  let
+    val all_arities = map (fn raw_tyco => Sign.read_arity thy
+      (raw_tyco, raw_sorts, raw_sort)) raw_tycos;
+    val tycos = map #1 all_arities;
+    val (_, sorts, sort) = hd all_arities;
+    val vs = Name.names Name.context Name.aT sorts;
+  in (tycos, vs, sort) end;
+
 fun gen_overloading prep_const raw_ops thy =
   let
     val ctxt = ProofContext.init thy;
@@ -377,6 +387,8 @@
       (NewLocale.test_locale thy (NewLocale.intern thy target)) thy target)) thy;
 
 fun instantiation arities = init_lthy_ctxt (make_target "" false false false arities []);
+fun instantiation_cmd raw_arities thy =
+  instantiation (read_multi_arity thy raw_arities) thy;
 
 val overloading = gen_overloading (fn ctxt => Syntax.check_term ctxt o Const);
 val overloading_cmd = gen_overloading Syntax.read_term;