From a356482acf3bf92eaa34d2b38ee2e81761f58520 Mon Sep 17 00:00:00 2001 From: Andrea Asperti Date: Thu, 21 Jan 2010 17:32:19 +0000 Subject: [PATCH] Esempio --- .../matita/nlibrary/arithmetics/nat.ma | 432 +++++++++++++++++- 1 file changed, 422 insertions(+), 10 deletions(-) diff --git a/helm/software/matita/nlibrary/arithmetics/nat.ma b/helm/software/matita/nlibrary/arithmetics/nat.ma index a1e6fbb23..8efe32043 100644 --- a/helm/software/matita/nlibrary/arithmetics/nat.ma +++ b/helm/software/matita/nlibrary/arithmetics/nat.ma @@ -36,12 +36,10 @@ ncoercion nat_to_pos: ∀n:nat. n ≠0 →pos ≝ mk_pos on *) ndefinition pred ≝ - λn. match n with - [ O ⇒ O - | (S p) ⇒ p]. + λn. match n with [ O ⇒ O | (S p) ⇒ p]. ntheorem pred_Sn : ∀n. n = pred (S n). -//. nqed. +//; nqed. ntheorem injective_S : injective nat nat S. //; nqed. @@ -55,8 +53,7 @@ ntheorem not_eq_S: ∀n,m:nat. n ≠ m → S n ≠ S m. ndefinition not_zero: nat → Prop ≝ λn: nat. match n with - [ O ⇒ False - | (S p) ⇒ True ]. + [ O ⇒ False | (S p) ⇒ True ]. ntheorem not_eq_O_S : ∀n:nat. O ≠ S n. #n; #eqOS; nchange with (not_zero O); nrewrite > eqOS; //. @@ -80,10 +77,10 @@ ntheorem nat_elim2 : #n0; #Rn0m; #m; ncases m;/2/; nqed. ntheorem decidable_eq_nat : \forall n,m:nat.decidable (n=m). -napply nat_elim2;#n; +napply nat_elim2; #n; ##[ ncases n; /2/; ##| /3/; - ##| #m; #Hind; ncases Hind; /3/; (* ??? /2/;#neqnm; /3/; *) + ##| #m; #Hind; ncases Hind; /3/; ##] nqed. @@ -125,7 +122,7 @@ ntheorem symmetric_plus: symmetric ? plus. ntheorem associative_plus : associative nat plus. #n; nelim n; nnormalize; //; nqed. -ntheorem assoc_plus1: ∀a,b,c. b + (a + c) = a + (b + c). +ntheorem assoc_plus1: ∀a,b,c. c + (b + a) = b + c + a. //; nqed. ntheorem injective_plus_r: ∀n:nat.injective nat nat (λm.n+m). @@ -206,5 +203,420 @@ n = (S(S O))*m \lor n = S ((S(S O))*m). nqed. *) -(************************** compare ****************************) +(******************** ordering relations ************************) + +ninductive le (n:nat) : nat → Prop ≝ + | le_n : le n n + | le_S : ∀ m:nat. le n m → le n (S m). + +interpretation "natural 'less or equal to'" 'leq x y = (le x y). + +interpretation "natural 'neither less nor equal to'" 'nleq x y = (Not (le x y)). + +ndefinition lt: nat → nat → Prop ≝ +λn,m:nat. S n ≤ m. + +interpretation "natural 'less than'" 'lt x y = (lt x y). + +interpretation "natural 'not less than'" 'nless x y = (Not (lt x y)). + +ndefinition ge: nat \to nat \to Prop \def +\lambda n,m:nat.m \leq n. + +interpretation "natural 'greater or equal to'" 'geq x y = (ge x y). + +ndefinition gt: nat \to nat \to Prop \def +\lambda n,m:nat.m (S_pred m); + [ apply le_S_S; + assumption + | assumption + ] +]. +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. + +(* not eq *) +theorem lt_to_not_eq : \forall n,m:nat. n H in H1. +apply (lt_to_not_eq b b) +[ assumption +| reflexivity +] +qed. + +theorem lt_n_m_to_not_lt_m_Sn: ∀n,m. n < m → m ≮ S n. +intros; +unfold Not; +intro; +unfold lt in H; +unfold lt in H1; +generalize in match (le_S_S ? ? H); +intro; +generalize in match (transitive_le ? ? ? H2 H1); +intro; +apply (not_le_Sn_n ? H3). +qed. + +(* le vs. lt *) +theorem lt_to_le : \forall n,m:nat. n H7. +apply H. +apply le_to_or_lt_eq.apply H6. +qed. +*) -- 2.39.2