src/HOLCF/Porder.thy
changeset 3842 b55686a7b22c
parent 2640 ee4dfce170a0
child 4721 c8a8482a8124
--- a/src/HOLCF/Porder.thy	Fri Oct 10 18:37:49 1997 +0200
+++ b/src/HOLCF/Porder.thy	Fri Oct 10 19:02:28 1997 +0200
@@ -24,7 +24,7 @@
 
 translations
 
-  "LUB x. t"	== "lub(range(%x.t))"
+  "LUB x. t"	== "lub(range(%x. t))"
 
 syntax (symbols)
 
@@ -33,14 +33,14 @@
 defs
 
 (* class definitions *)
-is_ub           "S  <| x == ! y.y:S --> y<<x"
+is_ub           "S  <| x == ! y. y:S --> y<<x"
 is_lub          "S <<| x == S <| x & (! u. S <| u  --> x << u)"
 
 (* Arbitrary chains are total orders    *)                  
 is_tord         "is_tord S == ! x y. x:S & y:S --> (x<<y | y<<x)"
 
 (* Here we use countable chains and I prefer to code them as functions! *)
-is_chain        "is_chain F == (! i.F(i) << F(Suc(i)))"
+is_chain        "is_chain F == (! i. F(i) << F(Suc(i)))"
 
 (* finite chains, needed for monotony of continouous functions *)
 max_in_chain_def "max_in_chain i C == ! j. i <= j --> C(i) = C(j)"