]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/library/nat/factorial2.ma
update in groud_2 and models
[helm.git] / helm / software / matita / library / nat / factorial2.ma
index 594feb51f7ec365fe4ea568f9723c2e33d489c13..bb6c72aebccd7092070199095efb69ea6b279fb3 100644 (file)
@@ -12,8 +12,6 @@
 (*                                                                        *)
 (**************************************************************************)
 
-set "baseuri" "cic:/matita/nat/factorial2".
-
 include "nat/exp.ma".
 include "nat/factorial.ma".
 
@@ -25,11 +23,6 @@ theorem exp_S: \forall n,m. exp m (S n) = m*exp m n.
 intros.reflexivity.
 qed.
 
-alias num (instance 0) = "natural number".
-lemma times_SSO: \forall n.2*(S n) = S(S(2*n)).
-intro.simplify.rewrite < plus_n_Sm.reflexivity.
-qed.
-
 theorem lt_O_to_fact1: \forall n.O<n \to
 fact (2*n) \le (exp 2 (pred (2*n)))*(fact n)*(fact n).
 intros.elim H
@@ -105,7 +98,7 @@ intro.elim n
 qed.
 
 theorem fact2: \forall n.O < n \to
-(exp (S(S O)) ((S(S O))*n))*(fact n)*(fact n) < fact (S((S(S O))*n)).
+(exp 2 (2*n))*(fact n)*(fact n) < fact (S(2*n)).
 intros.elim H
   [simplify.apply le_S.apply le_n
   |rewrite > times_SSO.
@@ -160,7 +153,7 @@ qed.
 
 (* a slightly better result *)
 theorem fact3: \forall n.O < n \to
-(exp (S(S O)) ((S(S O))*n))*(exp (fact n) (S(S O))) \le (S(S O))*n*fact ((S(S O))*n).
+(exp 2 (2*n))*(exp (fact n) 2) \le 2*n*fact (2*n).
 intros.
 elim H
   [simplify.apply le_n