]> matita.cs.unibo.it Git - helm.git/commitdiff
Nuova versione di not.
authorAndrea Asperti <andrea.asperti@unibo.it>
Thu, 18 Mar 2010 11:35:44 +0000 (11:35 +0000)
committerAndrea Asperti <andrea.asperti@unibo.it>
Thu, 18 Mar 2010 11:35:44 +0000 (11:35 +0000)
helm/software/matita/nlibrary/arithmetics/nat.ma

index 25745e370c19d3ed3072e525630234461c6a5725..fe9ef8267ffc064e1d50f9e3283b6f0c5e0c3660 100644 (file)
 (* 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
@@ -49,18 +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 ].
 
 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. 
@@ -78,7 +82,7 @@ ntheorem nat_elim2 :
 
 ntheorem decidable_eq_nat : ∀n,m:nat.decidable (n=m).
 napply nat_elim2; #n;
- ##[ ncases n; /3/;
+ ##[ ncases n; /2/;
  ##| /3/;
  ##| #m; #Hind; ncases Hind; /3/;
  ##]
@@ -112,9 +116,10 @@ ntheorem plus_Sn_m1: ∀n,m:nat. S m + n = n + S m.
 #n; nelim n; nnormalize; //; nqed.
 *)
 
-(* deleterio 
-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. 
@@ -159,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. *)
@@ -167,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.
@@ -223,13 +228,13 @@ 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 \to nat \to Prop \def
-\lambda n,m:nat.m \leq n.
+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 \to nat \to Prop \def
-\lambda n,m:nat.m<n.
+ndefinition gt: nat → nat → Prop ≝
+λn,m:nat.m<n.
 
 interpretation "natural 'greater than'" 'gt x y = (gt x y).
 
@@ -244,8 +249,8 @@ ntheorem trans_le: \forall n,m,p:nat. n \leq m \to m \leq p \to n \leq p
 \def transitive_le. *)
 
 
