Made global_note_qualified public.
--- a/src/Pure/Isar/locale.ML Wed Dec 03 14:02:24 2008 +0000
+++ b/src/Pure/Isar/locale.ML Wed Dec 03 15:26:46 2008 +0100
@@ -91,6 +91,9 @@
val intro_locales_tac: bool -> Proof.context -> thm list -> tactic
(* Storing results *)
+ val global_note_qualified: string ->
+ ((Name.binding * attribute list) * (thm list * attribute list) list) list ->
+ theory -> (string * thm list) list * theory
val local_note_qualified: string ->
((Name.binding * attribute list) * (thm list * attribute list) list) list ->
Proof.context -> (string * thm list) list * Proof.context
@@ -2409,7 +2412,6 @@
ListPair.map (fn ((_, props), thms) => map2 Element.make_witness props thms) (propps, thmss);
fun gen_interpretation prep_registration after_qed prfx raw_expr raw_insts thy =
- (* prfx = (flag indicating full qualification, name prefix) *)
let
val (propss, activate, morphs) = prep_registration thy prfx raw_expr raw_insts;
fun after_qed' results =