X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fnlibrary%2Farithmetics%2Fnat.ma;h=fe9ef8267ffc064e1d50f9e3283b6f0c5e0c3660;hb=fb0a4ba033da3ea305e0add8c2ec4573e95fa3d2;hp=a1e6fbb230c6f7cd93e444aa96a4fc681bfbacd1;hpb=9f0b17f553522373b51fc4761d5a3e1e9dbbad73;p=helm.git diff --git a/helm/software/matita/nlibrary/arithmetics/nat.ma b/helm/software/matita/nlibrary/arithmetics/nat.ma index a1e6fbb23..fe9ef8267 100644 --- a/helm/software/matita/nlibrary/arithmetics/nat.ma +++ b/helm/software/matita/nlibrary/arithmetics/nat.ma @@ -15,7 +15,11 @@ (* include "higher_order_defs/functions.ma". *) include "hints_declaration.ma". include "basics/functions.ma". -include "basics/eq.ma". +include "basics/eq.ma". + +ntheorem foo: ∀A:Type.∀a,b:A.∀f:A→A.∀g:A→A→A. +(∀x,y.f (g x y) = x) → ∀x. g (f a) x = b → f a = f b. +//; nqed. ninductive nat : Type[0] ≝ | O : nat @@ -36,12 +40,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. @@ -51,19 +53,18 @@ ntheorem inj_S : \forall n,m:nat.(S n)=(S m) \to n=m. //. nqed. *) ntheorem not_eq_S: ∀n,m:nat. n ≠ m → S n ≠ S m. -/2/; nqed. +/3/; nqed. 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; //. +#n; napply nmk; #eqOS; nchange with (not_zero O); nrewrite > eqOS; //. nqed. -ntheorem not_eq_n_Sn : ∀n:nat. n ≠ S n. -#n; nelim n; /2/; nqed. +ntheorem not_eq_n_Sn: ∀n:nat. n ≠ S n. +#n; nelim n;/2/; nqed. ntheorem nat_case: ∀n:nat.∀P:nat → Prop. @@ -79,11 +80,11 @@ ntheorem nat_elim2 : #R; #ROn; #RSO; #RSS; #n; nelim n;//; #n0; #Rn0m; #m; ncases m;/2/; nqed. -ntheorem decidable_eq_nat : \forall n,m:nat.decidable (n=m). -napply nat_elim2;#n; +ntheorem decidable_eq_nat : ∀n,m:nat.decidable (n=m). +napply nat_elim2; #n; ##[ ncases n; /2/; ##| /3/; - ##| #m; #Hind; ncases Hind; /3/; (* ??? /2/;#neqnm; /3/; *) + ##| #m; #Hind; ncases Hind; /3/; ##] nqed. @@ -115,9 +116,10 @@ ntheorem plus_Sn_m1: ∀n,m:nat. S m + n = n + S m. #n; nelim n; nnormalize; //; nqed. *) -(* -ntheorem plus_n_SO : ∀n:nat. S n = n+S O. -//; nqed. *) +(* deleterio? +ntheorem plus_n_1 : ∀n:nat. S n = n+1. +//; nqed. +*) ntheorem symmetric_plus: symmetric ? plus. #n; nelim n; nnormalize; //; nqed. @@ -125,7 +127,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). @@ -162,7 +164,7 @@ ntheorem times_n_Sm : ∀n,m:nat. n+(n*m) = n*(S m). #n; nelim n; nnormalize; //; nqed. ntheorem symmetric_times : symmetric nat times. -#n; nelim n; nnormalize; //; nqed. +#n; nelim n; nnormalize; //; nqed. (* variant sym_times : \forall n,m:nat. n*m = m*n \def symmetric_times. *) @@ -170,9 +172,9 @@ symmetric_times. *) ntheorem distributive_times_plus : distributive nat times plus. #n; nelim n; nnormalize; //; nqed. -ntheorem distributive_times_plus_r: -\forall a,b,c:nat. (b+c)*a = b*a + c*a. -//; nqed. +ntheorem distributive_times_plus_r : + ∀a,b,c:nat. (b+c)*a = b*a + c*a. +//; nqed. ntheorem associative_times: associative nat times. #n; nelim n; nnormalize; //; nqed. @@ -206,5 +208,1046 @@ 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)). + +(* nlemma eq_lt: ∀n,m. (n < m) = (S n ≤ m). +//; nqed. *) + +ndefinition ge: nat → nat → Prop ≝ +λn,m:nat.m ≤ n. + +interpretation "natural 'greater or equal to'" 'geq x y = (ge x y). + +ndefinition gt: nat → nat → Prop ≝ +λn,m:nat.m (S_pred m); + [ apply le_S_S; + assumption + | assumption + ] +]. +qed. + +*) + +(* le to lt or eq *) +ntheorem le_to_or_lt_eq: ∀n,m:nat. n ≤ m → n < m ∨ n = m. +#n; #m; #lenm; nelim lenm; /3/; nqed. + +(* not eq *) +ntheorem lt_to_not_eq : ∀n,m:nat. n < m → n ≠ m. +#n; #m; #H; napply not_to_not;/2/; nqed. + +(*not lt +ntheorem eq_to_not_lt: ∀a,b:nat. a = b → a ≮ b. +intros. +unfold Not. +intros. +rewrite > 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. *) + +ntheorem not_eq_to_le_to_lt: ∀n,m. n≠m → n≤m → n H7. +apply H. +apply le_to_or_lt_eq.apply H6. +qed. +*) + +(*********************** monotonicity ***************************) +ntheorem monotonic_le_plus_r: +∀n:nat.monotonic nat le (λm.n + m). +#n; #a; #b; nelim n; nnormalize; //; +#m; #H; #leab;napply le_S_S; /2/; nqed. + +(* +ntheorem le_plus_r: ∀p,n,m:nat. n ≤ m → p + n ≤ p + m +≝ monotonic_le_plus_r. *) + +ntheorem monotonic_le_plus_l: +∀m:nat.monotonic nat le (λn.n + m). +#m; #x; #y; #H; napplyS monotonic_le_plus_r; +/2/; nqed. +(* +ntheorem le_plus_l: \forall p,n,m:nat. n \le m \to n + p \le m + p +\def monotonic_le_plus_l. *) + +ntheorem le_plus: ∀n1,n2,m1,m2:nat. n1 ≤ n2 → m1 ≤ m2 +→ n1 + m1 ≤ n2 + m2. +#n1; #n2; #m1; #m2; #len; #lem; napply (transitive_le ? (n1+m2)); +/2/; nqed. + +ntheorem le_plus_n :∀n,m:nat. m ≤ n + m. +/2/; nqed. + +nlemma le_plus_a: ∀a,n,m. n ≤ m → n ≤ a + m. +/2/; nqed. + +nlemma le_plus_b: ∀b,n,m. n + b ≤ m → n ≤ m. +/2/; nqed. + +ntheorem le_plus_n_r :∀n,m:nat. m ≤ m + n. +/2/; nqed. + +ntheorem eq_plus_to_le: ∀n,m,p:nat.n=m+p → m ≤ n. +//; nqed. + +ntheorem le_plus_to_le: ∀a,n,m. a + n ≤ a + m → n ≤ m. +#a; nelim a; nnormalize; /3/; nqed. + +ntheorem le_plus_to_le_r: ∀a,n,m. n + a ≤ m +a → n ≤ m. +/2/; nqed. + +(* plus & lt *) + +ntheorem monotonic_lt_plus_r: +∀n:nat.monotonic nat lt (λm.n+m). +/2/; nqed. + +(* +variant lt_plus_r: \forall n,p,q:nat. p < q \to n + p < n + q \def +monotonic_lt_plus_r. *) + +ntheorem monotonic_lt_plus_l: +∀n:nat.monotonic nat lt (λm.m+n). +/2/; nqed. + +(* +variant lt_plus_l: \forall n,p,q:nat. p < q \to p + n < q + n \def +monotonic_lt_plus_l. *) + +ntheorem lt_plus: ∀n,m,p,q:nat. n < m → p < q → n + p < m + q. +#n; #m; #p; #q; #ltnm; #ltpq; +napply (transitive_lt ? (n+q));/2/; nqed. + +ntheorem lt_plus_to_lt_l :∀n,p,q:nat. p+n < q+n → p H2 in H1. + rewrite > (S_pred a) in H1 + [ apply False_ind. + apply (eq_to_not_lt O ((S (pred a))*(S m))) + [ apply sym_eq. + assumption + | apply lt_O_times_S_S + ] + | assumption + ] +] +qed. + +theorem O_lt_times_to_O_lt: \forall a,c:nat. +O \lt (a * c) \to O \lt a. +intros. +apply (nat_case1 a) +[ intros. + rewrite > H1 in H. + simplify in H. + assumption +| intros. + apply lt_O_S +] +qed. + +lemma lt_times_to_lt_O: \forall i,n,m:nat. i < n*m \to O < m. +intros. +elim (le_to_or_lt_eq O ? (le_O_n m)) + [assumption + |apply False_ind. + rewrite < H1 in H. + rewrite < times_n_O in H. + apply (not_le_Sn_O ? H) + ] +qed. *) + +(* +ntheorem monotonic_lt_times_r: +∀n:nat.monotonic nat lt (λm.(S n)*m). +/2/; +simplify. +intros.elim n. +simplify.rewrite < plus_n_O.rewrite < plus_n_O.assumption. +apply lt_plus.assumption.assumption. +qed. *) + +ntheorem monotonic_lt_times_l: + ∀c:nat. O < c → monotonic nat lt (λt.(t*c)). +#c; #posc; #n; #m; #ltnm; +nelim ltnm; nnormalize; + ##[/2/; + ##|#a; #_; #lt1; napply (transitive_le ??? lt1);//; + ##] +nqed. + +ntheorem monotonic_lt_times_r: + ∀c:nat. O < c → monotonic nat lt (λt.(c*t)). +/2/; nqed. + +ntheorem lt_to_le_to_lt_times: +∀n,m,p,q:nat. n < m → p ≤ q → O < q → n*p < m*q. +#n; #m; #p; #q; #ltnm; #lepq; #posq; +napply (le_to_lt_to_lt ? (n*q)); + ##[napply monotonic_le_times_r;//; + ##|napply monotonic_lt_times_l;//; + ##] +nqed. + +ntheorem lt_times:∀n,m,p,q:nat. n nat_compare_n_n.reflexivity. +intro.apply nat_compare_elim.intro. +absurd (p minus_Sn_m. +apply le_S.assumption. +apply lt_to_le.assumption. +qed. + +theorem minus_le_S_minus_S: \forall n,m:nat. m-n \leq S (m-(S n)). +intros. +apply (nat_elim2 (\lambda n,m.m-n \leq S (m-(S n)))). +intro.elim n1.simplify.apply le_n_Sn. +simplify.rewrite < minus_n_O.apply le_n. +intros.simplify.apply le_n_Sn. +intros.simplify.apply H. +qed. + +theorem lt_minus_S_n_to_le_minus_n : \forall n,m,p:nat. m-(S n) < p \to m-n \leq p. +intros 3.intro. +(* autobatch *) +(* end auto($Revision: 9739 $) proof: TIME=1.33 SIZE=100 DEPTH=100 *) +apply (trans_le (m-n) (S (m-(S n))) p). +apply minus_le_S_minus_S. +assumption. +qed. + +theorem le_minus_m: \forall n,m:nat. n-m \leq n. +intros.apply (nat_elim2 (\lambda m,n. n-m \leq n)). +intros.rewrite < minus_n_O.apply le_n. +intros.simplify.apply le_n. +intros.simplify.apply le_S.assumption. +qed. + +theorem lt_minus_m: \forall n,m:nat. O < n \to O < m \to n-m \lt n. +intros.apply (lt_O_n_elim n H).intro. +apply (lt_O_n_elim m H1).intro. +simplify.unfold lt.apply le_S_S.apply le_minus_m. +qed. + +theorem minus_le_O_to_le: \forall n,m:nat. n-m \leq O \to n \leq m. +intros 2. +apply (nat_elim2 (\lambda n,m:nat.n-m \leq O \to n \leq m)). +intros.apply le_O_n. +simplify.intros. assumption. +simplify.intros.apply le_S_S.apply H.assumption. +qed. +*) + +(* monotonicity and galois *) + +ntheorem monotonic_le_minus_l: +∀p,q,n:nat. q ≤ p → q-n ≤ p-n. +napply nat_elim2; #p; #q; + ##[#lePO; napply (le_n_O_elim ? lePO);//; + ##|//; + ##|#Hind; #n; ncases n; + ##[//; + ##|#a; #leSS; napply Hind; /2/; + ##] + ##] +nqed. + +ntheorem le_minus_to_plus: ∀n,m,p. n-m ≤ p → n≤ p+m. +#n; #m; #p; #lep; +napply transitive_le; + ##[##|napply le_plus_minus_m_m + ##|napply monotonic_le_plus_l;//; + ##] +nqed. + +ntheorem le_plus_to_minus: ∀n,m,p. n ≤ p+m → n-m ≤ p. +#n; #m; #p; #lep; +(* bello *) +napplyS monotonic_le_minus_l;//; +(* /2/; *) +nqed. + +ntheorem monotonic_le_minus_r: +∀p,q,n:nat. q ≤ p → n-p ≤ n-q. +#p; #q; #n; #lepq; +napply le_plus_to_minus; +napply (transitive_le ??? (le_plus_minus_m_m ? q));/2/; +nqed. + +(*********************** boolean arithmetics ********************) +include "basics/bool.ma". + +nlet rec eqb n m ≝ +match n with + [ O ⇒ match m with [ O ⇒ true | S q ⇒ false] + | S p ⇒ match m with [ O ⇒ false | S q ⇒ eqb p q] + ]. + +(* +ntheorem eqb_to_Prop: ∀n,m:nat. +match (eqb n m) with +[ true \Rightarrow n = m +| false \Rightarrow n \neq m]. +intros. +apply (nat_elim2 +(\lambda n,m:nat.match (eqb n m) with +[ true \Rightarrow n = m +| false \Rightarrow n \neq m])). +intro.elim n1. +simplify.reflexivity. +simplify.apply not_eq_O_S. +intro. +simplify.unfold Not. +intro. apply (not_eq_O_S n1).apply sym_eq.assumption. +intros.simplify. +generalize in match H. +elim ((eqb n1 m1)). +simplify.apply eq_f.apply H1. +simplify.unfold Not.intro.apply H1.apply inj_S.assumption. +qed. +*) + +naxiom eqb_elim : ∀ n,m:nat.∀ P:bool → Prop. +(n=m → (P true)) → (n ≠ m → (P false)) → (P (eqb n m)). +(* +napply nat_elim2; + ##[#n; ncases n; nnormalize; /3/; + ##|nnormalize; /3/; + ##|nnormalize; /4/; + ##] +nqed.*) + +ntheorem eqb_n_n: ∀n. eqb n n = true. +#n; nelim n; nnormalize; //. +nqed. + +ntheorem eqb_true_to_eq: ∀n,m:nat. eqb n m = true → n = m. +#n; #m; napply (eqb_elim n m);//; +#_; #abs; napply False_ind; /2/; +nqed. + +ntheorem eqb_false_to_not_eq: ∀n,m:nat. eqb n m = false → n ≠ m. +#n; #m; napply (eqb_elim n m);/2/; +nqed. + +ntheorem eq_to_eqb_true: ∀n,m:nat. + n = m → eqb n m = true. +//; nqed. + +ntheorem not_eq_to_eqb_false: ∀n,m:nat. + n ≠ m → eqb n m = false. +#n; #m; #noteq; +napply eqb_elim;//; +#Heq; napply False_ind; /2/; +nqed. + +nlet rec leb n m ≝ +match n with + [ O ⇒ true + | (S p) ⇒ + match m with + [ O ⇒ false + | (S q) ⇒ leb p q]]. + +ntheorem leb_elim: ∀n,m:nat. ∀P:bool → Prop. +(n ≤ m → P true) → (n ≰ m → P false) → P (leb n m). +napply nat_elim2; nnormalize; + ##[/2/ + ##|/3/; + ##|#n; #m; #Hind; #P; #Pt; #Pf; napply Hind; + ##[#lenm; napply Pt; napply le_S_S;//; + ##|#nlenm; napply Pf; /2/; + ##] + ##] +nqed. + +ntheorem leb_true_to_le:∀n,m.leb n m = true → n ≤ m. +#n; #m; napply leb_elim; + ##[//; + ##|#_; #abs; napply False_ind; /2/; + ##] +nqed. + +ntheorem leb_false_to_not_le:∀n,m. + leb n m = false → n ≰ m. +#n; #m; napply leb_elim; + ##[#_; #abs; napply False_ind; /2/; + ##|//; + ##] +nqed. + +ntheorem le_to_leb_true: ∀n,m. n ≤ m → leb n m = true. +#n; #m; napply leb_elim; //; +#H; #H1; napply False_ind; /2/; +nqed. + +ntheorem not_le_to_leb_false: ∀n,m. n ≰ m → leb n m = false. +#n; #m; napply leb_elim; //; +#H; #H1; napply False_ind; /2/; +nqed. + +ntheorem lt_to_leb_false: ∀n,m. m < n → leb n m = false. +/3/; nqed. + +(* serve anche ltb? +ndefinition ltb ≝λn,m. leb (S n) m. + +ntheorem ltb_elim: ∀n,m:nat. ∀P:bool → Prop. +(n < m → P true) → (n ≮ m → P false) → P (ltb n m). +#n; #m; #P; #Hlt; #Hnlt; +napply leb_elim; /3/; nqed. + +ntheorem ltb_true_to_lt:∀n,m.ltb n m = true → n < m. +#n; #m; #Hltb; napply leb_true_to_le; nassumption; +nqed. + +ntheorem ltb_false_to_not_lt:∀n,m. + ltb n m = false → n ≮ m. +#n; #m; #Hltb; napply leb_false_to_not_le; nassumption; +nqed. + +ntheorem lt_to_ltb_true: ∀n,m. n < m → ltb n m = true. +#n; #m; #Hltb; napply le_to_leb_true; nassumption; +nqed. + +ntheorem le_to_ltb_false: ∀n,m. m \le n → ltb n m = false. +#n; #m; #Hltb; napply lt_to_leb_false; /2/; +nqed. *) + +ninductive compare : Type[0] ≝ +| LT : compare +| EQ : compare +| GT : compare. + +ndefinition compare_invert: compare → compare ≝ + λc.match c with + [ LT ⇒ GT + | EQ ⇒ EQ + | GT ⇒ LT ]. + +nlet rec nat_compare n m: compare ≝ +match n with +[ O ⇒ match m with + [ O ⇒ EQ + | (S q) ⇒ LT ] +| S p ⇒ match m with + [ O ⇒ GT + | S q ⇒ nat_compare p q]]. + +ntheorem nat_compare_n_n: ∀n. nat_compare n n = EQ. +#n;nelim n +##[// +##|#m;#IH;nnormalize;//] +nqed. + +ntheorem nat_compare_S_S: ∀n,m:nat.nat_compare n m = nat_compare (S n) (S m). +//; +nqed. + +ntheorem nat_compare_pred_pred: + ∀n,m.O < n → O < m → nat_compare n m = nat_compare (pred n) (pred m). +#n;#m;#Hn;#Hm; +napply (lt_O_n_elim n Hn); +napply (lt_O_n_elim m Hm); +#p;#q;//; +nqed. + +ntheorem nat_compare_to_Prop: + ∀n,m.match (nat_compare n m) with + [ LT ⇒ n < m + | EQ ⇒ n = m + | GT ⇒ m < n ]. +#n;#m; +napply (nat_elim2 (λn,m.match (nat_compare n m) with + [ LT ⇒ n < m + | EQ ⇒ n = m + | GT ⇒ m < n ]) ?????) (* FIXME: don't want to put all these ?, especially when … does not work! *) +##[##1,2:#n1;ncases n1;//; +##|#n1;#m1;nnormalize;ncases (nat_compare n1 m1); + ##[##1,3:nnormalize;#IH;napply le_S_S;//; + ##|nnormalize;#IH;nrewrite > IH;//] +nqed. + +ntheorem nat_compare_n_m_m_n: + ∀n,m:nat.nat_compare n m = compare_invert (nat_compare m n). +#n;#m; +napply (nat_elim2 (λn,m. nat_compare n m = compare_invert (nat_compare m n))) +##[##1,2:#n1;ncases n1;//; +##|#n1;#m1;#IH;nnormalize;napply IH] +nqed. + +ntheorem nat_compare_elim : + ∀n,m. ∀P:compare → Prop. + (n < m → P LT) → (n=m → P EQ) → (m < n → P GT) → P (nat_compare n m). +#n;#m;#P;#Hlt;#Heq;#Hgt; +ncut (match (nat_compare n m) with + [ LT ⇒ n < m + | EQ ⇒ n=m + | GT ⇒ m < n] → + P (nat_compare n m)) +##[ncases (nat_compare n m); + ##[napply Hlt + ##|napply Heq + ##|napply Hgt] +##|#Hcut;napply Hcut;//; +nqed. + +ninductive cmp_cases (n,m:nat) : CProp[0] ≝ + | cmp_le : n ≤ m → cmp_cases n m + | cmp_gt : m < n → cmp_cases n m. + +ntheorem lt_to_le : ∀n,m:nat. n < m → n ≤ m. +#n;#m;#H;nelim H +##[// +##|/2/] +nqed. + +nlemma cmp_nat: ∀n,m.cmp_cases n m. +#n;#m; nlapply (nat_compare_to_Prop n m); +ncases (nat_compare n m);#H +##[@;napply lt_to_le;// +##|@;// +##|@2;//] +nqed.