1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| A.Asperti, C.Sacerdoti Coen, *)
8 (* ||A|| E.Tassi, S.Zacchiroli *)
10 (* \ / Matita is distributed under the terms of the *)
11 (* v GNU Lesser General Public License Version 2.1 *)
13 (**************************************************************************)
15 include "nat/factorial.ma".
17 include "nat/lt_arith.ma".
19 let rec sigma n f m \def
22 | (S p) \Rightarrow (f (S p+m))+(sigma p f m)].
27 | (S p) \Rightarrow (f (S p+m))*(pi p f m)].
29 theorem eq_sigma: \forall f,g:nat \to nat.
31 (\forall i:nat. m \le i \to i \le m+n \to f i = g i) \to
32 (sigma n f m) = (sigma n g m).
34 simplify.apply H.apply le_n.rewrite < plus_n_O.apply le_n.
37 change with (m \le (S n1)+m).apply le_plus_n.
38 rewrite > (sym_plus m).apply le_n.
39 apply H.intros.apply H1.assumption.
41 apply le_S.assumption.
44 theorem eq_pi: \forall f,g:nat \to nat.
46 (\forall i:nat. m \le i \to i \le m+n \to f i = g i) \to
47 (pi n f m) = (pi n g m).
49 simplify.apply H.apply le_n.rewrite < plus_n_O.apply le_n.
52 change with (m \le (S n1)+m).apply le_plus_n.
53 rewrite > (sym_plus m).apply le_n.
54 apply H.intros.apply H1.assumption.
56 apply le_S.assumption.
59 theorem eq_fact_pi: \forall n. (S n)! = pi n (\lambda m.m) (S O).
62 change with ((S(S n1))*(S n1)! = ((S n1)+(S O))*(pi n1 (\lambda m.m) (S O))).
63 rewrite < plus_n_Sm.rewrite < plus_n_O.
64 apply eq_f.assumption.
67 theorem exp_pi_l: \forall f:nat\to nat.\forall n,m,a:nat.
68 (exp a (S n))*pi n f m= pi n (\lambda p.a*(f p)) m.
69 intros.elim n.simplify.rewrite < times_n_SO.reflexivity.
72 rewrite > assoc_times.
73 rewrite > assoc_times in\vdash (? ? ? %).
74 apply eq_f.rewrite < assoc_times.
75 rewrite < assoc_times.
76 apply eq_f2.apply sym_times.reflexivity.