refined activate_notes: simultaneous transformation before activation;
authorwenzelm
Mon, 12 Mar 2012 20:44:10 +0100
changeset 46880 af8e516953ce
parent 46879 a8b1236e0837
child 46881 b81f1de9f57e
refined activate_notes: simultaneous transformation before activation; actually export all_registrations_of;
src/Pure/Isar/locale.ML
--- a/src/Pure/Isar/locale.ML	Mon Mar 12 19:09:38 2012 +0100
+++ b/src/Pure/Isar/locale.ML	Mon Mar 12 20:44:10 2012 +0100
@@ -72,6 +72,7 @@
   val amend_registration: string * morphism -> morphism * bool ->
     morphism -> Context.generic -> Context.generic
   val registrations_of: Context.generic -> string -> (string * morphism) list
+  val all_registrations_of: Context.generic -> (string * morphism) list
   val add_dependency: string -> string * morphism -> (morphism * bool) option ->
     morphism -> theory -> theory
 
@@ -205,12 +206,10 @@
 fun mixins_of thy name serial = the_locale thy name |>
   #mixins |> lookup_mixins serial;
 
-(* unused *)
+(* FIXME unused *)
 fun identity_on thy name morph =
   let val mk_instance = instance_of thy name
-  in
-    forall2 (curry Term.aconv_untyped) (mk_instance Morphism.identity) (mk_instance morph)
-  end;
+  in ListPair.all Term.aconv_untyped (mk_instance Morphism.identity, mk_instance morph) end;
 
 (* Print instance and qualifiers *)
 
@@ -382,7 +381,7 @@
 fun registrations_of context name =
   get_registrations context (filter (curry (op =) name o fst o fst));
 
-fun all_registrations context = get_registrations context I;
+fun all_registrations_of context = get_registrations context I;
 
 
 (*** Activate context elements of locale ***)
@@ -401,17 +400,15 @@
 fun activate_notes activ_elem transfer context export' (name, morph) input =
   let
     val thy = Context.theory_of context;
-    val {notes, ...} = the_locale thy name;
-    fun activate ((kind, facts), _) input =
-      let
-        val mixin =
-          (case export' of
-            NONE => Morphism.identity
-          | SOME export => collect_mixins context (name, morph $> export) $> export);
-        val facts' = facts |> Element.transform_facts (transfer input $> morph $> mixin);
-      in activ_elem (Notes (kind, facts')) input end;
+    val mixin =
+      (case export' of
+        NONE => Morphism.identity
+      | SOME export => collect_mixins context (name, morph $> export) $> export);
+    val morph' = transfer input $> morph $> mixin;
+    val notes' = map (Element.transform_ctxt morph' o Notes o #1) (#notes (the_locale thy name));
   in
-    fold_rev activate notes input
+    fold_rev (fn elem => fn res => activ_elem (Element.transform_ctxt (transfer res) elem) res)
+      notes' input
   end;
 
 fun activate_all name thy activ_elem transfer (marked, input) =