]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/nlibrary/arithmetics/nat.ma
some experiment filtering with height
[helm.git] / helm / software / matita / nlibrary / arithmetics / nat.ma
index 0418444e376e0ecd47ca07a83b8915b1aa8cbc57..393f16189a9001d8dacab469c43f285ce8123533 100644 (file)
@@ -78,7 +78,7 @@ ntheorem nat_elim2 :
 
 ntheorem decidable_eq_nat : ∀n,m:nat.decidable (n=m).
 napply nat_elim2; #n;
- ##[ ncases n; /2/;
+ ##[ ncases n; /3/;
  ##| /3/;
  ##| #m; #Hind; ncases Hind; /3/;
  ##]
@@ -259,13 +259,18 @@ ntheorem le_n_Sn : ∀n:nat. n ≤ S n.
 ntheorem le_pred_n : ∀n:nat. pred n ≤ n.
 #n; nelim n; //; nqed.
 
+(* XXX global problem *)
+nlemma my_trans_le : ∀x,y,z:nat.x ≤ y → y ≤ z → x ≤ z. 
+napply transitive_le.
+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.
+(* XXX *) nletin hint ≝ monotonic. /2/; nqed.
 
-ntheorem lt_S_S_to_lt: ∀n,m. S n < S m \to n < m.
+ntheorem lt_S_S_to_lt: ∀n,m. S n < S m  n < m.
 /2/; nqed. 
 
 ntheorem lt_to_lt_S_S: ∀n,m. n < m → S n < S m.
@@ -285,14 +290,14 @@ 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.
+napply nat_elim2; #n; /3/;
+#m; #dec; ncases dec;/4/; 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.
+#n; nelim n; /3/; nqed.
 
 ntheorem lt_S_to_le: ∀n,m:nat. n < S m → n ≤ m.
 /2/; nqed.
@@ -301,7 +306,7 @@ 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/;
+ ##|#m;#Hind;#HnotleSS; napply lt_to_lt_S_S;/4/;
  ##]
 nqed.
 
@@ -428,9 +433,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.
 
 (*
@@ -568,9 +573,10 @@ 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. 
+/2/; nqed.
 
 (*
 variant lt_plus_r: \forall n,p,q:nat. p < q \to n + p < n + q \def
@@ -749,8 +755,7 @@ ntheorem lt_times_n_to_lt_l:
 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; //;
+napply monotonic_le_times_l;/3/;
 nqed.
 
 ntheorem lt_times_n_to_lt_r: 
@@ -835,11 +840,18 @@ ntheorem minus_Sn_m: ∀m,n:nat. m ≤ n → S n -m = S (n-m).
 #n; #m; #lenm; nelim lenm; napplyS refl_eq. *)
 napply nat_elim2; 
   ##[//
-  ##|#n; #abs; napply False_ind;/2/;
-  ##|/3/;
+  ##|#n; #abs; napply False_ind; (* XXX *) napply not_le_Sn_O; /2/.
+  ##|#n; #m; #Hind; #c; napplyS Hind; /2/;
   ##]
 nqed.
 
+ntheorem not_eq_to_le_to_le_minus: 
+  ∀n,m.n ≠ m → n ≤ m → n ≤ m - 1.
+#n; #m; ncases m;//; #m; nnormalize;
+#H; #H1; napply le_S_S_to_le;
+napplyS (not_eq_to_le_to_lt n (S m) H H1);
+nqed.
+
 ntheorem eq_minus_S_pred: ∀n,m. n - (S m) = pred(n -m).
 napply nat_elim2; //; nqed.
 
@@ -847,7 +859,7 @@ ntheorem plus_minus:
 ∀m,n,p:nat. m ≤ n → (n-m)+p = (n+p)-m.
 napply nat_elim2; 
   ##[//
-  ##|#n; #p; #abs; napply False_ind;/2/;
+  ##|#n; #p; #abs; napply False_ind; (* XXX *) napply not_le_Sn_O; /2/;
   ##|nnormalize;/3/;
   ##]
 nqed.
@@ -981,3 +993,139 @@ 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.
+*)
+
+ntheorem 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; (* XXX *) nletin hint ≝ not_eq_O_S; /3/; 
+  ##|nnormalize; /4/; 
+  ##] (* XXX rimane aperto *) #m; #P; #_; #H; napply H; napply not_eq_O_S.
+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; (* XXX *) nletin hint ≝ not_eq_true_false; /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; 
+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/
+  ##| (* XXX *) nletin hint ≝ not_le_Sn_O; /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; (* XXX *) nletin hint ≝ not_eq_true_false; /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; (* XXX *) nletin hint ≝ not_eq_true_false; /2/;
+  ##|/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. *)