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 (* \ / This file is distributed under the terms of the *)
11 (* v GNU Lesser General Public License Version 2.1 *)
13 (**************************************************************************)
15 set "baseuri" "cic:/matita/nat/binomial".
17 include "nat/iteration2.ma".
18 include "nat/factorial2.ma".
20 definition bc \def \lambda n,k. n!/(k!*(n-k)!).
22 theorem bc_n_n: \forall n. bc n n = S O.
25 simplify in ⊢ (? ? (? ? (? ? %)) ?).
31 theorem bc_n_O: \forall n. bc n O = S O.
34 simplify in ⊢ (? ? (? ? %) ?).
40 theorem fact_minus: \forall n,k. k < n \to
41 (n-k)*(n - S k)! = (n - k)!.
43 [intro.apply False_ind.
44 apply (lt_to_not_le ? ? H).
54 theorem bc2 : \forall n.
55 \forall k. k \leq n \to divides (k!*(n-k)!) n!.
57 [rewrite < (le_n_O_to_eq ? H);apply reflexive_divides
58 |generalize in match H1;cases k
59 [intro;simplify in ⊢ (? (? ? (? %)) ?);simplify in ⊢ (? (? % ?) ?);
60 rewrite > sym_times;rewrite < times_n_SO;apply reflexive_divides
61 |intro;elim (decidable_eq_nat n2 n1)
62 [rewrite > H3;rewrite < minus_n_n;
63 rewrite > times_n_SO in ⊢ (? ? %);apply reflexive_divides
65 [lapply (H (n1 - (S n2)))
66 [change in ⊢ (? ? %) with ((S n1)*n1!);
67 rewrite > (plus_minus_m_m n1 n2) in ⊢ (? ? (? (? %) ?))
69 change in ⊢ (? ? (? % ?)) with ((S n2) + (n1 - n2));
70 rewrite > sym_times in \vdash (? ? %);
71 rewrite > distr_times_plus in ⊢ (? ? %);
72 simplify in ⊢ (? (? ? (? %)) ?);
75 change in ⊢ (? (? ? %) ?) with ((S n2)*n2!);
76 rewrite > sym_times in ⊢ (? (? ? %) ?);
77 rewrite < assoc_times;
79 [rewrite > sym_times;assumption
80 |apply reflexive_divides]
81 |rewrite < fact_minus in ⊢ (? (? ? %) ?)
82 [rewrite > sym_times in ⊢ (? (? ? %) ?);
83 rewrite < assoc_times;
85 [rewrite < eq_plus_minus_minus_minus in Hletin1;
86 [rewrite > sym_times;rewrite < minus_n_n in Hletin1;
87 rewrite < plus_n_O in Hletin1;assumption
88 |lapply (le_S_S_to_le ? ? H2);
89 elim (le_to_or_lt_eq ? ? Hletin2);
93 |apply reflexive_divides]
94 |lapply (le_S_S_to_le ? ? H2);
95 elim (le_to_or_lt_eq ? ? Hletin2);
98 |apply le_S_S_to_le;assumption]
99 |apply le_minus_m;apply le_S_S_to_le;assumption]
100 |apply le_S_S_to_le;assumption]]]]
103 theorem bc1: \forall n.\forall k. k < n \to bc (S n) (S k) = (bc n k) + (bc n (S k)).
105 rewrite > (lt_to_lt_to_eq_div_div_times_times ? ? (S k)) in ⊢ (? ? ? (? % ?))
106 [rewrite > sym_times in ⊢ (? ? ? (? (? ? %) ?)).
107 rewrite < assoc_times in ⊢ (? ? ? (? (? ? %) ?)).
108 rewrite > (lt_to_lt_to_eq_div_div_times_times ? ? (n - k)) in ⊢ (? ? ? (? ? %))
109 [rewrite > assoc_times in ⊢ (? ? ? (? ? (? ? %))).
110 rewrite > sym_times in ⊢ (? ? ? (? ? (? ? (? ? %)))).
112 [rewrite < eq_div_plus
113 [rewrite < distr_times_plus.
114 simplify in ⊢ (? ? ? (? (? ? %) ?)).
115 rewrite > sym_plus in ⊢ (? ? ? (? (? ? (? %)) ?)).
116 rewrite < plus_minus_m_m
117 [rewrite > sym_times in ⊢ (? ? ? (? % ?)).
122 |rewrite > (times_n_O O).
123 apply lt_times;apply lt_O_fact
124 |change in ⊢ (? (? % ?) ?) with ((S k)*k!);
125 rewrite < sym_times in ⊢ (? ? %);
126 rewrite > assoc_times;apply divides_times
127 [apply reflexive_divides
128 |apply bc2;apply le_S_S_to_le;constructor 2;assumption]
129 |rewrite < fact_minus
130 [rewrite > sym_times in ⊢ (? (? ? %) ?);rewrite < assoc_times;
132 [apply bc2;assumption
133 |apply reflexive_divides]
136 |apply lt_to_lt_O_minus;assumption
137 |rewrite > (times_n_O O).
138 apply lt_times;apply lt_O_fact]
140 |rewrite > (times_n_O O).
141 apply lt_times;apply lt_O_fact]
144 theorem lt_O_bc: \forall n,m. m \le n \to O < bc n m.
146 [apply (le_n_O_elim ? H).
148 |elim (le_to_or_lt_eq ? ? H1)
149 [generalize in match H2.cases m;intro
150 [rewrite > bc_n_O.apply le_n
152 [apply (trans_le ? (bc n1 n2))
153 [apply H.apply le_S_S_to_le.apply lt_to_le.assumption
156 |apply le_S_S_to_le.assumption
166 theorem exp_plus_sigma_p:\forall a,b,n.
167 exp (a+b) n = sigma_p (S n) (\lambda k.true)
168 (\lambda k.(bc n k)*(exp a (n-k))*(exp b k)).
170 [simplify.reflexivity
171 |simplify in ⊢ (? ? % ?).
172 rewrite > true_to_sigma_p_Sn
173 [rewrite > H;rewrite > sym_times in ⊢ (? ? % ?);
174 rewrite > distr_times_plus in ⊢ (? ? % ?);
175 rewrite < minus_n_n in ⊢ (? ? ? (? (? (? ? (? ? %)) ?) ?));
176 rewrite > sym_times in ⊢ (? ? (? % ?) ?);
177 rewrite > sym_times in ⊢ (? ? (? ? %) ?);
178 rewrite > distributive_times_plus_sigma_p in ⊢ (? ? (? % ?) ?);
179 rewrite > distributive_times_plus_sigma_p in ⊢ (? ? (? ? %) ?);
180 rewrite > true_to_sigma_p_Sn in ⊢ (? ? (? ? %) ?)
181 [rewrite < assoc_plus;rewrite < sym_plus in ⊢ (? ? (? % ?) ?);
182 rewrite > assoc_plus;
184 [rewrite > sym_times;rewrite > assoc_times;autobatch paramodulation
185 |rewrite > (sigma_p_gi ? ? O);
186 [rewrite < (eq_sigma_p_gh ? S pred n1 (S n1) (λx:nat.true)) in ⊢ (? ? (? (? ? %) ?) ?)
187 [rewrite > (sigma_p_gi ? ? O) in ⊢ (? ? ? %);
188 [rewrite > assoc_plus;apply eq_f2
189 [autobatch paramodulation;
190 |rewrite < sigma_p_plus_1;
191 rewrite < (eq_sigma_p_gh ? S pred n1 (S n1) (λx:nat.true)) in \vdash (? ? ? %);
194 |intros;rewrite > sym_times;rewrite > assoc_times;
195 rewrite > sym_times in ⊢ (? ? (? (? ? %) ?) ?);
196 rewrite > assoc_times;rewrite < assoc_times in ⊢ (? ? (? (? ? %) ?) ?);
197 rewrite > sym_times in ⊢ (? ? (? (? ? (? % ?)) ?) ?);
198 change in ⊢ (? ? (? (? ? (? % ?)) ?) ?) with (exp a (S (n1 - S x)));
199 rewrite < (minus_Sn_m ? ? H1);rewrite > minus_S_S;
200 rewrite > sym_times in \vdash (? ? (? ? %) ?);
201 rewrite > assoc_times;
202 rewrite > sym_times in ⊢ (? ? (? ? (? ? %)) ?);
203 change in \vdash (? ? (? ? (? ? %)) ?) with (exp b (S x));
204 rewrite > assoc_times in \vdash (? ? (? ? %) ?);
205 rewrite > sym_times in \vdash (? ? (? % ?) ?);
206 rewrite > sym_times in \vdash (? ? (? ? %) ?);
207 rewrite > assoc_times in \vdash (? ? ? %);
208 rewrite > sym_times in \vdash (? ? ? %);
209 rewrite < distr_times_plus;
210 rewrite > sym_plus;rewrite < bc1;
211 [reflexivity|assumption]]
212 |intros;simplify;reflexivity
213 |intros;simplify;reflexivity
214 |intros;apply le_S_S;assumption
217 [simplify in H2;destruct H2
218 |simplify;reflexivity]
220 [simplify in H2;destruct H2
221 |simplify;apply le_S_S_to_le;assumption]]]
222 |apply le_S_S;apply le_O_n
224 |intros;simplify;reflexivity
225 |intros;simplify;reflexivity
226 |intros;apply le_S_S;assumption
229 [simplify in H2;destruct H2
230 |simplify;reflexivity]
232 [simplify in H2;destruct H2
233 |simplify;apply le_S_S_to_le;assumption]]
234 |apply le_S_S;apply le_O_n
240 theorem exp_S_sigma_p:\forall a,n.
241 exp (S a) n = sigma_p (S n) (\lambda k.true) (\lambda k.(bc n k)*(exp a (n-k))).
244 rewrite > exp_plus_sigma_p.
245 apply eq_sigma_p;intros
248 rewrite < times_n_SO.
253 theorem exp_Sn_SSO: \forall n. exp (S n) 2 = S((exp n 2) + 2*n).
255 rewrite < times_n_SO.
257 rewrite < sym_times.simplify.
258 rewrite < assoc_plus.