src/HOL/Library/EfficientNat.thy
changeset 23591 d32a85385e17
parent 23438 dd824e86fa8a
child 23711 dc452e8641aa
--- a/src/HOL/Library/EfficientNat.thy	Thu Jul 05 20:01:26 2007 +0200
+++ b/src/HOL/Library/EfficientNat.thy	Thu Jul 05 20:01:28 2007 +0200
@@ -336,7 +336,7 @@
       | find_var (t $ u) = (case find_var t of NONE => find_var u | x => x)
       | find_var _ = NONE;
     fun find_thm th =
-      let val th' = ObjectLogic.atomize_thm th
+      let val th' = Conv.fconv_rule ObjectLogic.atomize th
       in Option.map (pair (th, th')) (find_var (prop_of th')) end
   in
     case get_first find_thm thms of