X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambda_delta%2FGround_2%2Farith.ma;h=39d28c959182d17c31fa4a851751928c3fc12b8d;hb=b5db76fe31ab35bae0257cb6684c511bcc531e45;hp=187e0f49bcfc21dfe6201094ae9030c3785b84c4;hpb=aec661d51ffa04b4248cdfece772b58780737e3f;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 187e0f49b..39d28c959 100644 --- a/matita/matita/contribs/lambda_delta/Ground_2/arith.ma +++ b/matita/matita/contribs/lambda_delta/Ground_2/arith.ma @@ -17,179 +17,60 @@ include "Ground_2/star.ma". (* ARITHMETICAL PROPERTIES **************************************************) -axiom eq_nat_dec: ∀n1,n2:nat. Decidable (n1 = n2). - -axiom lt_dec: ∀n1,n2. Decidable (n1 < n2). - -lemma plus_S_eq_O_false: ∀n,m. n + S m = 0 → False. -#n #m (commutative_plus p) (le_O_to_eq_O … H) -H // -| #b #IHb #c elim c -c // - #c #_ #a #Hcb - lapply (le_S_S_to_le … Hcb) -Hcb #Hcb - minus_plus @eq_f2 /2/ +#a #b #c1 #H >minus_minus_comm >minus_le_minus_minus_comm // qed. lemma arith_b2: ∀a,b,c1,c2. c1 + c2 ≤ b → a - c1 - c2 - (b - c1 - c2) = a - b. -#a #b #c1 #c2 #H >minus_plus >minus_plus >minus_plus /2/ +#a #b #c1 #c2 #H >minus_plus >minus_plus >minus_plus /2 width=1/ qed. -lemma arith_c1: ∀a,b,c1. a + c1 - (b + c1) = a - b. -// qed. - lemma arith_c1x: ∀x,a,b,c1. x + c1 + a - (b + c1) = x + a - b. -#x #a #b #c1 >plus_plus_comm_23 // -qed. - -lemma arith_d1: ∀a,b,c1. c1 ≤ b → a + c1 + (b - c1) = a + b. -/2/ qed. - -lemma arith_e2: ∀a,c1,c2. a ≤ c1 → c1 + c2 - (c1 - a + c2) = a. -/3/ qed. - -lemma arith_f1: ∀a,b,c1. a + b ≤ c1 → c1 - (c1 - a - b) = a + b. -#a #b #c1 #H >minus_plus plus_minus // /2 width=1/ qed. -lemma arith3: ∀a1,a2,b,c1. a1 + a2 ≤ b → a1 + c1 + a2 ≤ b + c1. -/2/ qed. +(* inversion & forward lemmas ***********************************************) -lemma arith4: ∀h,d,e1,e2. d ≤ e1 + e2 → d + h ≤ e1 + h + e2. -/2/ qed. +axiom eq_nat_dec: ∀n1,n2:nat. Decidable (n1 = n2). -lemma arith5: ∀a,b1,b2,c1. c1 ≤ b1 → c1 ≤ a → a < b1 + b2 → a - c1 < b1 - c1 + b2. -#a #b1 #b2 #c1 #H1 #H2 #H3 -minus_plus >minus_plus @monotonic_le_plus_r @monotonic_le_minus_l // -qed. +lemma lt_zero_false: ∀n. n < 0 → False. +#n #H elim (lt_to_not_le … H) -H /2 width=1/ +qed-. + +lemma false_lt_to_le: ∀x,y. (x < y → False) → y ≤ x. +#x #y #H elim (decidable_lt x y) /2 width=1/ +#Hxy elim (H Hxy) +qed-. + +(* +lemma pippo: ∀x,y,z. x < z → y < z - x → x + y < z. +/3 width=2/ + +lemma le_or_ge: ∀m,n. m ≤ n ∨ n ≤ m. +#m #n elim (lt_or_ge m n) /2 width=1/ /3 width=2/ +qed-. +*)