X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary_auto%2Fnat%2Ffactorial.ma;fp=helm%2Fsoftware%2Fmatita%2Flibrary_auto%2Fnat%2Ffactorial.ma;h=2c77ca5b5f7f41d07ee97597837f6a50788900b4;hb=e61cdf77937ad2ff62af178ae1cb14af0629d0f2;hp=0000000000000000000000000000000000000000;hpb=4989069e8c6910461913fc9a2bae104a0e0259b2;p=helm.git diff --git a/helm/software/matita/library_auto/nat/factorial.ma b/helm/software/matita/library_auto/nat/factorial.ma new file mode 100644 index 000000000..2c77ca5b5 --- /dev/null +++ b/helm/software/matita/library_auto/nat/factorial.ma @@ -0,0 +1,105 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| A.Asperti, C.Sacerdoti Coen, *) +(* ||A|| E.Tassi, S.Zacchiroli *) +(* \ / *) +(* \ / Matita is distributed under the terms of the *) +(* v GNU Lesser General Public License Version 2.1 *) +(* *) +(**************************************************************************) + +set "baseuri" "cic:/matita/nat/factorial". + +include "nat/le_arith.ma". + +let rec fact n \def + match n with + [ O \Rightarrow (S O) + | (S m) \Rightarrow (S m)*(fact m)]. + +interpretation "factorial" 'fact n = (cic:/matita/nat/factorial/fact.con n). + +theorem le_SO_fact : \forall n. (S O) \le n!. +intro. +elim n +[ auto + (*simplify. + apply le_n*) +| change with ((S O) \le (S n1)*n1!). + auto + (*apply (trans_le ? ((S n1)*(S O))) + [ simplify. + apply le_S_S. + apply le_O_n + | apply le_times_r. + assumption + ]*) +] +qed. + +theorem le_SSO_fact : \forall n. (S O) < n \to (S(S O)) \le n!. +intro. +apply (nat_case n) +[ intro. + auto + (*apply False_ind. + apply (not_le_Sn_O (S O) H).*) +| intros. + change with ((S (S O)) \le (S m)*m!). + apply (trans_le ? ((S(S O))*(S O)));auto + (*[ apply le_n + | apply le_times + [ exact H + | apply le_SO_fact + ] + ]*) +] +qed. + +theorem le_n_fact_n: \forall n. n \le n!. +intro. +elim n +[ apply le_O_n +| change with (S n1 \le (S n1)*n1!). + apply (trans_le ? ((S n1)*(S O)));auto + (*[ rewrite < times_n_SO. + apply le_n + | apply le_times. + apply le_n. + apply le_SO_fact + ]*) +] +qed. + +theorem lt_n_fact_n: \forall n. (S(S O)) < n \to n < n!. +intro. +apply (nat_case n) +[ intro. + auto + (*apply False_ind. + apply (not_le_Sn_O (S(S O)) H)*) +| intros. + change with ((S m) < (S m)*m!). + apply (lt_to_le_to_lt ? ((S m)*(S (S O)))) + [ rewrite < sym_times. + simplify. + unfold lt. + apply le_S_S. + auto + (*rewrite < plus_n_O. + apply le_plus_n*) + | apply le_times_r. + auto + (*apply le_SSO_fact. + simplify. + unfold lt. + apply le_S_S_to_le. + exact H*) + ] +] +qed. +