src/ZF/OrdQuant.thy
changeset 56250 2c9f841f36b8
parent 54998 8601434fa334
child 58860 fee7cfa69c50
--- a/src/ZF/OrdQuant.thy	Sat Mar 22 15:58:27 2014 +0100
+++ b/src/ZF/OrdQuant.thy	Sat Mar 22 16:50:52 2014 +0100
@@ -357,9 +357,8 @@
 ML
 {*
 val Ord_atomize =
-    atomize ([("OrdQuant.oall", [@{thm ospec}]),("OrdQuant.rall", [@{thm rspec}])]@
-                 ZF_conn_pairs,
-             ZF_mem_pairs);
+  atomize ([(@{const_name oall}, @{thms ospec}), (@{const_name rall}, @{thms rspec})] @
+    ZF_conn_pairs, ZF_mem_pairs);
 *}
 declaration {* fn _ =>
   Simplifier.map_ss (Simplifier.set_mksimps (K (map mk_eq o Ord_atomize o gen_all)))