]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/nlibrary/arithmetics/nat.ma
Curryfication of binary morphisms.
[helm.git] / helm / software / matita / nlibrary / arithmetics / nat.ma
index 3152b0dbe3503aeee43bbd0f872b847f87979239..d1e401ae5f87de195cfdc4edbdd3f30d532b9470 100644 (file)
@@ -76,7 +76,7 @@ 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).
+ntheorem decidable_eq_nat : n,m:nat.decidable (n=m).
 napply nat_elim2; #n;
  ##[ ncases n; /2/;
  ##| /3/;
@@ -260,7 +260,7 @@ ntheorem le_pred_n : ∀n:nat. pred n ≤ n.
 #n; nelim n; //; nqed.
 
 ntheorem monotonic_pred: monotonic ? le pred.
-#n; #m; #lenm; nelim lenm; /2/; nqed.
+#n; #m; #lenm; nelim lenm; //; /2/; nqed.
 
 ntheorem le_S_S_to_le: ∀n,m:nat. S n ≤ S m → n ≤ m.
 /2/; nqed.
@@ -428,9 +428,9 @@ ntheorem le_n_Sm_elim : ∀n,m:nat.n ≤ S m →
 (* le and eq *)
 
 ntheorem le_to_le_to_eq: ∀n,m. n ≤ m → m ≤ n → n = m.
-napply nat_elim2; /3/; nqed.
+napply nat_elim2; /4/; nqed.
 
