From 8899a3f240f62633f4df58b2ee358fa285a82d1d Mon Sep 17 00:00:00 2001 From: Andrea Asperti Date: Tue, 26 Jan 2010 09:20:14 +0000 Subject: [PATCH] le_arith --- .../matita/nlibrary/arithmetics/nat.ma | 442 +++++++++--------- 1 file changed, 233 insertions(+), 209 deletions(-) diff --git a/helm/software/matita/nlibrary/arithmetics/nat.ma b/helm/software/matita/nlibrary/arithmetics/nat.ma index 8efe32043..b6683b267 100644 --- a/helm/software/matita/nlibrary/arithmetics/nat.ma +++ b/helm/software/matita/nlibrary/arithmetics/nat.ma @@ -262,41 +262,90 @@ ntheorem le_pred_n : ∀n:nat. pred n ≤ n. ntheorem monotonic_pred: monotonic ? le pred. #n; #m; #lenm; nelim lenm; /2/; nqed. -ntheorem le_S_S_to_le : ∀n,m:nat. S n ≤ S m → n ≤ m. +ntheorem le_S_S_to_le: ∀n,m:nat. S n ≤ S m → n ≤ m. /2/; nqed. -(* -nchange with (pred (S n) ≤ pred (S m)); -nelim leSS; apply le_n.apply (trans_le ? (pred n1)).assumption. -apply le_pred_n. -qed. +ntheorem lt_S_S_to_lt: ∀n,m. S n < S m \to n < m. +/2/; nqed. -theorem lt_S_S_to_lt: \forall n,m. - S n < S m \to n < m. -intros. apply le_S_S_to_le. assumption. -qed. +ntheorem lt_to_lt_S_S: ∀n,m. n < m → S n < S m. +/2/; nqed. -theorem lt_to_lt_S_S: ∀n,m. n < m → S n < S m. -intros; -unfold lt in H; -apply (le_S_S ? ? H). -qed. +ntheorem lt_to_not_zero : ∀n,m:nat. n < m → not_zero m. +#n; #m; #Hlt; nelim Hlt;//; nqed. -theorem leS_to_not_zero : \forall n,m:nat. S n \leq m \to not_zero m. -intros.elim H.exact I.exact I. -qed. +(* lt vs. le *) +ntheorem not_le_Sn_O: ∀ n:nat. S n ≰ O. +#n; #Hlen0; napply (lt_to_not_zero ?? Hlen0); nqed. -(* not le *) -theorem not_le_Sn_O: \forall n:nat. S n \nleq O. -intros.unfold Not.simplify.intros.apply (leS_to_not_zero ? ? H). -qed. +ntheorem not_le_to_not_le_S_S: ∀ n,m:nat. n ≰ m → S n ≰ S m. +/3/; nqed. -theorem not_le_Sn_n: \forall n:nat. S n \nleq n. -intros.elim n.apply not_le_Sn_O.unfold Not.simplify.intros.cut (S n1 \leq n1). -apply H.assumption. -apply le_S_S_to_le.assumption. -qed. +ntheorem not_le_S_S_to_not_le: ∀ n,m:nat. S n ≰ S m → n ≰ m. +/3/; nqed. + +ntheorem decidable_le: ∀n,m. decidable (n≤m). +napply nat_elim2; #n; /2/; +#m; #dec; ncases dec;/3/; nqed. + +ntheorem decidable_lt: ∀n,m. decidable (n < m). +#n; #m; napply decidable_le ; nqed. + +ntheorem not_le_Sn_n: ∀n:nat. S n ≰ n. +#n; nelim n; /2/; nqed. + +ntheorem lt_S_to_le: ∀n,m:nat. n < S m → n ≤ m. +/2/; nqed. + +ntheorem not_le_to_lt: ∀n,m. n ≰ m → m < n. +napply nat_elim2; #n; + ##[#abs; napply False_ind;/2/; + ##|/2/; + ##|#m;#Hind;#HnotleSS; napply lt_to_lt_S_S;/3/; + ##] +nqed. + +ntheorem lt_to_not_le: ∀n,m. n < m → m ≰ n. +#n; #m; #Hltnm; nelim Hltnm;/3/; nqed. + +ntheorem not_lt_to_le: ∀n,m:nat. n ≮ m → m ≤ n. +#n; #m; #Hnlt; napply lt_S_to_le; +(* something strange here: /2/ fails *) +napply not_le_to_lt; napply Hnlt; nqed. + +ntheorem le_to_not_lt: ∀n,m:nat. n ≤ m → m ≮ n. +/2/; nqed. + +(* lt and le trans *) + +ntheorem lt_to_le_to_lt: ∀n,m,p:nat. n < m → m ≤ p → n < p. +#n; #m; #p; #H; #H1; nelim H1; /2/; nqed. + +ntheorem le_to_lt_to_lt: ∀n,m,p:nat. n ≤ m → m < p → n < p. +#n; #m; #p; #H; nelim H; /3/; nqed. + +ntheorem lt_S_to_lt: ∀n,m. S n < m → n < m. +/2/; nqed. + +ntheorem ltn_to_ltO: ∀n,m:nat. n < m → O < m. +/2/; nqed. + +(* +theorem lt_SO_n_to_lt_O_pred_n: \forall n:nat. +(S O) \lt n \to O \lt (pred n). +intros. +apply (ltn_to_ltO (pred (S O)) (pred n) ?). + apply (lt_pred (S O) n); + [ apply (lt_O_S O) + | assumption + ] +qed. *) + +ntheorem lt_O_n_elim: ∀n:nat. O < n → + ∀P:nat → Prop.(∀m:nat.P (S m)) → P n. +#n; nelim n; //; #abs; napply False_ind; /2/; nqed. +(* theorem lt_pred: \forall n,m. O < n \to n < m \to pred n < pred m. apply nat_elim2 @@ -325,48 +374,18 @@ elim n; ]. qed. -theorem le_to_le_pred: - ∀n,m. n ≤ m → pred n ≤ pred m. -intros 2; -elim n; -[ simplify; - apply le_O_n -| simplify; - elim m in H1 ⊢ %; - [ elim (not_le_Sn_O ? H1) - | simplify; - apply le_S_S_to_le; - assumption - ] -]. -qed. +*) (* le to lt or eq *) -theorem le_to_or_lt_eq : \forall n,m:nat. -n \leq m \to n < m \lor n = m. -intros.elim H. -right.reflexivity. -left.unfold lt.apply le_S_S.assumption. -qed. - -theorem Not_lt_n_n: ∀n. n ≮ n. -intro; -unfold Not; -intro; -unfold lt in H; -apply (not_le_Sn_n ? H). -qed. +ntheorem le_to_or_lt_eq: ∀n,m:nat. n ≤ m → n < m ∨ n = m. +#n; #m; #lenm; nelim lenm; /3/; nqed. (* not eq *) -theorem lt_to_not_eq : \forall n,m:nat. n plus_n_O. +simplify.rewrite > plus_n_Sm. +apply le_plus + [assumption + |rewrite < plus_n_O. + assumption + ] +qed. +(*0 and times *) +theorem O_lt_const_to_le_times_const: \forall a,c:nat. +O \lt c \to a \le a*c. +intros. +rewrite > (times_n_SO a) in \vdash (? % ?). +apply le_times +[ apply le_n +| assumption +] +qed. *) \ No newline at end of file -- 2.39.2