-naxiom transitive_lt: transitive nat lt.
-(* #a; #b; #c; #ltab; #ltbc;nelim ltbc;/2/;nqed.*)
+ntheorem transitive_lt: transitive nat lt.
+#a; #b; #c; #ltab; #ltbc;nelim ltbc;/2/;nqed.
 
 (*
 theorem trans_lt: \forall n,m,p:nat. lt n m \to lt m p \to lt n p
@@ -263,29 +268,31 @@ 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 *)
+(* XXX global problem 
 nlemma my_trans_le : ∀x,y,z:nat.x ≤ y → y ≤ z → x ≤ z. 
 napply transitive_le.
-nqed.
+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.
-(* XXX *) nletin hint ≝ monotonic. /2/; nqed.
+(* XXX *) nletin hint ≝ monotonic. 
+/2/; nqed.
 
+(* this are instances of the le versions 
 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.
-/2/; nqed.
+/2/; nqed. *)
 
 ntheorem lt_to_not_zero : ∀n,m:nat. n < m → not_zero m.
 #n; #m; #Hlt; nelim Hlt;//; nqed.
 
 (* lt vs. le *)
 ntheorem not_le_Sn_O: ∀ n:nat. S n ≰ O.
-#n; #Hlen0; napply (lt_to_not_zero ?? Hlen0); nqed.
+#n; napply nmk; #Hlen0; napply (lt_to_not_zero ?? Hlen0); nqed.
 
 ntheorem not_le_to_not_le_S_S: ∀ n,m:nat. n ≰ m → S n ≰ S m.
 /3/; nqed.
@@ -293,26 +300,26 @@ ntheorem not_le_to_not_le_S_S: ∀ n,m:nat. n ≰ m → S n ≰ S m.
 ntheorem not_le_S_S_to_not_le: ∀ n,m:nat. S n ≰ S m → n ≰ m.
 /3/; nqed.
 
-naxiom decidable_le: ∀n,m. decidable (n≤m).
-(*
 ntheorem decidable_le: ∀n,m. decidable (n≤m).
-napply nat_elim2; #n; /3/;
-#m; #dec; ncases dec;/4/; nqed. *)
+napply nat_elim2; #n; /2/;
+#m; *; /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; /3/; nqed.
+#n; nelim n; /2/; nqed.
 
+(* this is le_S_S_to_le
 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;/4/;
+ ##|#m;#Hind;#HnotleSS; napply le_S_S;/3/;
  ##]
 nqed.
 
@@ -320,13 +327,15 @@ 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: 
-   we need an extra depths for unfolding not *)
-napply not_le_to_lt; napply Hnlt; nqed. 
+/4/; nqed.
+
+(*
+#n; #m; #Hnlt; napply le_S_S_to_le;/2/;
+(* 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.
+#n; #m; #H;napply lt_to_not_le; /2/; (* /3/ *) nqed.
 
 (* lt and le trans *)
 
@@ -355,7 +364,8 @@ 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.
+#n; nelim n; //; #abs; napply False_ind;/2/;
+nqed.
 
 (*
 theorem lt_pred: \forall n,m. 
@@ -394,7 +404,7 @@ ntheorem le_to_or_lt_eq: ∀n,m:nat. n ≤ m → n < m ∨ n = m.
 
 (* not eq *)
 ntheorem lt_to_not_eq : ∀n,m:nat. n < m → n ≠ m.
-/2/; nqed.
+#n; #m; #H; napply not_to_not;/2/; nqed.
 
 (*not lt 
 ntheorem eq_to_not_lt: ∀a,b:nat. a = b → a ≮ b.
@@ -423,14 +433,18 @@ qed. *)
 
 ntheorem not_eq_to_le_to_lt: ∀n,m. n≠m → n≤m → n<m.
 #n; #m; #Hneq; #Hle; ncases (le_to_or_lt_eq ?? Hle); //;
-#Heq; nelim (Hneq Heq); nqed.
+#Heq; /3/; nqed.
+(*
+nelim (Hneq Heq); nqed. *)
 
 (* le elimination *)
 ntheorem le_n_O_to_eq : ∀n:nat. n ≤ O → O=n.
-#n; ncases n; //; #a ; #abs; nelim (not_le_Sn_O ? abs); nqed.
+#n; ncases n; //; #a ; #abs;
+napply False_ind; /2/;nqed.
 
 ntheorem le_n_O_elim: ∀n:nat. n ≤ O → ∀P: nat →Prop. P O → P n.
-#n; ncases n; //; #a; #abs; nelim (not_le_Sn_O ? abs); nqed. 
+#n; ncases n; //; #a; #abs; 
+napply False_ind; /2/; nqed. 
 
 ntheorem le_n_Sm_elim : ∀n,m:nat.n ≤ S m → 
 ∀P:Prop. (S n ≤ S m → P) → (n=S m → P) → P.
@@ -439,7 +453,7 @@ 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; /4/; nqed.
+napply nat_elim2; /4/; nqed. 
 
 ntheorem lt_O_S : ∀n:nat. O < S n.
 /2/; nqed.
@@ -479,7 +493,7 @@ nelim n;
  ##[#q; #HleO; (* applica male *) 
     napply (le_n_O_elim ? HleO);
     napply H; #p; #ltpO;
-    napply False_ind; /2/; 
+    napply False_ind; /2/; (* 3 *)
  ##|#p; #Hind; #q; #HleS; 
     napply H; #a; #lta; napply Hind;
     napply le_S_S_to_le;/2/;
@@ -552,13 +566,14 @@ ntheorem le_plus_r: ∀p,n,m:nat. n ≤ m → p + n ≤ p + m
 
 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  \to m1 ≤ m2 
+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.
@@ -566,6 +581,12 @@ ntheorem le_plus: ∀n1,n2,m1,m2:nat. n1 ≤ n2  \to m1 ≤ m2
 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.
 
@@ -590,7 +611,7 @@ monotonic_lt_plus_r. *)
 
 ntheorem monotonic_lt_plus_l: 
 ∀n:nat.monotonic nat lt (λm.m+n).
-/2/;nqed.
+/2/; nqed.
 
 (*
 variant lt_plus_l: \forall n,p,q:nat. p < q \to p + n < q + n \def
@@ -606,11 +627,13 @@ ntheorem lt_plus_to_lt_l :∀n,p,q:nat. p+n < q+n → p<q.
 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.
+(*
+ntheorem le_to_lt_to_lt_plus: ∀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: 
@@ -623,9 +646,11 @@ nqed.
 ntheorem le_times_r: \forall p,n,m:nat. n \le m \to p*n \le p*m
 \def monotonic_le_times_r. *)
 
+(*
 ntheorem monotonic_le_times_l: 
 ∀m:nat.monotonic nat le (λn.n*m).
 /2/; nqed.
+*)
 
 (*
 theorem le_times_l: \forall p,n,m:nat. n \le m \to n*p \le m*p
@@ -634,32 +659,30 @@ theorem le_times_l: \forall p,n,m:nat. n \le m \to n*p \le m*p
 ntheorem le_times: ∀n1,n2,m1,m2:nat. 
 n1 ≤ n2  → m1 ≤ m2 → n1*m1 ≤ n2*m2.
 #n1; #n2; #m1; #m2; #len; #lem; 
-napply transitive_le; (* /2/ slow *)
- ##[ ##| napply monotonic_le_times_l;//; 
    ##| napply monotonic_le_times_r;//;
+napply (transitive_le ? (n1*m2)); (* /2/ slow *)
+ ##[ napply monotonic_le_times_r;//; 
##| napplyS monotonic_le_times_r;//;
  ##]
 nqed.
 
+(* interesssante *)
 ntheorem lt_times_n: ∀n,m:nat. O < n → m ≤ n*m.
-#n; #m; #H; napplyS monotonic_le_times_l;
-/2/; nqed.
+#n; #m; #H; /2/; nqed.
 
 ntheorem le_times_to_le: 
 ∀a,n,m. O < a → a * n ≤ a * m → n ≤ m.
 #a; napply nat_elim2; nnormalize;
   ##[//;
-  ##|#n; #H1; #H2; napply False_ind;
-     ngeneralize in match H2;
-     napply lt_to_not_le;
-     napply (transitive_le ? (S n));/2/;
+  ##|#n; #H1; #H2; 
+     napply (transitive_le ? (a*S n));/2/;
   ##|#n; #m; #H; #lta; #le;
      napply le_S_S; napply H; /2/;
   ##]
 nqed.
 
 ntheorem le_S_times_2: ∀n,m.O < m → n ≤ m → S n ≤ 2*m.
-#n; #m; #posm; #lenm; (* interessante *)
-napplyS (le_plus n); //; nqed.
+#n; #m; #posm; #lenm;  (* interessante *)
+napplyS (le_plus n m); //; nqed.
 
 (* times & lt *)
 (*
@@ -726,20 +749,14 @@ 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;//;
+  ##[/2/; 
   ##|#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.
+/2/; nqed.
 
 ntheorem lt_to_le_to_lt_times: 
 ∀n,m,p,q:nat. n < m → p ≤ q → O < q → n*p < m*q.
@@ -756,19 +773,17 @@ 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;
+∀n,p,q:nat. p*n < q*n → p < q.
+#n; #p; #q; #Hlt;
 nelim (decidable_lt p q);//;
-#nltpq;napply False_ind; 
-napply (lt_to_not_le ? ? Hlt);
-napply monotonic_le_times_l;/3/;
+#nltpq; napply False_ind; 
+napply (absurd ? ? (lt_to_not_le ? ? Hlt));
+napplyS monotonic_le_times_r;/2/;
 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.
+∀n,p,q:nat. n*p < n*q → p < q.
+/2/; nqed.
 
 (*
 theorem nat_compare_times_l : \forall n,p,q:nat. 
@@ -839,14 +854,14 @@ ntheorem minus_n_n: ∀n:nat.O=n-n.
 #n; nelim n; //; nqed.
 
 ntheorem minus_Sn_n: ∀n:nat. S O = (S n)-n.
-#n; nelim n; //; nqed.
+#n; nelim n; nnormalize; //; nqed.
 
 ntheorem minus_Sn_m: ∀m,n:nat. m ≤ n → S n -m = S (n-m).
 (* qualcosa da capire qui 
 #n; #m; #lenm; nelim lenm; napplyS refl_eq. *)
 napply nat_elim2; 
   ##[//
-  ##|#n; #abs; napply False_ind; /2/.
+  ##|#n; #abs; napply False_ind; /2/ 
   ##|#n; #m; #Hind; #c; napplyS Hind; /2/;
   ##]
 nqed.
@@ -859,7 +874,7 @@ 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.
+napply nat_elim2; nnormalize; //; nqed.
 
 ntheorem plus_minus:
 ∀m,n,p:nat. m ≤ n → (n-m)+p = (n+p)-m.
@@ -874,7 +889,7 @@ ntheorem minus_plus_m_m: ∀n,m:nat.n = (n+m)-m.
 #n; #m; napplyS (plus_minus m m n); //; nqed.
 
 ntheorem plus_minus_m_m: ∀n,m:nat.
-m \leq n \to n = (n-m)+m.
+  m ≤ n → n = (n-m)+m.
 #n; #m; #lemn; napplyS symmetric_eq; 
 napplyS (plus_minus m n m); //; nqed.
 
@@ -882,7 +897,7 @@ ntheorem le_plus_minus_m_m: ∀n,m:nat. n ≤ (n-m)+m.
 #n; nelim n;
   ##[//
   ##|#a; #Hind; #m; ncases m;//;  
-     nnormalize; #n;napplyS le_S_S;//  
+     nnormalize; #n;/2/;  
   ##]
 nqed.
 
@@ -901,7 +916,7 @@ nqed.
 ntheorem minus_pred_pred : ∀n,m:nat. O < n → O < m → 
 pred n - pred m = n - m.
 #n; #m; #posn; #posm;
-napply (lt_O_n_elim n posn);
+napply (lt_O_n_elim n posn); 
 napply (lt_O_n_elim m posm);//.
 nqed.
 
@@ -991,6 +1006,7 @@ 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: 
@@ -1033,14 +1049,15 @@ simplify.unfold Not.intro.apply H1.apply inj_S.assumption.
 qed.
 *)
 
-ntheorem eqb_elim : ∀ n,m:nat.∀ P:bool → Prop.
+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; /3/;
   ##|nnormalize; /4/; 
   ##] 
-nqed.
+nqed.*)
 
 ntheorem eqb_n_n: ∀n. eqb n n = true.
 #n; nelim n; nnormalize; //.
@@ -1062,8 +1079,8 @@ ntheorem eq_to_eqb_true: ∀n,m:nat.
 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/;
+napply eqb_elim;//;
+#Heq; napply False_ind; /2/; 
 nqed.
 
 nlet rec leb n m ≝ 
@@ -1078,10 +1095,10 @@ 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/;
+  ##|/3/;
   ##|#n; #m; #Hind; #P; #Pt; #Pf; napply Hind;
     ##[#lenm; napply Pt; napply le_S_S;//;
-    ##|#nlenm; napply Pf; #leSS; /3/;
+    ##|#nlenm; napply Pf; /2/; 
     ##]
   ##]
 nqed.
@@ -1097,7 +1114,7 @@ ntheorem leb_false_to_not_le:∀n,m.
   leb n m = false → n ≰ m.
 #n; #m; napply leb_elim;
   ##[#_; #abs; napply False_ind; /2/;
-  ##|/2/;
+  ##|//;
   ##]
 nqed.
 
@@ -1106,11 +1123,14 @@ ntheorem le_to_leb_true: ∀n,m. n ≤ m → leb n m = true.
 #H; #H1; napply False_ind; /2/;
 nqed.
 
-ntheorem lt_to_leb_false: ∀n,m. m < n → leb n m = false.
+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.