-ntheorem lt_O_S : \forall n:nat. O < S n.
+ntheorem lt_O_S : n:nat. O < S n.
 /2/; nqed.
 
 (*
@@ -567,6 +567,39 @@ ntheorem le_plus_to_le: ∀a,n,m. a + n ≤ a + m → n ≤ m.
 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<q.
+/2/; nqed.
+
+ntheorem lt_plus_to_lt_r :∀n,p,q:nat. n+p < n+q → p<q.
+/2/; nqed.
+
+ntheorem le_to_lt_to_plus_lt: ∀a,b,c,d:nat.
+a ≤ c → b < d → a + b < c+d.
+(* bello /2/ un po' lento *)
+#a; #b; #c; #d; #leac; #lebd; 
+nnormalize; napplyS le_plus; //; nqed.
+
 (* times *)
 ntheorem monotonic_le_times_r: 
 ∀n:nat.monotonic nat le (λm. n * m).
@@ -616,6 +649,160 @@ ntheorem le_S_times_2: ∀n,m.O < m → n ≤ m → n < 2*m.
 #n; #m; #posm; #lenm; (* interessante *)
 nnormalize; napplyS (le_plus n); //; nqed.
 
+(* times & lt *)
+(*
+ntheorem lt_O_times_S_S: ∀n,m:nat.O < (S n)*(S m).
+intros.simplify.unfold lt.apply le_S_S.apply le_O_n.
+qed. *)
+
+(*
+ntheorem lt_times_eq_O: \forall a,b:nat.
+O < a → a * b = O → b = O.
+intros.
+apply (nat_case1 b)
+[ intros.
+  reflexivity
+| intros.
+  rewrite > 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;
+  ##[napplyS monotonic_lt_plus_l;//;
+  ##|#a; #_; #lt1; napply (transitive_le ??? lt1);//;
+  ##]
+nqed.
+
+ntheorem monotonic_lt_times_r: 
+  ∀c:nat. O < c → monotonic nat lt (λt.(c*t)).
+(* /2/ lentissimo *)
+#c; #posc; #n; #m; #ltnm;
+(* why?? napplyS (monotonic_lt_times_l c posc n m ltnm); *)
+nrewrite > (symmetric_times c n);
+nrewrite > (symmetric_times c m);
+napply monotonic_lt_times_l;//;
+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<m → p<q → n*p < m*q.
+#n; #m; #p; #q; #ltnm; #ltpq;
+napply lt_to_le_to_lt_times;/2/;
+nqed.
+
+ntheorem lt_times_n_to_lt_l: 
+∀n,p,q:nat. O < n → p*n < q*n → p < q.
+#n; #p; #q; #posn; #Hlt;
+nelim (decidable_lt p q);//;
+#nltpq;napply False_ind; 
+napply (lt_to_not_le ? ? Hlt);
+napply monotonic_le_times_l.
+napply not_lt_to_le; //;
+nqed.
+
+ntheorem lt_times_n_to_lt_r: 
+∀n,p,q:nat. O < n → n*p < n*q → p < q.
+#n; #p; #q; #posn; #Hlt;
+napply (lt_times_n_to_lt_l ??? posn);//;
+nqed.
+
+(*
+theorem nat_compare_times_l : \forall n,p,q:nat. 
+nat_compare p q = nat_compare ((S n) * p) ((S n) * q).
+intros.apply nat_compare_elim.intro.
+apply nat_compare_elim.
+intro.reflexivity.
+intro.absurd (p=q).
+apply (inj_times_r n).assumption.
+apply lt_to_not_eq. assumption.
+intro.absurd (q<p).
+apply (lt_times_to_lt_r n).assumption.
+apply le_to_not_lt.apply lt_to_le.assumption.
+intro.rewrite < H.rewrite > nat_compare_n_n.reflexivity.
+intro.apply nat_compare_elim.intro.
+absurd (p<q).
+apply (lt_times_to_lt_r n).assumption.
+apply le_to_not_lt.apply lt_to_le.assumption.
+intro.absurd (q=p).
+symmetry.
+apply (inj_times_r n).assumption.
+apply lt_to_not_eq.assumption.
+intro.reflexivity.
+qed. *)
+
+(* times and plus 
+theorem lt_times_plus_times: \forall a,b,n,m:nat. 
+a < n \to b < m \to a*m + b < n*m.
+intros 3.
+apply (nat_case n)
+  [intros.apply False_ind.apply (not_le_Sn_O ? H)
+  |intros.simplify.
+   rewrite < sym_plus.
+   unfold.
+   change with (S b+a*m1 \leq m1+m*m1).
+   apply le_plus
+    [assumption
+    |apply le_times
+      [apply le_S_S_to_le.assumption
+      |apply le_n
+      ]
+    ]
+  ]
+qed. *)
+
 (************************** minus ******************************)
 
 nlet rec minus n m ≝ 
@@ -649,7 +836,7 @@ ntheorem minus_Sn_m: ∀m,n:nat. m ≤ n → S n -m = S (n-m).
 napply nat_elim2; 
   ##[//
   ##|#n; #abs; napply False_ind;/2/;
-  ##|/3/;
+  ##|#n; #m; #Hind; #c; napplyS Hind; /2/;
   ##]
 nqed.
 
@@ -676,7 +863,8 @@ napplyS (plus_minus m n m); //; nqed.
 ntheorem le_plus_minus_m_m: ∀n,m:nat. n ≤ (n-m)+m.
 #n; nelim n;
   ##[//
-  ##|#a; #Hind; #m; ncases m;/2/;  
+  ##|#a; #Hind; #m; ncases m;//;  
+     nnormalize; #n;napplyS le_S_S;//  
   ##]
 nqed.
 
@@ -793,3 +981,154 @@ ntheorem monotonic_le_minus_r:
 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.
+
+theorem eqb_elim : \forall n,m:nat.\forall P:bool \to Prop.
+(n=m \to (P true)) \to (n \neq m \to (P false)) \to (P (eqb n m)). 
+intros.
+cut 
+(match (eqb n m) with
+[ true  \Rightarrow n = m
+| false \Rightarrow n \neq m] \to (P (eqb n m))).
+apply Hcut.apply eqb_to_Prop.
+elim (eqb n m).
+apply ((H H2)).
+apply ((H1 H2)).
+qed. 
+
+*)
+
+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.
+napply nat_elim2; 
+  ##[#n; ncases n; nnormalize; //; 
+     #m; #abs; napply False_ind;/2/;
+  ##|nnormalize; #m; #abs; napply False_ind;/2/;
+  ##|nnormalize; 
+     #n; #m; #Hind; #eqnm; napply eq_f; napply Hind; //;
+  ##]
+nqed.
+
+ntheorem eqb_false_to_not_eq: ∀n,m:nat.
+  eqb n m = false → n ≠ m.
+napply nat_elim2; 
+  ##[#n; ncases n; nnormalize; /2/; 
+  ##|/2/;
+  ##|nnormalize;/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; 
+nelim (true_or_false (eqb n m)); //;
+#Heq; napply False_ind; napply noteq;/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; #leSS; /3/;
+    ##]
+  ##]
+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 lt_to_leb_false: ∀n,m. m < n → leb n m = false.
+#n; #m; napply leb_elim; //;
+#H; #H1; napply False_ind; /2/;
+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. *)