src/HOLCF/Pcpo.thy
changeset 33523 96730ad673be
parent 31076 99fe356cbbc2
child 35492 5d200f2d7a4f
--- a/src/HOLCF/Pcpo.thy	Sun Nov 08 18:43:42 2009 +0100
+++ b/src/HOLCF/Pcpo.thy	Sun Nov 08 19:15:37 2009 +0100
@@ -196,11 +196,11 @@
 text {* Simproc to rewrite @{term "\<bottom> = x"} to @{term "x = \<bottom>"}. *}
 
 setup {*
-  ReorientProc.add
+  Reorient_Proc.add
     (fn Const(@{const_name UU}, _) => true | _ => false)
 *}
 
-simproc_setup reorient_bottom ("\<bottom> = x") = ReorientProc.proc
+simproc_setup reorient_bottom ("\<bottom> = x") = Reorient_Proc.proc
 
 context pcpo
 begin