X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambda_delta%2Fground_2%2Farith.ma;h=8afc445b8f4343d68494ce3e774ebc8beb196c9c;hb=5613a25cee29ef32a597cb4b44e8f2f4d71c4df0;hp=c51873baa09bb9d8f0c46e2f502e378cb05af921;hpb=b074ebf6441993694c6e39e4eaeeb58a3186f479;p=helm.git diff --git a/matita/matita/contribs/lambda_delta/ground_2/arith.ma b/matita/matita/contribs/lambda_delta/ground_2/arith.ma index c51873baa..8afc445b8 100644 --- a/matita/matita/contribs/lambda_delta/ground_2/arith.ma +++ b/matita/matita/contribs/lambda_delta/ground_2/arith.ma @@ -44,7 +44,7 @@ lemma arith_h1: ∀a1,a2,b,c1. c1 ≤ a1 → c1 ≤ b → #a1 #a2 #b #c1 #H1 #H2 >plus_minus // /2 width=1/ qed. -(* inversion & forward lemmas ***********************************************) +(* Inversion & forward lemmas ***********************************************) axiom eq_nat_dec: ∀n1,n2:nat. Decidable (n1 = n2). @@ -69,7 +69,18 @@ lemma false_lt_to_le: ∀x,y. (x < y → ⊥) → y ≤ x. #Hxy elim (H Hxy) qed-. -(* iterators ****************************************************************) +lemma le_plus_xySz_x_false: ∀y,z,x. x + y + S z ≤ x → ⊥. +#y #z #x elim x -x +[ #H lapply (le_n_O_to_eq … H) -H +