(* *)
(**************************************************************************)
-include "ground/notation/functions/two_0.ma".
+include "ground/arith/pnat_two.ma".
include "ground/arith/nat_le_minus_plus.ma".
include "ground/arith/nat_lt.ma".
(* ARITHMETICAL PROPERTIES FOR λδ-2A ****************************************)
-interpretation
- "zero (non-negative integers)"
- 'Two = (ninj (psucc punit)).
-
(* Equalities ***************************************************************)
lemma plus_n_2: ∀n. (n + 𝟐) = n + 𝟏 + 𝟏.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "ground/notation/functions/two_0.ma".
+include "ground/arith/pnat.ma".
+
+(* POSITIVE INTEGERS ********************************************************)
+
+interpretation
+ "two (positive integers)"
+ 'Two = (psucc punit).