X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground%2Farith%2Farith_2a.ma;fp=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground%2Farith%2Farith_2a.ma;h=b5ff26116ef36e2b41f8e3c4abd9a51fc374f45c;hb=888840f6b3a71d3d686b53b702d362ab90ab0038;hp=0000000000000000000000000000000000000000;hpb=19b0a814861157ba05f23877d5cd94059f52c2e8;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/ground/arith/arith_2a.ma b/matita/matita/contribs/lambdadelta/ground/arith/arith_2a.ma new file mode 100644 index 000000000..b5ff26116 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/ground/arith/arith_2a.ma @@ -0,0 +1,93 @@ +(**************************************************************************) +(* ___ *) +(* ||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/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 + 𝟏 + 𝟏. +// qed. + +lemma arith_b1: ∀a,b,c1. c1 ≤ b → a - c1 - (b - c1) = a - b. +#a #b #c1 #H >nminus_comm (nminus_plus_assoc ? c1 c2) >(nminus_plus_assoc ? c1 c2) +/2 width=1 by arith_b1/ +qed-. + +lemma arith_c1x: ∀x,a,b,c1. x + c1 + a - (b + c1) = x + a - b. +#x #a #b #c1 +nminus_plus_comm_23 +/2 width=1 by arith_b1/ +qed-. + +lemma arith_i: ∀x,y,z. y < x → x+z-y-(𝟏) = x-y-(𝟏)+z. +/2 width=1 by nminus_plus_comm_23/ qed-. + +(* Constructions ************************************************************) + +fact le_repl_sn_conf_aux: ∀x,y,z:nat. x ≤ z → x = y → y ≤ z. +// qed-. + +fact le_repl_sn_trans_aux: ∀x,y,z:nat. x ≤ z → y = x → y ≤ z. +// qed-. + +lemma monotonic_le_minus_l2: ∀x1,x2,y,z. x1 ≤ x2 → x1 - y - z ≤ x2 - y - z. +/3 width=1 by nle_minus_bi_dx/ qed. + +lemma arith_j: ∀x,y,z. x-y-(𝟏) ≤ x-(y-z)-𝟏. +/3 width=1 by nle_minus_bi_dx, nle_minus_bi_sn/ qed. + +lemma arith_k_sn: ∀z,x,y,n. z < x → x+n ≤ y → x-z-(𝟏)+n ≤ y-z-𝟏. +#z #x #y #n #Hzx #Hxny +>nminus_plus_comm_23 [2: /2 width=1 by nle_minus_bi_sn/ ] +>nminus_plus_comm_23 [2: /2 width=1 by nlt_des_le/ ] +/2 width=1 by monotonic_le_minus_l2/ +qed. + +lemma arith_k_dx: ∀z,x,y,n. z < x → y ≤ x+n → y-z-(𝟏) ≤ x-z-(𝟏)+n. +#z #x #y #n #Hzx #Hyxn +>nminus_plus_comm_23 [2: /2 width=1 by nle_minus_bi_sn/ ] +>nminus_plus_comm_23 [2: /2 width=1 by nlt_des_le/ ] +/2 width=1 by monotonic_le_minus_l2/ +qed. + +(* Inversions ***************************************************************) + +lemma lt_plus_SO_to_le: ∀x,y. x < y + (𝟏) → x ≤ y. +/2 width=1 by nle_inv_succ_bi/ qed-. + +(* Iterators ****************************************************************) + +lemma iter_SO: ∀B:Type[0]. ∀f:B→B. ∀b,l. f^(l+𝟏) b = f (f^l b). +#B #f #b #l +