1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 include "ground/notation/functions/zero_0.ma".
16 include "ground/arith/pnat.ma".
18 (* NON-NEGATIVE INTEGERS ****************************************************)
21 inductive nat: Type[0] ≝
29 "zero (non-negative integers)"
32 (* Basic inversions *********************************************************)
35 lemma eq_inv_ninj_bi: injective … ninj.
39 (* Basic constructions ******************************************************)
42 lemma eq_nat_dec (n1,n2:nat): Decidable (n1 = n2).
43 * [| #p1 ] * [2,4: #p2 ]
44 [1,4: @or_intror #H destruct
45 | elim (eq_pnat_dec p1 p2)
46 /4 width=1 by eq_inv_ninj_bi, or_intror, or_introl/
47 | /2 width=1 by or_introl/