From: Ferruccio Guidi Date: Sat, 22 Feb 2014 11:50:51 +0000 (+0000) Subject: addition of unused material :) X-Git-Tag: make_still_working~971 X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=commitdiff_plain;h=e500cfb0c28718d44972a119f55f152018d18e62;p=helm.git addition of unused material :) --- diff --git a/matita/matita/contribs/lambdadelta/ground_2/etc/lib/arith.etc b/matita/matita/contribs/lambdadelta/ground_2/etc/lib/arith.etc new file mode 100644 index 000000000..656a44255 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/ground_2/etc/lib/arith.etc @@ -0,0 +1,18 @@ + +lemma zero_or_gt: ∀n. n = 0 ∨ 0 < n. +#n elim (lt_or_eq_or_gt 0 n) /2/ +#H elim (lt_zero_false … H) +qed. + +lemma pippo: ∀x,y,z. x < z → y < z - x → x + y < z. +/3 width=2/ + +lemma arith_b1x: ∀e,x1,x2,y. y ≤ x2 → x2 ≤ x1 → + e + (x1 - y) - (x2 - y) = e + (x1 - x2). +#e #x1 #x2 #y #H1 #H2 +<(arith_b1 … H1) >le_plus_minus // /2 width=1/ +qed-. + +lemma arith1: ∀x,y,z,w. z < y → x + (y + w) - z = x + y - z + w. +#x #y #z #w #H H -H // +qed-. + +lemma discr_YO_YS: ∀n. ynat_of_nat 0 = ⫯n → ⊥. (**) (* explicit coercion *) +/2 width=2 by discr_YS_YO/ qed-. diff --git a/matita/matita/contribs/lambdadelta/ground_2/etc/ynat_old/ynat_le.etc b/matita/matita/contribs/lambdadelta/ground_2/etc/ynat_old/ynat_le.etc new file mode 100644 index 000000000..bde84cf79 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/ground_2/etc/ynat_old/ynat_le.etc @@ -0,0 +1,84 @@ +(**************************************************************************) +(* ___ *) +(* ||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_2/star.ma". +include "ground_2/ynat/ynat_iszero.ma". +include "ground_2/ynat/ynat_pred.ma". + +(* INFINITARY NATURAL NUMBERS ***********************************************) + +(* order relation *) +coinductive yle: relation ynat ≝ +| yle_O: ∀n. yle 0 n +| yle_S: ∀m,n. yle m n → yle (⫯m) (⫯n) +. + +interpretation "natural 'less or equal to'" 'leq x y = (yle x y). + +(* Inversion lemmas *********************************************************) + +fact yle_inv_O2_aux: ∀m,x. m ≤ x → x = 0 → m = 0. +#m #x * -m -x // +#m #x #_ #H elim (discr_YS_YO … H) (**) (* destructing lemma needed *) +qed-. + +lemma yle_inv_O2: ∀m. m ≤ 0 → m = 0. +/2 width =3 by yle_inv_O2_aux/ qed-. + +fact yle_inv_S1_aux: ∀x,y. x ≤ y → ∀m. x = ⫯m → ∃∃n. m ≤ n & y = ⫯n. +#x #y * -x -y +[ #y #m #H elim (discr_YO_YS … H) (**) (* destructing lemma needed *) +| #x #y #Hxy #m #H destruct /2 width=3 by ex2_intro/ +] +qed-. + +lemma yle_inv_S1: ∀m,y. ⫯m ≤ y → ∃∃n. m ≤ n & y = ⫯n. +/2 width=3 by yle_inv_S1_aux/ qed-. + +lemma yle_inv_S: ∀m,n. ⫯m ≤ ⫯n → m ≤ n. +#m #n #H elim (yle_inv_S1 … H) -H +#x #Hx #H destruct // +qed-. + +(* Properties ***************************************************************) + +let corec yle_refl: reflexive … yle ≝ ?. +* [ @yle_O | #x @yle_S // ] +qed. + +let corec yle_Y: ∀m. m ≤ ∞ ≝ ?. +* [ @yle_O | #m