src/Pure/Pure.thy
changeset 78798 200daaab2578
parent 78797 fc598652fb8a
child 78803 577835250124
--- a/src/Pure/Pure.thy	Wed Oct 18 22:09:25 2023 +0200
+++ b/src/Pure/Pure.thy	Thu Oct 19 11:30:16 2023 +0200
@@ -336,7 +336,7 @@
     (Simplifier.simproc_setup_parser >> (fn arg =>
       Context.proof_map
         (ML_Context.expression (Input.pos_of (#proc arg))
-          (ML_Lex.read "Simplifier.simproc_setup_local " @ Simplifier.simproc_setup_ml arg))));
+          (ML_Lex.read "Simplifier.simproc_setup " @ Simplifier.simproc_setup_ml arg))));
 
 in end\<close>