use the noted theorems in 'BNF_Comp'
authorblanchet
Thu, 24 Jul 2014 00:24:00 +0200
changeset 57632 231e90cf2892
parent 57631 959caab43a3d
child 57633 4ff8c090d580
use the noted theorems in 'BNF_Comp'
src/HOL/Tools/BNF/bnf_comp.ML
src/HOL/Tools/BNF/bnf_def.ML
--- a/src/HOL/Tools/BNF/bnf_comp.ML	Thu Jul 24 00:24:00 2014 +0200
+++ b/src/HOL/Tools/BNF/bnf_comp.ML	Thu Jul 24 00:24:00 2014 +0200
@@ -878,9 +878,10 @@
 
     val notes = (map_b, map_def) :: (rel_b, rel_def) :: (set_bs ~~ set_defs)
       |> map (fn (b, def) => ((b, []), [([def], [])]))
-    val lthy'' = lthy' |> Local_Theory.notes notes |> snd
+
+    val (noted, lthy'') = lthy' |> Local_Theory.notes notes
   in
-    ((bnf'', (all_deads, absT_info)), lthy'')
+    ((morph_bnf (substitute_noted_thm noted) bnf'', (all_deads, absT_info)), lthy'')
   end;
 
 exception BAD_DEAD of typ * typ;
--- a/src/HOL/Tools/BNF/bnf_def.ML	Thu Jul 24 00:24:00 2014 +0200
+++ b/src/HOL/Tools/BNF/bnf_def.ML	Thu Jul 24 00:24:00 2014 +0200
@@ -752,7 +752,6 @@
 
     val phi = Proof_Context.export_morphism lthy_old lthy;
 
-
     val bnf_map_def = Morphism.thm phi raw_map_def;
     val bnf_set_defs = map (Morphism.thm phi) raw_set_defs;
     val bnf_bd_def = Morphism.thm phi raw_bd_def;