src/HOL/HOLCF/Pcpo.thy
changeset 78099 4d9349989d94
parent 74305 28a582aa25dd
--- a/src/HOL/HOLCF/Pcpo.thy	Tue May 23 20:11:15 2023 +0200
+++ b/src/HOL/HOLCF/Pcpo.thy	Tue May 23 21:43:36 2023 +0200
@@ -157,7 +157,7 @@
 
 text \<open>Simproc to rewrite \<^term>\<open>\<bottom> = x\<close> to \<^term>\<open>x = \<bottom>\<close>.\<close>
 setup \<open>Reorient_Proc.add (fn \<^Const_>\<open>bottom _\<close> => true | _ => false)\<close>
-simproc_setup reorient_bottom ("\<bottom> = x") = Reorient_Proc.proc
+simproc_setup reorient_bottom ("\<bottom> = x") = \<open>K Reorient_Proc.proc\<close>
 
 text \<open>useful lemmas about \<^term>\<open>\<bottom>\<close>\<close>