First version of global registration command.
authorballarin
Wed, 09 Mar 2005 18:44:52 +0100
changeset 15596 8665d08085df
parent 15595 dc8a41c7cefc
child 15597 b5f5722ed703
First version of global registration command.
etc/isar-keywords-ZF.el
etc/isar-keywords.el
src/FOL/IsaMakefile
src/FOL/ex/LocaleTest.thy
src/HOL/Algebra/UnivPoly.thy
src/HOL/Algebra/poly/UnivPoly2.thy
src/Pure/Isar/args.ML
src/Pure/Isar/attrib.ML
src/Pure/Isar/delta_data.ML
src/Pure/Isar/isar_cmd.ML
src/Pure/Isar/isar_syn.ML
src/Pure/Isar/isar_thy.ML
src/Pure/Isar/locale.ML
src/Pure/Isar/proof.ML
src/Pure/Isar/proof_context.ML
src/Pure/Isar/toplevel.ML
src/Pure/ROOT.ML
src/Pure/axclass.ML
src/Pure/logic.ML
--- a/etc/isar-keywords-ZF.el	Tue Mar 08 16:02:52 2005 +0100
+++ b/etc/isar-keywords-ZF.el	Wed Mar 09 18:44:52 2005 +0100
@@ -113,6 +113,7 @@
     "print_locale"
     "print_locales"
     "print_methods"
+    "print_registrations"
     "print_rules"
     "print_simpset"
     "print_syntax"
@@ -131,6 +132,7 @@
     "realizability"
     "realizers"
     "redo"
+    "registration"
     "remove_thy"
     "rep_datatype"
     "sect"
@@ -263,6 +265,7 @@
     "print_locale"
     "print_locales"
     "print_methods"
+    "print_registrations"
     "print_rules"
     "print_simpset"
     "print_syntax"
@@ -362,6 +365,7 @@
   '("corollary"
     "instance"
     "lemma"
+    "registration"
     "theorem"))
 
 (defconst isar-keywords-qed
--- a/etc/isar-keywords.el	Tue Mar 08 16:02:52 2005 +0100
+++ b/etc/isar-keywords.el	Wed Mar 09 18:44:52 2005 +0100
@@ -116,6 +116,7 @@
     "print_locale"
     "print_locales"
     "print_methods"
+    "print_registrations"
     "print_rules"
     "print_simpset"
     "print_syntax"
@@ -136,6 +137,7 @@
     "recdef_tc"
     "record"
     "redo"
+    "registration"
     "refute"
     "refute_params"
     "remove_thy"
@@ -288,6 +290,7 @@
     "print_locale"
     "print_locales"
     "print_methods"
+    "print_registrations"
     "print_rules"
     "print_simpset"
     "print_syntax"
@@ -394,6 +397,7 @@
     "instance"
     "lemma"
     "recdef_tc"
+    "registration"
     "specification"
     "theorem"
     "typedef"))
--- a/src/FOL/IsaMakefile	Tue Mar 08 16:02:52 2005 +0100
+++ b/src/FOL/IsaMakefile	Wed Mar 09 18:44:52 2005 +0100
@@ -46,6 +46,7 @@
 
 $(LOG)/FOL-ex.gz: $(OUT)/FOL ex/First_Order_Logic.thy \
   ex/If.thy ex/IffOracle.ML ex/IffOracle.thy ex/List.ML ex/List.thy	\
+  ex/LocaleTest.thy \
   ex/LocaleInst.thy \
   ex/Nat.ML ex/Nat.thy ex/Nat2.ML ex/Nat2.thy ex/Natural_Numbers.thy	\
   ex/Prolog.ML ex/Prolog.thy ex/ROOT.ML ex/Classical.thy ex/document/root.tex\
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/FOL/ex/LocaleTest.thy	Wed Mar 09 18:44:52 2005 +0100
@@ -0,0 +1,76 @@
+(*  Title:      FOL/ex/LocaleTest.thy
+    ID:         $Id$
+    Author:     Clemens Ballarin
+    Copyright (c) 2005 by Clemens Ballarin
+
+Collection of regression tests for locales.
+*)
+
+header {* Test of Locale instantiation *}
+
+theory LocaleTest = FOL:
+
+(* registration input syntax *)
+
+locale L
+locale M = fixes a and b and c
+
+registration test [simp]: L + M a b c [x y z] .
+
+print_registrations L
+print_registrations M
+
+registration test [simp]: L .
+
+registration L .
+
+(* processing of locale expression *)
+
+ML {* reset show_hyps *}
+
+locale A = fixes a assumes asm_A: "a = a"
+
+locale (open) B = fixes b assumes asm_B [simp]: "b = b"
+
+locale C = A + B + assumes asm_C: "c = c"
+  (* TODO: independent type var in c, prohibit locale declaration *)
+
+locale D = A + B + fixes d defines def_D: "d == (a = b)"
+
+ML {* set show_sorts *}
+
+typedecl i
+arities i :: "term"
+
+ML {* set Toplevel.debug *}
+
+registration p1: C ["X::'b" "Y"] by (auto intro: A.intro C_axioms.intro)
+  (* both X and Y get type 'b since 'b is the internal type of parameter b,
+     not wanted, but put up with for now. *)
+
+ML {* set show_hyps *}
+
+print_registrations A
+
+(* thm asm_A a.asm_A p1.a.asm_A *)
+
+(*
+registration p2: D [X Y Z] sorry
+
+print_registrations D
+*)
+
+registration p3: D [X Y] by (auto intro: A.intro)
+
+print_registrations A
+print_registrations B
+print_registrations C
+print_registrations D
+
+(* not permitted
+registration p4: A ["?x10"] apply (rule A.intro) apply rule done
+
+print_registrations A
+*)
+
+end
--- a/src/HOL/Algebra/UnivPoly.thy	Tue Mar 08 16:02:52 2005 +0100
+++ b/src/HOL/Algebra/UnivPoly.thy	Wed Mar 09 18:44:52 2005 +0100
@@ -369,6 +369,29 @@
     done
 qed (simp_all add: R)
 
