src/HOL/Integ/IntDef.thy
changeset 5540 0f16c3b66ab4
parent 5508 691c70898518
child 5562 02261e6880d1
--- a/src/HOL/Integ/IntDef.thy	Wed Sep 23 10:17:11 1998 +0200
+++ b/src/HOL/Integ/IntDef.thy	Wed Sep 23 10:25:37 1998 +0200
@@ -23,11 +23,11 @@
 
 constdefs
 
-  znat        :: nat => int                                  ("$# _" [80] 80)
+  nat        :: nat => int                                  ("$# _" [80] 80)
   "$# m == Abs_Integ(intrel ^^ {(m,0)})"
 
-  znegative   :: int => bool
-  "znegative(Z) == EX x y. x<y & (x,y::nat):Rep_Integ(Z)"
+  neg   :: int => bool
+  "neg(Z) == EX x y. x<y & (x,y::nat):Rep_Integ(Z)"
 
   (*For simplifying equalities*)
   iszero :: int => bool
@@ -41,7 +41,7 @@
 
   zdiff_def "Z1 - Z2 == Z1 + -(Z2::int)"
 
-  zless_def "Z1<Z2 == znegative(Z1 - Z2)"
+  zless_def "Z1<Z2 == neg(Z1 - Z2)"
 
   zle_def   "Z1 <= (Z2::int) == ~(Z2 < Z1)"