1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| A.Asperti, C.Sacerdoti Coen, *)
8 (* ||A|| E.Tassi, S.Zacchiroli *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU Lesser General Public License Version 2.1 *)
13 (**************************************************************************)
15 include "arithmetics/nat.ma".
22 interpretation "Integers" 'Z = Z.
24 (* TODO: move the following two to datatypes/compare.ma *)
25 ninductive compare : Type[0] ≝
30 nlet rec nat_compare n m: compare ≝
37 | S q ⇒ nat_compare p q]].
39 ndefinition Zplus : Z → Z → Z ≝
45 | pos n ⇒ (pos (pred ((S m)+(S n))))
47 match nat_compare m n with
48 [ LT ⇒ (neg (pred (n-m)))
50 | GT ⇒ (pos (pred (m-n)))] ]
55 match nat_compare m n with
56 [ LT ⇒ (pos (pred (n-m)))
58 | GT ⇒ (neg (pred (m-n)))]
59 | neg n ⇒ (neg (pred ((S m)+(S n))))] ].
61 interpretation "integer plus" 'plus x y = (Zplus x y).
63 ndefinition Z_of_nat ≝
68 nlemma fails : ∀p. p + Z_of_nat 1 = Z_of_nat 1 + p.