src/HOL/Tools/BNF/bnf_gfp.ML
changeset 56348 2653b2fdad06
parent 56346 42533f8f4729
child 56516 a13c2ccc160b
--- a/src/HOL/Tools/BNF/bnf_gfp.ML	Tue Apr 01 10:51:29 2014 +0200
+++ b/src/HOL/Tools/BNF/bnf_gfp.ML	Tue Apr 01 10:51:29 2014 +0200
@@ -2433,10 +2433,10 @@
 
         val (Jbnfs, lthy) =
           fold_map7 (fn b => fn tacs => fn map_b => fn rel_b => fn set_bs => fn wit_thms =>
-              fn consts => fn lthy =>
+              fn consts =>
             bnf_def Hardly_Inline (user_policy Note_Some) false I tacs (wit_tac wit_thms)
-              (SOME deads) map_b rel_b set_bs consts lthy
-            |> (fn (bnf, lthy) => (bnf, register_bnf (Local_Theory.full_name lthy b) bnf lthy)))
+              (SOME deads) map_b rel_b set_bs consts
+            #> (fn (bnf, lthy) => (bnf, register_bnf (Local_Theory.full_name lthy b) bnf lthy)))
           bs tacss map_bs rel_bs set_bss wit_thmss
           ((((((bs ~~ Ts) ~~ Jmaps) ~~ Jsetss_by_bnf) ~~ Jbds) ~~ all_witss) ~~ map SOME Jrels)
           lthy;