+(*
+Strange phenomenon in Isar:
+
+theorem (in UP_cring) UP_cring:
+  "cring P"
+proof (rule cringI)
+  show "abelian_group P" proof (rule abelian_groupI)
+  fix x y z
+  assume "x \<in> carrier P" and "y \<in> carrier P" and "z \<in> carrier P"
+  {
+  show "x \<oplus>\<^bsub>P\<^esub> y \<in> carrier P" sorry
+  next
+  show "x \<oplus>\<^bsub>P\<^esub> y \<oplus>\<^bsub>P\<^esub> z = x \<oplus>\<^bsub>P\<^esub> (y \<oplus>\<^bsub>P\<^esub> z)" sorry
+  next
+  show "x \<oplus>\<^bsub>P\<^esub> y = y \<oplus>\<^bsub>P\<^esub> x" sorry
+  next
+  show "\<zero>\<^bsub>P\<^esub> \<oplus>\<^bsub>P\<^esub> x = x" sorry
+  next
+  show "\<exists>y\<in>carrier P. y \<oplus>\<^bsub>P\<^esub> x = \<zero>\<^bsub>P\<^esub>" sorry
+  next
+  show "\<zero>\<^bsub>P\<^esub> \<in> carrier P" sorry  last goal rejected!!!
+*)
+
 theorem (in UP_cring) UP_cring:
   "cring P"
   by (auto intro!: cringI abelian_groupI comm_monoidI UP_a_assoc UP_l_zero
--- a/src/HOL/Algebra/poly/UnivPoly2.thy	Tue Mar 08 16:02:52 2005 +0100
+++ b/src/HOL/Algebra/poly/UnivPoly2.thy	Wed Mar 09 18:44:52 2005 +0100
@@ -249,7 +249,6 @@
 instance up :: (ring) ring
 proof
   fix p q r :: "'a::ring up"
-  fix n
   show "(p + q) + r = p + (q + r)"
     by (rule up_eqI) simp
   show "0 + p = p"
@@ -318,6 +317,7 @@
     by (simp add: up_inverse_def)
   show "p / q = p * inverse q"
     by (simp add: up_divide_def)
+  fix n
   show "p ^ n = nat_rec 1 (%u b. b * p) n"
     by (simp add: up_power_def)
   qed
--- a/src/Pure/Isar/args.ML	Tue Mar 08 16:02:52 2005 +0100
+++ b/src/Pure/Isar/args.ML	Wed Mar 09 18:44:52 2005 +0100
@@ -43,11 +43,11 @@
   val global_typ: theory * T list -> typ * (theory * T list)
   val global_term: theory * T list -> term * (theory * T list)
   val global_prop: theory * T list -> term * (theory * T list)
-  val local_typ_raw: Proof.context * T list -> typ * (Proof.context * T list)
-  val local_typ: Proof.context * T list -> typ * (Proof.context * T list)
-  val local_term: Proof.context * T list -> term * (Proof.context * T list)
-  val local_prop: Proof.context * T list -> term * (Proof.context * T list)
-  val bang_facts: Proof.context * T list -> thm list * (Proof.context * T list)
+  val local_typ_raw: ProofContext.context * T list -> typ * (ProofContext.context * T list)
+  val local_typ: ProofContext.context * T list -> typ * (ProofContext.context * T list)
+  val local_term: ProofContext.context * T list -> term * (ProofContext.context * T list)
+  val local_prop: ProofContext.context * T list -> term * (ProofContext.context * T list)
+  val bang_facts: ProofContext.context * T list -> thm list * (ProofContext.context * T list)
   val goal_spec: ((int -> tactic) -> tactic) -> ('a * T list)
     -> ((int -> tactic) -> tactic) * ('a * T list)
   type src
--- a/src/Pure/Isar/attrib.ML	Tue Mar 08 16:02:52 2005 +0100
+++ b/src/Pure/Isar/attrib.ML	Wed Mar 09 18:44:52 2005 +0100
@@ -18,6 +18,7 @@
   exception ATTRIB_FAIL of (string * Position.T) * exn
   val global_attribute: theory -> Args.src -> theory attribute
   val local_attribute: theory -> Args.src -> Proof.context attribute
+  val multi_attribute: theory -> Args.src -> Locale.multi_attribute
   val local_attribute': Proof.context -> Args.src -> Proof.context attribute
   val undef_global_attribute: theory attribute
   val undef_local_attribute: Proof.context attribute
@@ -96,6 +97,8 @@
 
 val global_attribute = gen_attribute fst;
 val local_attribute = gen_attribute snd;
+fun multi_attribute thy src =
+      (global_attribute thy src, local_attribute thy src);
 val local_attribute' = local_attribute o ProofContext.theory_of;
 
 val undef_global_attribute: theory attribute =
--- a/src/Pure/Isar/delta_data.ML	Tue Mar 08 16:02:52 2005 +0100
+++ b/src/Pure/Isar/delta_data.ML	Wed Mar 09 18:44:52 2005 +0100
@@ -1,4 +1,5 @@
 (*  Author: Jia Meng, Cambridge University Computer Laboratory (06/01/05)
+    $Id$
     Copyright 2004 University of Cambridge
 
 Used for delta_simpset and delta_claset
--- a/src/Pure/Isar/isar_cmd.ML	Tue Mar 08 16:02:52 2005 +0100
+++ b/src/Pure/Isar/isar_cmd.ML	Wed Mar 09 18:44:52 2005 +0100
@@ -49,6 +49,7 @@
   val print_locales: Toplevel.transition -> Toplevel.transition
   val print_locale: Locale.expr * Args.src Locale.element list
     -> Toplevel.transition -> Toplevel.transition
+  val print_registrations: string -> Toplevel.transition -> Toplevel.transition
   val print_attributes: Toplevel.transition -> Toplevel.transition
   val print_rules: Toplevel.transition -> Toplevel.transition
   val print_induct_rules: Toplevel.transition -> Toplevel.transition
@@ -218,7 +219,12 @@
 
 fun print_locale (import, body) = Toplevel.unknown_theory o Toplevel.keep (fn state =>
   let val thy = Toplevel.theory_of state in
-    Locale.print_locale thy import (map (Locale.map_attrib_element (Attrib.local_attribute thy)) body)
+    Locale.print_locale thy import (map (Locale.map_attrib_element (Attrib.multi_attribute thy)) body)
+  end);
+
+fun print_registrations name = Toplevel.unknown_theory o
+  Toplevel.keep (fn state => let val thy = Toplevel.theory_of state in
+    Locale.print_global_registrations thy name
   end);
 
 val print_attributes = Toplevel.unknown_theory o
--- a/src/Pure/Isar/isar_syn.ML	Tue Mar 08 16:02:52 2005 +0100
+++ b/src/Pure/Isar/isar_syn.ML	Wed Mar 09 18:44:52 2005 +0100
@@ -310,6 +310,17 @@
         -- Scan.optional (P.$$$ "=" |-- P.!!! locale_val) (Locale.empty, [])
       >> (Toplevel.theory o IsarThy.add_locale o (fn ((x, y), (z, w)) => (x, y, z, w))));
 
+val uterm = P.underscore >> K NONE || P.term >> SOME;
+
+val view_val =
+  Scan.optional (P.$$$ "[" |-- Scan.repeat1 uterm --| P.$$$ "]") [];
+
+val registrationP =
+  OuterSyntax.command "registration"
+  "prove and register instance of locale expression" K.thy_goal
+  ((P.opt_thm_name ":" -- P.locale_expr -- P.!!! view_val
+     >> IsarThy.register_globally)
+   >> (Toplevel.print oo Toplevel.theory_to_proof));
 
 
 (** proof commands **)
@@ -579,9 +590,14 @@
     (Scan.succeed (Toplevel.no_timing o IsarCmd.print_locales));
 
 val print_localeP =
-  OuterSyntax.improper_command "print_locale" "print named locale of this theory" K.diag
+  OuterSyntax.improper_command "print_locale" "print locale expression in this theory" K.diag
     (locale_val >> (Toplevel.no_timing oo IsarCmd.print_locale));
 
+val print_registrationsP =
+  OuterSyntax.improper_command "print_registrations"
+    "print registrations of named locale in this theory" K.diag
+    (P.xname >> (Toplevel.no_timing oo IsarCmd.print_registrations));
+
 val print_attributesP =
   OuterSyntax.improper_command "print_attributes" "print attributes of this theory" K.diag
     (Scan.succeed (Toplevel.no_timing o IsarCmd.print_attributes));
@@ -784,10 +800,11 @@
   default_proofP, immediate_proofP, done_proofP, skip_proofP,
   forget_proofP, deferP, preferP, applyP, apply_endP, proofP, alsoP,
   finallyP, moreoverP, ultimatelyP, backP, cannot_undoP, clear_undosP,
-  redoP, undos_proofP, undoP, killP, instantiateP,
+  redoP, undos_proofP, undoP, killP, registrationP, instantiateP,
   (*diagnostic commands*)
   pretty_setmarginP, print_commandsP, print_contextP, print_theoryP,
   print_syntaxP, print_theoremsP, print_localesP, print_localeP,
+  print_registrationsP,
   print_attributesP, print_rulesP, print_induct_rulesP,
   print_trans_rulesP, print_methodsP, print_antiquotationsP,
   thms_containingP, thm_depsP, print_bindsP, print_lthmsP,
--- a/src/Pure/Isar/isar_thy.ML	Tue Mar 08 16:02:52 2005 +0100
+++ b/src/Pure/Isar/isar_thy.ML	Wed Mar 09 18:44:52 2005 +0100
@@ -22,8 +22,8 @@
     ((bstring * Args.src list) * (thmref * Args.src list) list) list
     -> theory -> theory * (bstring * thm list) list
   val locale_theorems_i: string -> string ->
-    ((bstring * Proof.context attribute list)
-      * (thm list * Proof.context attribute list) list) list
+    ((bstring * Locale.multi_attribute list)
+      * (thm list * Locale.multi_attribute list) list) list
     -> theory -> theory * (string * thm list) list
   val smart_theorems: string -> xstring option ->
     ((bstring * Args.src list) * (thmref * Args.src list) list) list
@@ -58,21 +58,21 @@
     -> bool -> theory -> ProofHistory.T
   val theorem_i: string -> (bstring * theory attribute list) *
     (term * (term list * term list)) -> bool -> theory -> ProofHistory.T
-  val multi_theorem: string -> bstring * Args.src list
+  val multi_theorem: string -> (theory -> theory) -> bstring * Args.src list
     -> Args.src Locale.element Locale.elem_expr list
     -> ((bstring * Args.src list) * (string * (string list * string list)) list) list
     -> bool -> theory -> ProofHistory.T
-  val multi_theorem_i: string -> bstring * theory attribute list
-    -> Proof.context attribute Locale.element_i Locale.elem_expr list
+  val multi_theorem_i: string -> (theory -> theory) -> bstring * theory attribute list
+    -> Locale.multi_attribute Locale.element_i Locale.elem_expr list
     -> ((bstring * theory attribute list) * (term * (term list * term list)) list) list
     -> bool -> theory -> ProofHistory.T
   val locale_multi_theorem: string -> xstring
     -> bstring * Args.src list -> Args.src Locale.element Locale.elem_expr list
     -> ((bstring * Args.src list) * (string * (string list * string list)) list) list
     -> bool -> theory -> ProofHistory.T
-  val locale_multi_theorem_i: string -> string -> bstring * Proof.context attribute list
-    -> Proof.context attribute Locale.element_i Locale.elem_expr list
-    -> ((bstring * Proof.context attribute list) * (term * (term list * term list)) list) list
+  val locale_multi_theorem_i: string -> string -> bstring * Locale.multi_attribute list
+    -> Locale.multi_attribute Locale.element_i Locale.elem_expr list
+    -> ((bstring * Locale.multi_attribute list) * (term * (term list * term list)) list) list
     -> bool -> theory -> ProofHistory.T
   val smart_multi_theorem: string -> xstring option
     -> bstring * Args.src list -> Args.src Locale.element Locale.elem_expr list
@@ -153,12 +153,16 @@
     -> Toplevel.transition -> Toplevel.transition
   val init_context: ('a -> unit) -> 'a -> Toplevel.transition -> Toplevel.transition
   val context: string -> Toplevel.transition -> Toplevel.transition
+
+  val register_globally:
+    ((string * Args.src list) * Locale.expr) * string option list ->
+    bool -> theory -> ProofHistory.T
+
 end;
 
 structure IsarThy: ISAR_THY =
 struct
 
-
 (** derived theory and proof operations **)
 
 (* name spaces *)
@@ -213,6 +217,7 @@
 
 fun global_attribs thy = prep_attribs (Attrib.global_attribute thy);
 fun local_attribs thy = prep_attribs (Attrib.local_attribute thy);
+fun multi_attribs thy = prep_attribs (Attrib.multi_attribute thy);
 
 end;
 
@@ -236,13 +241,13 @@
   |> present_thmss k;
 
 fun locale_theorems k loc args thy = thy
-  |> Locale.note_thmss k loc (local_attribs thy args)
+  |> Locale.note_thmss k loc (multi_attribs thy args)
   |> present_thmss k;
 
 fun smart_theorems k opt_loc args thy = thy
   |> (case opt_loc of
     NONE => PureThy.note_thmss (Drule.kind k) (global_attribs thy args)
-  | SOME loc => Locale.note_thmss k loc (local_attribs thy args))
+  | SOME loc => Locale.note_thmss k loc (multi_attribs thy args))
   |> present_thmss k;
 
 fun declare_theorems opt_loc args = #1 o smart_theorems "" opt_loc [(("", []), args)];
@@ -370,27 +375,27 @@
 
 in
 
-fun multi_theorem k a elems concl int thy =
-  global_statement (Proof.multi_theorem k NONE (apsnd (map (Attrib.global_attribute thy)) a)
-    (map (Locale.map_attrib_elem_or_expr (Attrib.local_attribute thy)) elems)) concl int thy;
+fun multi_theorem k thy_mod a elems concl int thy =
+  global_statement (Proof.multi_theorem k thy_mod NONE (apsnd (map (Attrib.global_attribute thy)) a)
+    (map (Locale.map_attrib_elem_or_expr (Attrib.multi_attribute thy)) elems)) concl int thy;
 
-fun multi_theorem_i k a = global_statement_i o Proof.multi_theorem_i k NONE a;
+fun multi_theorem_i k thy_mod a = global_statement_i o Proof.multi_theorem_i k thy_mod NONE a;
 
 fun locale_multi_theorem k locale (name, atts) elems concl int thy =
-  global_statement (Proof.multi_theorem k
-      (SOME (locale, (map (Attrib.local_attribute thy) atts,
-          map (map (Attrib.local_attribute thy) o snd o fst) concl)))
-      (name, []) (map (Locale.map_attrib_elem_or_expr (Attrib.local_attribute thy)) elems))
+  global_statement (Proof.multi_theorem k I
+      (SOME (locale, (map (Attrib.multi_attribute thy) atts,
+          map (map (Attrib.multi_attribute thy) o snd o fst) concl)))
+      (name, []) (map (Locale.map_attrib_elem_or_expr (Attrib.multi_attribute thy)) elems))
       (map (apfst (apsnd (K []))) concl) int thy;
 
 fun locale_multi_theorem_i k locale (name, atts) elems concl =
-  global_statement_i (Proof.multi_theorem_i k (SOME (locale, (atts, map (snd o fst) concl)))
+  global_statement_i (Proof.multi_theorem_i k I (SOME (locale, (atts, map (snd o fst) concl)))
       (name, []) elems) (map (apfst (apsnd (K []))) concl);
 
-fun theorem k (a, t) = multi_theorem k a [] [(("", []), [t])];
-fun theorem_i k (a, t) = multi_theorem_i k a [] [(("", []), [t])];
+fun theorem k (a, t) = multi_theorem k I a [] [(("", []), [t])];
+fun theorem_i k (a, t) = multi_theorem_i k I a [] [(("", []), [t])];
 
-fun smart_multi_theorem k NONE a elems = multi_theorem k a elems
+fun smart_multi_theorem k NONE a elems = multi_theorem k I a elems
   | smart_multi_theorem k (SOME locale) a elems = locale_multi_theorem k locale a elems;
 
 val assume      = local_statement Proof.assume I;
@@ -582,8 +587,7 @@
 
 fun add_locale (do_pred, name, import, body) thy =
   Locale.add_locale do_pred name import
-    (map (Locale.map_attrib_element (Attrib.local_attribute thy)) body) thy;
-
+    (map (Locale.map_attrib_element (Attrib.multi_attribute thy)) body) thy;
 
 (* theory init and exit *)
 
@@ -621,4 +625,23 @@
 val context = init_context (ThyInfo.quiet_update_thy true);
 
 
+(* global locale registration *)
+
+fun register_globally (((prfx, atts), expr), insts) b (* bool *) thy =
+  let
+    val (thy', propss, activate) =
+          Locale.prep_registration (prfx, []) expr insts thy;
+(* TODO: convert atts *)
+    fun register id (thy, thm) = let
+        val thm' = Drule.freeze_all thm;
+      in
+        (Locale.global_activate_thm id thm' thy, thm')
+      end;
+  in
+    multi_theorem_i Drule.internalK activate ("", []) [] 
+      (map (fn (id as (n, _), props) => ((NameSpace.base n, [register id]), 
+        map (fn prop => (prop, ([], []))) props)) propss) b thy'
+  end;
+
+
 end;
--- a/src/Pure/Isar/locale.ML	Tue Mar 08 16:02:52 2005 +0100
+++ b/src/Pure/Isar/locale.ML	Wed Mar 09 18:44:52 2005 +0100
@@ -21,6 +21,7 @@
 signature LOCALE =
 sig
   type context
+  type multi_attribute
 
   (* Constructors for elem, expr and elem_expr are
      currently only used for inputting locales (outer_parse.ML). *)
@@ -43,43 +44,55 @@
   val intern: Sign.sg -> xstring -> string
   val cond_extern: Sign.sg -> string -> xstring
   val the_locale: theory -> string -> locale
-  val map_attrib_element: ('att -> context attribute) -> 'att element ->
-    context attribute element
-  val map_attrib_element_i: ('att -> context attribute) -> 'att element_i ->
-    context attribute element_i
-  val map_attrib_elem_or_expr: ('att -> context attribute) ->
-    'att element elem_expr -> context attribute element elem_expr
-  val map_attrib_elem_or_expr_i: ('att -> context attribute) ->
-    'att element_i elem_expr -> context attribute element_i elem_expr
+  val map_attrib_element: ('att -> multi_attribute) -> 'att element ->
+    multi_attribute element
+  val map_attrib_element_i: ('att -> multi_attribute) -> 'att element_i ->
+    multi_attribute element_i
+  val map_attrib_elem_or_expr: ('att -> multi_attribute) ->
+    'att element elem_expr -> multi_attribute element elem_expr
+  val map_attrib_elem_or_expr_i: ('att -> multi_attribute) ->
+    'att element_i elem_expr -> multi_attribute element_i elem_expr
 
+  (* Processing of locale statements *)
   val read_context_statement: xstring option ->
-    context attribute element elem_expr list ->
+    multi_attribute element elem_expr list ->
     (string * (string list * string list)) list list -> context ->
     string option * (cterm list * cterm list) * context * context * 
       (term * (term list * term list)) list list
   val cert_context_statement: string option ->
-    context attribute element_i elem_expr list ->
+    multi_attribute element_i elem_expr list ->
     (term * (term list * term list)) list list -> context ->
     string option * (cterm list * cterm list) * context * context *
       (term * (term list * term list)) list list
+
+  (* Diagnostic functions *)
   val print_locales: theory -> unit
-  val print_locale: theory -> expr -> context attribute element list -> unit
-  val add_locale: bool -> bstring -> expr -> context attribute element list -> theory -> theory
-  val add_locale_i: bool -> bstring -> expr -> context attribute element_i list
+  val print_locale: theory -> expr -> multi_attribute element list -> unit
+  val print_global_registrations: theory -> string -> unit
+
+  val add_locale: bool -> bstring -> expr -> multi_attribute element list -> theory -> theory
+  val add_locale_i: bool -> bstring -> expr -> multi_attribute element_i list
     -> theory -> theory
   val smart_note_thmss: string -> (string * 'a) option ->
     ((bstring * theory attribute list) * (thm list * theory attribute list) list) list ->
     theory -> theory * (bstring * thm list) list
   val note_thmss: string -> xstring ->
-    ((bstring * context attribute list) * (thmref * context attribute list) list) list ->
+    ((bstring * multi_attribute list) * (thmref * multi_attribute list) list) list ->
     theory -> theory * (bstring * thm list) list
   val note_thmss_i: string -> string ->
-    ((bstring * context attribute list) * (thm list * context attribute list) list) list ->
+    ((bstring * multi_attribute list) * (thm list * multi_attribute list) list) list ->
     theory -> theory * (bstring * thm list) list
-  val add_thmss: string -> ((string * thm list) * context attribute list) list ->
+  val add_thmss: string -> ((string * thm list) * multi_attribute list) list ->
     theory * context -> (theory * context) * (string * thm list) list
+
   val instantiate: string -> string * context attribute list
     -> thm list option -> context -> context
+  val prep_registration:
+    string * theory attribute list -> expr -> string option list -> theory ->
+    theory * ((string * term list) * term list) list * (theory -> theory)
+  val global_activate_thm:
+    string * term list -> thm -> theory -> theory
+
   val setup: (theory -> theory) list
 end;
 
@@ -90,6 +103,10 @@
 
 type context = ProofContext.context;
 
+(* Locales store theory attributes (for activation in theories)
+   and context attributes (for activation in contexts) *)
+type multi_attribute = theory attribute * context attribute;
+
 datatype ('typ, 'term, 'fact, 'att) elem =
   Fixes of (string * 'typ option * mixfix option) list |
   Assumes of ((string * 'att list) * ('term * ('term list * 'term list)) list) list |
@@ -121,29 +138,44 @@
        (cf. [1], normalisation of locale expressions.)
     *)
   import: expr,                                       (*dynamic import*)
-  elems: (context attribute element_i * stamp) list,  (*static content*)
+  elems: (multi_attribute element_i * stamp) list,  (*static content*)
   params: (string * typ option) list * string list}   (*all/local params*)
 
 
 (** theory data **)
 
+structure Termlisttab = TableFun(type key = term list
+  val ord = Library.list_ord Term.term_ord);
+
 structure LocalesArgs =
 struct
   val name = "Isar/locales";
-  type T = NameSpace.T * locale Symtab.table;
+  type T = NameSpace.T * locale Symtab.table *
+      ((string * theory attribute list) * thm list) Termlisttab.table
+        Symtab.table;
+    (* 1st entry: locale namespace,
+       2nd entry: locales of the theory,
+       3rd entry: registrations: theorems instantiating locale assumptions,
+         with prefix and attributes, indexed by locale name and parameter
+         instantiation *)
 
-  val empty = (NameSpace.empty, Symtab.empty);
+  val empty = (NameSpace.empty, Symtab.empty, Symtab.empty);
   val copy = I;
   val prep_ext = I;
 
-  (*joining of locale elements: only facts may be added later!*)
-  fun join ({predicate, import, elems, params}: locale, {elems = elems', ...}: locale) =
-    SOME {predicate = predicate, import = import, elems = gen_merge_lists eq_snd elems elems',
+  fun join_locs ({predicate, import, elems, params}: locale,
+      {elems = elems', ...}: locale) =
+    SOME {predicate = predicate, import = import,
+      elems = gen_merge_lists eq_snd elems elems',
       params = params};
-  fun merge ((space1, locs1), (space2, locs2)) =
-    (NameSpace.merge (space1, space2), Symtab.join join (locs1, locs2));
+  (* joining of registrations: prefix and attributes of left theory,
+     thmsss are equal *)
+  fun join_regs (reg, _) = SOME reg;
+  fun merge ((space1, locs1, regs1), (space2, locs2, regs2)) =
+    (NameSpace.merge (space1, space2), Symtab.join join_locs (locs1, locs2),
+     Symtab.join (SOME o Termlisttab.join join_regs) (regs1, regs2));
 
-  fun print _ (space, locs) =
+  fun print _ (space, locs, _) =
     Pretty.strs ("locales:" :: map (NameSpace.cond_extern space o #1) (Symtab.dest locs))
     |> Pretty.writeln;
 end;
@@ -158,9 +190,13 @@
 (* access locales *)
 
 fun declare_locale name =
-  LocalesData.map (apfst (fn space => (NameSpace.extend (space, [name]))));
+  LocalesData.map (fn (space, locs, regs) =>
+    (NameSpace.extend (space, [name]), locs, regs));
 
-fun put_locale name loc = LocalesData.map (apsnd (fn locs => Symtab.update ((name, loc), locs)));
+fun put_locale name loc =
+  LocalesData.map (fn (space, locs, regs) =>
+    (space, Symtab.update ((name, loc), locs), regs));
+
 fun get_locale thy name = Symtab.lookup (#2 (LocalesData.get thy), name);
 
 fun the_locale thy name =
@@ -169,6 +205,38 @@
   | NONE => error ("Unknown locale " ^ quote name));
 
 
+(* access registrations *)
+
+(* add registration to theory, ignored if already present *)
+
+fun global_put_registration (name, ps) attn =
+  LocalesData.map (fn (space, locs, regs) =>
+    (space, locs, let
+        val tab = getOpt (Symtab.lookup (regs, name), Termlisttab.empty);
+      in
+        Symtab.update ((name, Termlisttab.update_new ((ps, (attn, [])), tab)),
+          regs)
+      end handle Termlisttab.DUP _ => regs));
+
+(* add theorem to registration in theory,
+   ignored if registration not present *)
+
+fun global_put_registration_thm (name, ps) thm =
+  LocalesData.map (fn (space, locs, regs) =>
+    (space, locs, let
+        val tab = valOf (Symtab.lookup (regs, name));
+        val (x, thms) = valOf (Termlisttab.lookup (tab, ps));
+      in
+        Symtab.update ((name, Termlisttab.update ((ps, (x, thm::thms)), tab)),
+          regs)
+      end handle Option => regs))
+
+fun global_get_registration thy (name, ps) =
+  case Symtab.lookup (#3 (LocalesData.get thy), name) of
+      NONE => NONE
+    | SOME tab => Termlisttab.lookup (tab, ps);
+
+
 (* import hierarchy
    implementation could be more efficient, eg. by maintaining a database
    of dependencies *)
@@ -187,6 +255,29 @@
   end;
 
 
+(* registrations *)
+
+fun print_global_registrations thy loc =
+  let
+    val sg = Theory.sign_of thy;
+    val loc_int = intern sg loc;
+    val (_, _, regs) = LocalesData.get thy;
+    val prt_term = Pretty.quote o Sign.pretty_term sg;
+    fun prt_inst (ts, ((prfx, _), thms)) = let
+in
+      Pretty.block [Pretty.str prfx, Pretty.str ":", Pretty.brk 1,
+        Pretty.list "(" ")" (map prt_term ts)]
+end;
+    val loc_regs = Symtab.lookup (regs, loc_int);
+  in
+    (case loc_regs of
+        NONE => Pretty.str "No registrations."
+      | SOME r => Pretty.big_list "registrations:"
+          (map prt_inst (Termlisttab.dest r)))
+    |> Pretty.writeln
+  end;
+
+
 (* diagnostics *)
 
 fun err_in_locale ctxt msg ids =
@@ -425,9 +516,34 @@
             map (inst_thm ctxt env) axs), map (inst_elem ctxt env) elems);
       in map inst (elemss ~~ envs) end;
 
+(* flatten_expr:
+   Extend list of identifiers by those new in locale expression expr.
+   Compute corresponding list of lists of locale elements (one entry per
+   identifier).
+
+   Identifiers represent locale fragments and are in an extended form:
+     ((name, ps), (ax_ps, axs))
+   (name, ps) is the locale name with all its parameters.
+   (ax_ps, axs) is the locale axioms with its parameters;
+     axs are always taken from the top level of the locale hierarchy,
+     hence axioms may contain additional parameters from later fragments:
+     ps subset of ax_ps.  axs is either singleton or empty.
+
+   Elements are enriched by identifier-like information:
+     (((name, ax_ps), axs), elems)
+   The parameters in ax_ps are the axiom parameters, but enriched by type
+   info: now each entry is a pair of string and typ option.  Axioms are
+   type-instantiated.
+
+*)
+
 fun flatten_expr ctxt (prev_idents, expr) =
   let
     val thy = ProofContext.theory_of ctxt;
+    (* thy used for retrieval of locale info,
+       ctxt for error messages, parameter unification and instantiation
+       of axioms *)
+    (* TODO: consider err_in_locale with thy argument *)
 
     fun renaming (SOME x :: xs) (y :: ys) = (y, x) :: renaming xs ys
       | renaming (NONE :: xs) (y :: ys) = renaming xs ys
@@ -444,7 +560,7 @@
       end;
 
     fun identify top (Locale name) =
-    (* CB: ids is a list of tuples of the form ((name, ps)  axs),
+    (* CB: ids_ax is a list of tuples of the form ((name, ps), axs),
        where name is a locale name, ps a list of parameter names and axs
        a list of axioms relating to the identifier, axs is empty unless
        identify at top level (top = true);
@@ -519,6 +635,36 @@
 end;
 
 
+(* attributes *)
+
+local
+
+fun read_att attrib (x, srcs) = (x, map attrib srcs)
+
+(* CB: Map attrib over
+   * A context element: add attrib to attribute lists of assumptions,
+     definitions and facts (on both sides for facts).
+   * Locale expression: no effect. *)
+
+fun gen_map_attrib_elem _ (Fixes fixes) = Fixes fixes
+  | gen_map_attrib_elem attrib (Assumes asms) = Assumes (map (apfst (read_att attrib)) asms)
+  | gen_map_attrib_elem attrib (Defines defs) = Defines (map (apfst (read_att attrib)) defs)
+  | gen_map_attrib_elem attrib (Notes facts) =
+      Notes (map (apfst (read_att attrib) o apsnd (map (read_att attrib))) facts)
+
+fun gen_map_attrib_elem_expr attrib (Elem elem) = Elem (gen_map_attrib_elem attrib elem)
+  | gen_map_attrib_elem_expr _ (Expr expr) = Expr expr;
+
+in
+
+val map_attrib_element = gen_map_attrib_elem;
+val map_attrib_element_i = gen_map_attrib_elem;
+val map_attrib_elem_or_expr = gen_map_attrib_elem_expr;
+val map_attrib_elem_or_expr_i = gen_map_attrib_elem_expr;
+
+end;
+
+
 (* activate elements *)
 
 local
@@ -534,6 +680,8 @@
       ((ctxt |> ProofContext.add_fixes fixes, axs), [])
   | activate_elem _ ((ctxt, axs), Assumes asms) =
       let
+        (* extract context attributes *)
+        val (Assumes asms) = map_attrib_element_i snd (Assumes asms);
         val ts = List.concat (map (map #1 o #2) asms);
         val (ps,qs) = splitAt (length ts, axs);
         val (ctxt', _) =
@@ -541,14 +689,20 @@
           |> ProofContext.assume_i (export_axioms ps) asms;
       in ((ctxt', qs), []) end
   | activate_elem _ ((ctxt, axs), Defines defs) =
-      let val (ctxt', _) =
+      let
+        (* extract context attributes *)
+        val (Defines defs) = map_attrib_element_i snd (Defines defs);
+        val (ctxt', _) =
         ctxt |> ProofContext.assume_i ProofContext.export_def
           (defs |> map (fn ((name, atts), (t, ps)) =>
             let val (c, t') = ProofContext.cert_def ctxt t
             in ((if name = "" then Thm.def_name c else name, atts), [(t', (ps, []))]) end))
       in ((ctxt', axs), []) end
   | activate_elem is_ext ((ctxt, axs), Notes facts) =
-      let val (ctxt', res) = ctxt |> ProofContext.note_thmss_i facts
+      let
+        (* extract context attributes *)
+        val (Notes facts) = map_attrib_element_i snd (Notes facts);
+        val (ctxt', res) = ctxt |> ProofContext.note_thmss_i facts
       in ((ctxt', axs), if is_ext then res else []) end;
 
 fun activate_elems (((name, ps), axs), elems) ctxt =
@@ -591,36 +745,6 @@
   | intern_expr sg (Rename (expr, xs)) = Rename (intern_expr sg expr, xs);
 
 
-(* attributes *)
-
-local
-
-fun read_att attrib (x, srcs) = (x, map attrib srcs)
-
-(* CB: Map attrib over
-   * A context element: add attrib to attribute lists of assumptions,
-     definitions and facts (on both sides for facts).
-   * Locale expression: no effect. *)
-
-fun gen_map_attrib_elem _ (Fixes fixes) = Fixes fixes
-  | gen_map_attrib_elem attrib (Assumes asms) = Assumes (map (apfst (read_att attrib)) asms)
-  | gen_map_attrib_elem attrib (Defines defs) = Defines (map (apfst (read_att attrib)) defs)
-  | gen_map_attrib_elem attrib (Notes facts) =
-      Notes (map (apfst (read_att attrib) o apsnd (map (read_att attrib))) facts)
-
-fun gen_map_attrib_elem_expr attrib (Elem elem) = Elem (gen_map_attrib_elem attrib elem)
-  | gen_map_attrib_elem_expr _ (Expr expr) = Expr expr;
-
-in
-
-val map_attrib_element = gen_map_attrib_elem;
-val map_attrib_element_i = gen_map_attrib_elem;
-val map_attrib_elem_or_expr = gen_map_attrib_elem_expr;
-val map_attrib_elem_or_expr_i = gen_map_attrib_elem_expr;
-
-end;
-
-
 (* parameters *)
 
 local
@@ -644,19 +768,20 @@
 
 datatype ('a, 'b) int_ext = Int of 'a | Ext of 'b;
 
-(* CB: flatten (ids, expr) normalises expr (which is either a locale
+(* flatten (ids, expr) normalises expr (which is either a locale
    expression or a single context element) wrt.
    to the list ids of already accumulated identifiers.
    It returns (ids', elemss) where ids' is an extension of ids
    with identifiers generated for expr, and elemss is the list of
-   context elements generated from expr, decorated with additional
-   information (for expr it is the identifier, where parameters additionially
-   contain type information (extracted from the locale record), for a Fixes
-   element, it is an identifier with name = "" and parameters with type
-   information NONE, for other elements it is simply ("", [])).
+   context elements generated from expr.  For details, see flatten_expr.
+   Additionally, for a locale expression, the elems are grouped into a single
+   Int; individual context elements are marked Ext.  In this case, the
+   identifier-like information of the element is as follows:
+   - for Fixes: (("", ps), []) where the ps have type info NONE
+   - for other elements: (("", []), []).
    The implementation of activate_facts relies on identifier names being
    empty strings for external elements.
-TODO: correct this comment wrt axioms. *)
+*)
 
 fun flatten _ (ids, Elem (Fixes fixes)) =
       (ids, [((("", map (rpair NONE o #1) fixes), []), Ext (Fixes fixes))])
@@ -711,11 +836,13 @@
 
 local
 
-(* CB: following code (norm_term, abstract_term, abstract_thm, bind_def)
-   used in eval_text for defines elements. *)
+(* CB: normalise Assumes and Defines wrt. previous definitions *)
 
 val norm_term = Envir.beta_norm oo Term.subst_atomic;
 
+(* CB: following code (abstract_term, abstract_thm, bind_def)
+   used in eval_text for Defines elements. *)
+
 fun abstract_term eq =    (*assumes well-formedness according to ProofContext.cert_def*)
   let
     val body = Term.strip_all_body eq;
@@ -997,7 +1124,7 @@
       the_locale thy (intern sign loc_name);
     val fixed_params = param_types ps;
     val init = ProofContext.init thy;
-    val (ids, raw_elemss) =
+    val (_, raw_elemss) =
           flatten (init, intern_expr sign) ([], Expr (Locale loc_name));
     val ((parms, all_elemss, concl),
          (spec as (_, (ints, _)), (xs, env, defs))) =
@@ -1127,7 +1254,9 @@
     fun inst_elem (ctxt, (Ext _)) = ctxt
       | inst_elem (ctxt, (Int (Notes facts))) =
               (* instantiate fact *)
-          let val facts' =
+          let (* extract context attributes *)
+              val (Notes facts) = map_attrib_element_i snd (Notes facts);
+              val facts' =
                 map (apsnd (map (apfst (map inst_thm')))) facts
 		handle THM (msg, n, thms) => error ("Exception THM " ^
 		  string_of_int n ^ " raised\n" ^
@@ -1210,6 +1339,31 @@
   |>> hide_bound_names (map (#1 o #1) args)
   |>> Theory.parent_path;
 
+fun note_thms_qualified' kind (arg as ((name, atts), thms)) thy =
+  let
+    val qname = NameSpace.unpack name
+  in
+    if length qname <= 1
+    then PureThy.note_thmss_i kind [arg] thy
+  else let val (prfx, n) = split_last qname
+    in thy
+      |> Theory.add_path (NameSpace.pack prfx)
+      |> PureThy.note_thmss_i kind [((n, atts), thms)]
+      |>> funpow (length prfx) Theory.parent_path
+    end
+  end;
+
+(* prfx may be empty (not yet), names (in args) may be qualified *)
+
+fun note_thmss_qualified' kind prfx args thy =
+  thy
+  |> Theory.add_path (Sign.base_name prfx)
+  |> (fn thy => Library.foldl (fn ((thy, res), arg) =>
+        let val (thy', res') = note_thms_qualified' (Drule.kind kind) arg thy
+        in (thy', res @ res') end) ((thy, []), args))
+(*  |>> hide_bound_names (map (#1 o #1) args) *)
+  |>> Theory.parent_path;
+
 fun smart_note_thmss kind NONE = PureThy.note_thmss_i (Drule.kind kind)
   | smart_note_thmss kind (SOME (loc, _)) = note_thmss_qualified kind loc;
   (* CB: only used in Proof.finish_global. *)
@@ -1233,8 +1387,13 @@
     val (_, (stmt, _), loc_ctxt, _, _) =
       cert_context_statement (SOME loc) [] [] thy_ctxt;
     val args = map (apsnd (map (apfst (get_thms loc_ctxt)))) raw_args;
+    (* convert multi attributes to context attributes for
+       ProofContext.note_thmss_i *)
+    val args'' = args
+          |> map (apfst (apsnd (map snd)))
+          |> map (apsnd (map (apsnd (map snd))));
     val export = ProofContext.export_standard stmt loc_ctxt thy_ctxt;
-    val results = map (map export o #2) (#2 (ProofContext.note_thmss_i args loc_ctxt));
+    val results = map (map export o #2) (#2 (ProofContext.note_thmss_i args'' loc_ctxt));
     val args' = map (rpair [] o #1 o #1) args ~~ map (single o Thm.no_attributes) results;
   in
     thy
@@ -1264,7 +1423,7 @@
 
 
 (* predicate text *)
-(* CB: generate locale predicates (and delta predicates) *)
+(* CB: generate locale predicates and delta predicates *)
 
 local
 
@@ -1403,8 +1562,8 @@
 local
 
 fun gen_add_locale prep_ctxt prep_expr do_pred bname raw_import raw_body thy =
-  (* CB: do_pred = false means old-style locale, declared with (open).
-     Old-style locales don't define predicates. *)
+  (* CB: do_pred controls generation of predicates.
+         true -> with, false -> without predicates. *)
   let
     val sign = Theory.sign_of thy;
     val name = Sign.full_name sign bname;
@@ -1451,6 +1610,272 @@
 
 
 
+(** Registration commands **)
+
+local
+
+(** instantiate free vars **)
+
+(* instantiate TFrees *)
+
+fun tinst_type tinst T = if Symtab.is_empty tinst
+      then T
+      else Term.map_type_tfree
+        (fn (v as (x, _)) => getOpt (Symtab.lookup (tinst, x), (TFree v))) T;
+
+fun tinst_term tinst t = if Symtab.is_empty tinst
+      then t
+      else Term.map_term_types (tinst_type tinst) t;
+
+fun tinst_thm sg tinst thm = if Symtab.is_empty tinst
+      then thm
+      else let
+          val cert = Thm.cterm_of sg;
+          val certT = Thm.ctyp_of sg;
+          val {hyps, prop, ...} = Thm.rep_thm thm;
+          val tfrees = foldr Term.add_term_tfree_names [] (prop :: hyps);
+          val tinst' = Symtab.dest tinst |>
+                List.filter (fn (a, _) => a mem_string tfrees);
+        in
+          if null tinst' then thm
+          else thm |> Drule.implies_intr_list (map cert hyps)
+            |> Drule.tvars_intr_list (map #1 tinst')
+            |> (fn (th, al) => th |> Thm.instantiate
+                ((map (fn (a, T) => (valOf (assoc (al, a)), certT T)) tinst'),
+                  []))
+            |> (fn th => Drule.implies_elim_list th
+                 (map (Thm.assume o cert o tinst_term tinst) hyps))
+        end;
+
+fun tinst_elem _ tinst (Fixes fixes) =
+      Fixes (map (fn (x, T, mx) => (x, Option.map (tinst_type tinst) T, mx)) fixes)
+  | tinst_elem _ tinst (Assumes asms) =
+      Assumes (map (apsnd (map (fn (t, (ps, qs)) =>
+        (tinst_term tinst t,
+          (map (tinst_term tinst) ps, map (tinst_term tinst) qs))))) asms)
+  | tinst_elem _ tinst (Defines defs) =
+      Defines (map (apsnd (fn (t, ps) =>
+        (tinst_term tinst t, map (tinst_term tinst) ps))) defs)
+  | tinst_elem sg tinst (Notes facts) =
+      Notes (map (apsnd (map (apfst (map (tinst_thm sg tinst))))) facts);
+
+(* instantiate TFrees and Frees *)
+
+fun inst_term (inst, tinst) = if Symtab.is_empty inst
+      then tinst_term tinst
+      else (* instantiate terms and types simultaneously *)
+        let
+          fun instf (Const (x, T)) = Const (x, tinst_type tinst T)
+            | instf (Free (x, T)) = (case Symtab.lookup (inst, x) of
+                 NONE => Free (x, tinst_type tinst T)
+               | SOME t => t)
+            | instf (Var (xi, T)) = Var (xi, tinst_type tinst T)
+            | instf (b as Bound _) = b
+            | instf (Abs (x, T, t)) = Abs (x, tinst_type tinst T, instf t)
+            | instf (s $ t) = instf s $ instf t
+        in instf end;
+
+fun inst_thm sg (inst, tinst) thm = if Symtab.is_empty inst
+      then tinst_thm sg tinst thm
+      else let
+          val cert = Thm.cterm_of sg;
+          val certT = Thm.ctyp_of sg;
+          val {hyps, prop, ...} = Thm.rep_thm thm;
+          (* type instantiations *)
+          val tfrees = foldr Term.add_term_tfree_names [] (prop :: hyps);
+          val tinst' = Symtab.dest tinst |>
+                List.filter (fn (a, _) => a mem_string tfrees);
+          (* term instantiations;
+             note: lhss are type instantiated, because
+                   type insts will be done first*)
+          val frees = foldr Term.add_term_frees [] (prop :: hyps);
+          val inst' = Symtab.dest inst |>
+                List.mapPartial (fn (a, t) =>
+                  get_first (fn (Free (x, T)) => 
+                    if a = x then SOME (Free (x, tinst_type tinst T), t)
+                    else NONE) frees);
+        in
+          if null tinst' andalso null inst' then tinst_thm sg tinst thm
+          else thm |> Drule.implies_intr_list (map cert hyps)
+            |> Drule.tvars_intr_list (map #1 tinst')
+            |> (fn (th, al) => th |> Thm.instantiate
+                ((map (fn (a, T) => (valOf (assoc (al, a)), certT T)) tinst'),
+                  []))
+            |> Drule.forall_intr_list (map (cert o #1) inst')
+            |> Drule.forall_elim_list (map (cert o #2) inst') 
+            |> (fn th => Drule.implies_elim_list th
+                 (map (Thm.assume o cert o inst_term (inst, tinst)) hyps))
+        end;
+
+fun inst_elem _ (_, tinst) (Fixes fixes) =
+      Fixes (map (fn (x, T, mx) => (x, Option.map (tinst_type tinst) T, mx)) fixes)
+  | inst_elem _ inst (Assumes asms) =
+      Assumes (map (apsnd (map (fn (t, (ps, qs)) =>
+        (inst_term inst t,
+          (map (inst_term inst) ps, map (inst_term inst) qs))))) asms)
+  | inst_elem _ inst (Defines defs) =
+      Defines (map (apsnd (fn (t, ps) =>
+        (inst_term inst t, map (inst_term inst) ps))) defs)
+  | inst_elem sg inst (Notes facts) =
+      Notes (map (apsnd (map (apfst (map (inst_thm sg inst))))) facts);
+
+fun inst_elems sign inst ((n, ps), elems) =
+      ((n, map (inst_term inst) ps), map (inst_elem sign inst) elems);
+
+(* extract proof obligations (assms and defs) from elements *)
+
+(* check if defining equation has become t == t, these are dropped
+   in extract_asms_elem, as they lead to trivial proof obligations *)
+(* currently disabled *)
+fun check_def (_, (def, _)) = SOME def;
+(*
+fun check_def (_, (def, _)) =
+      if def |> Logic.dest_equals |> op aconv
+      then NONE else SOME def;
+*)
+
+fun extract_asms_elem (ts, Fixes _) = ts
+  | extract_asms_elem (ts, Assumes asms) =
+      ts @ List.concat (map (fn (_, ams) => map (fn (t, _) => t) ams) asms)
+  | extract_asms_elem (ts, Defines defs) =
+      ts @ List.mapPartial check_def defs
+  | extract_asms_elem (ts, Notes _) = ts;
+
+fun extract_asms_elems (id, elems) =
+      (id, Library.foldl extract_asms_elem ([], elems));
+
+fun extract_asms_elemss elemss =
+      map extract_asms_elems elemss;
+
+(* add registration, without theorems, to theory *)
+
+fun prep_reg_global attn (thy, (id, _)) =
+      global_put_registration id attn thy;
+
+(* activate instantiated facts in theory *)
+
+fun activate_facts_elem _ _ (thy, Fixes _) = thy
+  | activate_facts_elem _ _ (thy, Assumes _) = thy
+  | activate_facts_elem _ _ (thy, Defines _) = thy
+  | activate_facts_elem disch (prfx, atts) (thy, Notes facts) =
+      let
+        (* extract theory attributes *)
+        val (Notes facts) = map_attrib_element_i fst (Notes facts);
+        val facts' = map (apfst (apsnd (fn a => a @ atts))) facts;
+        val facts'' = map (apsnd (map (apfst (map disch)))) facts';
+      in
+        fst (note_thmss_qualified' "" prfx facts thy)
+      end;
+
+fun activate_facts_elems disch (thy, (id, elems)) =
+      let
+        val ((prfx, atts2), _) = valOf (global_get_registration thy id)
+          handle Option => error ("(internal) unknown registration of " ^
+            quote (fst id) ^ " while activating facts.");
+      in
+        Library.foldl (activate_facts_elem disch (prfx, atts2)) (thy, elems)
+      end;
+
+fun activate_facts_elemss all_elemss new_elemss thy =
+      let
+        val prems = List.concat (List.mapPartial (fn (id, _) =>
+              Option.map snd (global_get_registration thy id)
+                handle Option => error ("(internal) unknown registration of " ^
+                  quote (fst id) ^ " while activating facts.")) all_elemss);
+        fun disch thm = let
+          in Drule.satisfy_hyps prems thm end;
+      in Library.foldl (activate_facts_elems disch) (thy, new_elemss) end;
+
+in
+
+fun prep_registration attn expr insts thy =
+  let
+    val ctxt = ProofContext.init thy;
+    val sign = Theory.sign_of thy;
+    val tsig = Sign.tsig_of sign;
+
+    val (ids, raw_elemss) =
+          flatten (ctxt, intern_expr sign) ([], Expr expr);
+    val do_close = false;  (* effect unknown *)
+    val ((parms, all_elemss, _), (spec, (xs, defs, _))) =
+          read_elemss do_close ctxt [] raw_elemss [];
+
+
+    (** compute instantiation **)
+
+    (* user input *)
+    val insts = if length parms < length insts
+         then error "More arguments than parameters in instantiation."
+         else insts @ replicate (length parms - length insts) NONE;
+    val (ps, pTs) = split_list parms;
+    val pvTs = map Type.varifyT pTs;
+    val given = List.mapPartial (fn (_, (NONE, _)) => NONE
+         | (x, (SOME inst, T)) => SOME (x, (inst, T))) (ps ~~ (insts ~~ pvTs));
+    val (given_ps, given_insts) = split_list given;
+    val tvars = foldr Term.add_typ_tvars [] pvTs;
+    val used = foldr Term.add_typ_varnames [] pvTs;
+    fun sorts (a, i) = assoc (tvars, (a, i));
+    val (tvs, tvinst) = Sign.read_def_terms (sign, K NONE, sorts) used true
+         given_insts;
+    val tinst = Symtab.make (map (fn ((x, 0), T) => (x, Type.unvarifyT T)
+                  | ((_, n), _) => error "Var in prep_registration") tvinst);
+    val inst = Symtab.make (given_ps ~~ map Logic.unvarify tvs);
+
+    (* defined params without user input *)
+    val not_given = List.mapPartial (fn (x, (NONE, T)) => SOME (x, T)
+         | (_, (SOME _, _)) => NONE) (ps ~~ (insts ~~ pTs));
+    fun add_def ((inst, tinst), (p, pT)) =
+      let
+        val (t, T) = case find_first (fn (Free (a, _), _) => a = p) defs of
+               NONE => error ("Instance missing for parameter " ^ quote p)
+             | SOME (Free (_, T), t) => (t, T);
+        val d = t |> inst_term (inst, tinst) |> Envir.beta_norm;
+      in (Symtab.update_new ((p, d), inst), tinst) end;
+    val (inst, tinst) = Library.foldl add_def ((inst, tinst), not_given);
+  
+
+    (** compute proof obligations **)
+
+    (* restore small ids *)
+    val ids' = map (fn ((n, ps), _) =>
+          (n, map (fn p => Free (p, valOf (assoc (parms, p)))) ps)) ids;
+
+    (* instantiate ids and elements *)
+    val inst_elemss = map
+          (fn (id, (_, elems)) => inst_elems sign (inst, tinst) (id, 
+            map (fn Int e => e) elems)) 
+          (ids' ~~ all_elemss);
+
+(*
+    (* varify ids, props are varified after they are proved *)
+    val inst_elemss' = map (fn ((n, ps), elems) =>
+          ((n, map Logic.varify ps), elems)) inst_elemss;
+*)
+
+    (* remove fragments already registered with theory *)
+    val new_inst_elemss = List.filter (fn (id, _) =>
+          is_none (global_get_registration thy id)) inst_elemss;
+
+
+    val propss = extract_asms_elemss new_inst_elemss;
+
+
+    (** add registrations to theory,
+        without theorems, these are added after the proof **)
+
+    val thy' = Library.foldl (prep_reg_global attn) (thy, new_inst_elemss);
+
+  in (thy', propss, activate_facts_elemss inst_elemss new_inst_elemss) end;
+
+
+(* Add registrations and theorems to theory context *)
+
+val global_activate_thm = global_put_registration_thm
+
+end;  (* local *)
+
+
+
 (** locale theory setup **)
 
 val setup =
--- a/src/Pure/Isar/proof.ML	Tue Mar 08 16:02:52 2005 +0100
+++ b/src/Pure/Isar/proof.ML	Wed Mar 09 18:44:52 2005 +0100
@@ -90,15 +90,15 @@
   val def: string -> context attribute list -> string *  (string * string list) -> state -> state
   val def_i: string -> context attribute list -> string * (term * term list) -> state -> state
   val invoke_case: string * string option list * context attribute list -> state -> state
-  val multi_theorem: string
-    -> (xstring * (context attribute list * context attribute list list)) option
-    -> bstring * theory attribute list -> context attribute Locale.element Locale.elem_expr list
+  val multi_theorem: string -> (theory -> theory)
+    -> (xstring * (Locale.multi_attribute list * Locale.multi_attribute list list)) option
+    -> bstring * theory attribute list -> Locale.multi_attribute Locale.element Locale.elem_expr list
     -> ((bstring * theory attribute list) * (string * (string list * string list)) list) list
     -> theory -> state
-  val multi_theorem_i: string
-    -> (string * (context attribute list * context attribute list list)) option
+  val multi_theorem_i: string -> (theory -> theory)
+    -> (string * (Locale.multi_attribute list * Locale.multi_attribute list list)) option
     -> bstring * theory attribute list
-    -> context attribute Locale.element_i Locale.elem_expr list
+    -> Locale.multi_attribute Locale.element_i Locale.elem_expr list
     -> ((bstring * theory attribute list) * (term * (term list * term list)) list) list
     -> theory -> state
   val chain: state -> state
@@ -153,9 +153,11 @@
 datatype kind =
   Theorem of {kind: string,
     theory_spec: (bstring * theory attribute list) * theory attribute list list,
-    locale_spec: (string * (context attribute list * context attribute list list)) option,
-    view: cterm list * cterm list} |
-(* target view and includes view *)
+    locale_spec: (string * (Locale.multi_attribute list * Locale.multi_attribute list list)) option,
+    view: cterm list * cterm list, (* target view and includes view *)
+    thy_mod: theory -> theory} |   (* used in finish_global to modify final
+      theory, normally set to I;
+      used by registration command to activate registrations *)
   Show of context attribute list list |
   Have of context attribute list list;
 
@@ -163,9 +165,9 @@
    the proof state *)
 
 fun kind_name _ (Theorem {kind = s, theory_spec = ((a, _), _),
-        locale_spec = NONE, view = _}) = s ^ (if a = "" then "" else " " ^ a)
+        locale_spec = NONE, ...}) = s ^ (if a = "" then "" else " " ^ a)
   | kind_name sg (Theorem {kind = s, theory_spec = ((a, _), _),
-        locale_spec = SOME (name, _), view = _}) =
+        locale_spec = SOME (name, _), ...}) =
       s ^ " (in " ^ Locale.cond_extern sg name ^ ")" ^ (if a = "" then "" else " " ^ a)
   | kind_name _ (Show _) = "show"
   | kind_name _ (Have _) = "have";
@@ -782,7 +784,7 @@
   end;
 
 (*global goals*)
-fun global_goal prep kind raw_locale a elems concl thy =
+fun global_goal prep kind thy_mod raw_locale a elems concl thy =
   let
     val init = init_state thy;
     val (opt_name, view, locale_ctxt, elems_ctxt, propp) =
@@ -799,7 +801,8 @@
       (Theorem {kind = kind,
         theory_spec = (a, map (snd o fst) concl),
         locale_spec = (case raw_locale of NONE => NONE | SOME (_, x) => SOME (valOf opt_name, x)),
-        view = view})
+        view = view,
+        thy_mod = thy_mod})
       Seq.single true (map (fst o fst) concl) propp
   end;
 
@@ -900,7 +903,7 @@
     val (goal_ctxt, (((kind, names, tss), (_, raw_thm)), _)) = current_goal state;
     val locale_ctxt = context_of (state |> close_block);
     val theory_ctxt = context_of (state |> close_block |> close_block);
-    val {kind = k, theory_spec = ((name, atts), attss), locale_spec, view = (target_view, elems_view)} =
+    val {kind = k, theory_spec = ((name, atts), attss), locale_spec, view = (target_view, elems_view), thy_mod} =
       (case kind of Theorem x => x | _ => err_malformed "finish_global" state);
 
     val ts = List.concat tss;
@@ -926,7 +929,7 @@
       |> (fn (thy, res) => (thy, res)
           |>> (#1 o Locale.smart_note_thmss k locale_spec
             [((name, atts), [(List.concat (map #2 res), [])])]));
-  in (theory', ((k, name), results')) end;
+  in (thy_mod theory', ((k, name), results')) end;
 
 
 (*note: clients should inspect first result only, as backtracking may destroy theory*)
--- a/src/Pure/Isar/proof_context.ML	Tue Mar 08 16:02:52 2005 +0100
+++ b/src/Pure/Isar/proof_context.ML	Wed Mar 09 18:44:52 2005 +0100
@@ -635,7 +635,10 @@
 
 in
 
+(* CB: for attribute where.  See attrib.ML. *)
 val read_termTs           = gen_read' (read_def_termTs false) (apfst o map) false false false;
+
+(* CB: for rule_tac etc.  See method.ML. *)
 val read_termTs_schematic = gen_read' (read_def_termTs false) (apfst o map) false false true;
 
 fun read_term_pats T ctxt =
--- a/src/Pure/Isar/toplevel.ML	Tue Mar 08 16:02:52 2005 +0100
+++ b/src/Pure/Isar/toplevel.ML	Wed Mar 09 18:44:52 2005 +0100
@@ -425,6 +425,11 @@
      then raised_msg "THEORY" (cat_lines ("" :: msg ::
        map (Pretty.string_of o Display.pretty_theory) thys))
      else msg
+  | msg_on_debug (AST (msg, asts)) =
+     if !debug
+     then raised_msg "AST" (cat_lines ("" :: msg ::
+       map (Pretty.string_of o Syntax.pretty_ast) asts))
+     else msg
   | msg_on_debug (TERM (msg, ts)) =
      (case (!debug, Context.get_context ()) of
          (true, SOME thy) =>
@@ -445,21 +450,6 @@
            end
        | _ => raised_msg "TYPE" msg)
 
-(*
-	 (writeln ("Exception THM " ^ string_of_int i ^ " raised:\n" ^ msg);
-	  seq print_thm thms)
-   | THEORY (msg,thys) =>
-	 (writeln ("Exception THEORY raised:\n" ^ msg);
-	  seq (Pretty.writeln o Display.pretty_theory) thys)
-   | TERM (msg,ts) =>
-	 (writeln ("Exception TERM raised:\n" ^ msg);
-	  seq (writeln o Sign.string_of_term sign) ts)
-   | TYPE (msg,Ts,ts) =>
-	 (writeln ("Exception TYPE raised:\n" ^ msg);
-	  seq (writeln o Sign.string_of_typ sign) Ts;
-	  seq (writeln o Sign.string_of_term sign) ts)
-*)
-
 fun exn_message TERMINATE = "Exit."
   | exn_message RESTART = "Restart."
   | exn_message (EXCURSION_FAIL (exn, msg)) = cat_lines [exn_message exn, msg]
@@ -475,7 +465,7 @@
   | exn_message (Attrib.ATTRIB_FAIL info) = fail_message "attribute" info
   | exn_message (Method.METHOD_FAIL info) = fail_message "method" info
   | exn_message (Antiquote.ANTIQUOTE_FAIL info) = fail_message "antiquotation" info
-  | exn_message (Syntax.AST (msg, _)) = raised_msg "AST" msg
+  | exn_message (Syntax.AST (msg, asts)) = msg_on_debug (AST (msg, asts))
   | exn_message (TYPE (msg, Ts, ts)) = msg_on_debug (TYPE (msg, Ts, ts))
   | exn_message (TERM (msg, ts)) = msg_on_debug (TERM (msg, ts))
   | exn_message (THM (msg, i, thms)) = msg_on_debug (THM (msg, i, thms))
--- a/src/Pure/ROOT.ML	Tue Mar 08 16:02:52 2005 +0100
+++ b/src/Pure/ROOT.ML	Wed Mar 09 18:44:52 2005 +0100
@@ -11,6 +11,7 @@
 
 
 print_depth 10;
+error_depth 30;
 
 (*fake hiding of private structures*)
 structure Hidden = struct end;
--- a/src/Pure/axclass.ML	Tue Mar 08 16:02:52 2005 +0100
+++ b/src/Pure/axclass.ML	Wed Mar 09 18:44:52 2005 +0100
@@ -407,7 +407,7 @@
 
 fun inst_proof mk_prop add_thms inst int theory =
   theory
-  |> IsarThy.multi_theorem_i Drule.internalK
+  |> IsarThy.multi_theorem_i Drule.internalK I
     ("", [fn (thy, th) => (add_thms [th] thy, th)]) []
     (map (fn t => (("", []), [(t, ([], []))])) (mk_prop (Theory.sign_of theory) inst)) int;
 
--- a/src/Pure/logic.ML	Tue Mar 08 16:02:52 2005 +0100
+++ b/src/Pure/logic.ML	Wed Mar 09 18:44:52 2005 +0100
@@ -336,6 +336,7 @@
 
 (*Inverse of varify.  Converts axioms back to their original form.*)
 fun unvarify (Const(a,T))    = Const(a, Type.unvarifyT T)
+  | unvarify (Free(a,T))     = Free(a, Type.unvarifyT T)  (* CB: added *)
   | unvarify (Var((a,0), T)) = Free(a, Type.unvarifyT T)
   | unvarify (Var(ixn,T))    = Var(ixn, Type.unvarifyT T)  (*non-0 index!*)
   | unvarify (Abs (a,T,body)) = Abs (a, Type.unvarifyT T, unvarify body)