From 48cd63fc7f4415925582eae224a36a9c1bb3cc06 Mon Sep 17 00:00:00 2001 From: Ferruccio Guidi Date: Sat, 12 Nov 2022 23:21:34 +0100 Subject: [PATCH] update in ground + positive integer two --- .../lambdadelta/ground/arith/arith_2a.ma | 6 +---- .../lambdadelta/ground/arith/pnat_two.ma | 22 +++++++++++++++++++ 2 files changed, 23 insertions(+), 5 deletions(-) create mode 100644 matita/matita/contribs/lambdadelta/ground/arith/pnat_two.ma diff --git a/matita/matita/contribs/lambdadelta/ground/arith/arith_2a.ma b/matita/matita/contribs/lambdadelta/ground/arith/arith_2a.ma index 92224c4c9..a136174bc 100644 --- a/matita/matita/contribs/lambdadelta/ground/arith/arith_2a.ma +++ b/matita/matita/contribs/lambdadelta/ground/arith/arith_2a.ma @@ -12,16 +12,12 @@ (* *) (**************************************************************************) -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 + 𝟏 + 𝟏. diff --git a/matita/matita/contribs/lambdadelta/ground/arith/pnat_two.ma b/matita/matita/contribs/lambdadelta/ground/arith/pnat_two.ma new file mode 100644 index 000000000..d1be445c1 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/ground/arith/pnat_two.ma @@ -0,0 +1,22 @@ +(**************************************************************************) +(* ___ *) +(* ||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). -- 2.39.2