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/;
##]
/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.
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.
/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
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:
##]
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.
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_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.
-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; //;
- ##]
+#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.
-napply nat_elim2;
- ##[#n; ncases n; nnormalize; /2/;
- ##|/2/;
- ##|nnormalize;/2/;
- ##]
+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.
leb n m = false → n ≰ m.
#n; #m; napply leb_elim;
##[#_; #abs; napply False_ind; /2/;
- ##|//;
+ ##|/2/;
##]
nqed.