]> matita.cs.unibo.it Git - helm.git/blob - matita/library/nat/pi_p.ma
Towards chebyshev.
[helm.git] / matita / library / nat / pi_p.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||       A.Asperti, C.Sacerdoti Coen,                          *)
8 (*      ||A||       E.Tassi, S.Zacchiroli                                 *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU Lesser General Public License Version 2.1         *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 set "baseuri" "cic:/matita/nat/pi_p".
16
17 include "nat/primes.ma".
18 (* include "nat/ord.ma". *)
19 include "nat/generic_iter_p.ma".
20 (* include "nat/count.ma". necessary just to use bool_to_nat and bool_to_nat_andb*)
21 include "nat/iteration2.ma".
22
23 (* pi_p on nautral numbers is a specialization of iter_p_gen *)
24 definition pi_p: nat \to (nat \to bool) \to (nat \to nat) \to nat \def
25 \lambda n, p, g. (iter_p_gen n p nat g (S O) times).
26
27 theorem true_to_pi_p_Sn: 
28 \forall n:nat. \forall p:nat \to bool. \forall g:nat \to nat.
29 p n = true \to pi_p (S n) p g = (g n)*(pi_p n p g).
30 intros.
31 unfold pi_p.
32 apply true_to_iter_p_gen_Sn.
33 assumption.
34 qed.
35    
36 theorem false_to_pi_p_Sn: 
37 \forall n:nat. \forall p:nat \to bool. \forall g:nat \to nat.
38 p n = false \to pi_p (S n) p g = pi_p n p g.
39 intros.
40 unfold pi_p.
41 apply false_to_iter_p_gen_Sn.
42 assumption.
43 qed.  
44
45 theorem eq_pi_p: \forall p1,p2:nat \to bool.
46 \forall g1,g2: nat \to nat.\forall n.
47 (\forall x.  x < n \to p1 x = p2 x) \to
48 (\forall x.  x < n \to g1 x = g2 x) \to
49 pi_p n p1 g1 = pi_p n p2 g2.
50 intros.
51 unfold pi_p.
52 apply eq_iter_p_gen;
53 assumption.
54 qed.
55
56 theorem eq_pi_p1: \forall p1,p2:nat \to bool.
57 \forall g1,g2: nat \to nat.\forall n.
58 (\forall x.  x < n \to p1 x = p2 x) \to
59 (\forall x.  x < n \to p1 x = true \to g1 x = g2 x) \to
60 pi_p n p1 g1 = pi_p n p2 g2.
61 intros.
62 unfold pi_p.
63 apply eq_iter_p_gen1;
64 assumption.
65 qed.
66
67 theorem pi_p_false: 
68 \forall g: nat \to nat.\forall n.pi_p n (\lambda x.false) g = S O.
69 intros.
70 unfold pi_p.
71 apply iter_p_gen_false.
72 qed.
73
74 theorem pi_p_times: \forall n,k:nat.\forall p:nat \to bool.
75 \forall g: nat \to nat.
76 pi_p (k+n) p g 
77 = pi_p k (\lambda x.p (x+n)) (\lambda x.g (x+n)) * pi_p n p g.
78 intros.
79 unfold pi_p.
80 apply (iter_p_gen_plusA nat n k p g (S O) times)
81 [ apply sym_times.
82 | intros.
83   apply sym_eq.
84   apply times_n_SO
85 | apply associative_times
86 ]
87 qed.
88
89 theorem false_to_eq_pi_p: \forall n,m:nat.n \le m \to
90 \forall p:nat \to bool.
91 \forall g: nat \to nat. (\forall i:nat. n \le i \to i < m \to
92 p i = false) \to pi_p m p g = pi_p n p g.
93 intros.
94 unfold pi_p.
95 apply (false_to_eq_iter_p_gen);
96 assumption.
97 qed.
98
99 theorem or_false_eq_SO_to_eq_pi_p: 
100 \forall n,m:nat.\forall p:nat \to bool.
101 \forall g: nat \to nat.
102 n \le m \to (\forall i:nat. n \le i \to i < m \to p i = false \lor g i = S O)
103 \to pi_p m p g = pi_p n p g.
104 intros.
105 unfold pi_p.
106 apply or_false_eq_baseA_to_eq_iter_p_gen
107   [intros.simplify.rewrite < plus_n_O.reflexivity
108   |assumption
109   |assumption
110   ]
111 qed.
112
113 theorem pi_p2 : 
114 \forall n,m:nat.
115 \forall p1,p2:nat \to bool.
116 \forall g: nat \to nat \to nat.
117 pi_p (n*m) 
118   (\lambda x.andb (p1 (div x m)) (p2 (mod x m))) 
119   (\lambda x.g (div x m) (mod x m)) =
120 pi_p n p1 
121   (\lambda x.pi_p m p2 (g x)).
122 intros.
123 unfold pi_p.
124 apply (iter_p_gen2 n m p1 p2 nat g (S O) times)
125 [ apply sym_times
126 | apply associative_times
127 | intros.
128   apply sym_eq.
129   apply times_n_SO
130 ]
131 qed.
132
133 theorem pi_p2' : 
134 \forall n,m:nat.
135 \forall p1:nat \to bool.
136 \forall p2:nat \to nat \to bool.
137 \forall g: nat \to nat \to nat.
138 pi_p (n*m) 
139   (\lambda x.andb (p1 (div x m)) (p2 (div x m) (mod x  m))) 
140   (\lambda x.g (div x m) (mod x m)) =
141 pi_p n p1 
142   (\lambda x.pi_p m (p2 x) (g x)).
143 intros.
144 unfold pi_p.
145 apply (iter_p_gen2' n m p1 p2 nat g (S O) times)
146 [ apply sym_times
147 | apply associative_times
148 | intros.
149   apply sym_eq.
150   apply times_n_SO
151 ]
152 qed.
153
154 lemma pi_p_gi: \forall g: nat \to nat.
155 \forall n,i.\forall p:nat \to bool.i < n \to p i = true \to 
156 pi_p n p g = g i * pi_p n (\lambda x. andb (p x) (notb (eqb x i))) g.
157 intros.
158 unfold pi_p.
159 apply (iter_p_gen_gi)
160 [ apply sym_times
161 | apply associative_times
162 | intros.
163   apply sym_eq.
164   apply times_n_SO
165 | assumption
166 | assumption
167 ]
168 qed.
169
170 theorem eq_pi_p_gh: 
171 \forall g,h,h1: nat \to nat.\forall n,n1.
172 \forall p1,p2:nat \to bool.
173 (\forall i. i < n \to p1 i = true \to p2 (h i) = true) \to
174 (\forall i. i < n \to p1 i = true \to h1 (h i) = i) \to 
175 (\forall i. i < n \to p1 i = true \to h i < n1) \to 
176 (\forall j. j < n1 \to p2 j = true \to p1 (h1 j) = true) \to
177 (\forall j. j < n1 \to p2 j = true \to h (h1 j) = j) \to 
178 (\forall j. j < n1 \to p2 j = true \to h1 j < n) \to 
179 pi_p n p1 (\lambda x.g(h x)) = pi_p n1 p2 g.
180 intros.
181 unfold pi_p.
182 apply (eq_iter_p_gen_gh nat (S O) times ? ? ? g h h1 n n1 p1 p2)
183 [ apply sym_times
184 | apply associative_times
185 | intros.
186   apply sym_eq.
187   apply times_n_SO
188 | assumption
189 | assumption
190 | assumption
191 | assumption
192 | assumption
193 | assumption
194 ]
195 qed.
196
197 (* monotonicity *)
198 theorem le_pi_p: 
199 \forall n:nat. \forall p:nat \to bool. \forall g1,g2:nat \to nat.
200 (\forall i. i < n \to p i = true \to g1 i \le g2 i ) \to 
201 pi_p n p g1 \le pi_p n p g2.
202 intros.
203 generalize in match H.
204 elim n
205   [apply le_n.
206   |apply (bool_elim ? (p n1));intros
207     [rewrite > true_to_pi_p_Sn
208       [rewrite > true_to_pi_p_Sn in ⊢ (? ? %)
209         [apply le_times
210           [apply H2[apply le_n|assumption]
211           |apply H1.
212            intros.
213            apply H2[apply le_S.assumption|assumption]
214           ]
215         |assumption
216         ]
217       |assumption
218       ]
219     |rewrite > false_to_pi_p_Sn
220       [rewrite > false_to_pi_p_Sn in ⊢ (? ? %)
221         [apply H1.
222          intros.
223          apply H2[apply le_S.assumption|assumption]
224         |assumption
225         ]
226       |assumption
227       ]
228     ]
229   ]
230 qed.
231      
232 theorem exp_sigma_p: \forall n,a,p. 
233 pi_p n p (\lambda x.a) = (exp a (sigma_p n p (\lambda x.S O))).
234 intros.
235 elim n
236   [reflexivity
237   |apply (bool_elim ? (p n1))
238     [intro.
239      rewrite > true_to_pi_p_Sn
240       [rewrite > true_to_sigma_p_Sn
241         [simplify.
242          rewrite > H.
243          reflexivity.
244         |assumption
245         ]
246       |assumption
247       ]
248     |intro.
249      rewrite > false_to_pi_p_Sn
250       [rewrite > false_to_sigma_p_Sn
251         [simplify.assumption
252         |assumption
253         ]
254       |assumption
255       ]
256     ]
257   ]
258 qed.
259
260 theorem times_pi_p: \forall n,p,f,g. 
261 pi_p n p (\lambda x.f x*g x) = pi_p n p f * pi_p n p  g. 
262 intros.
263 elim n
264   [simplify.reflexivity
265   |apply (bool_elim ? (p n1))
266     [intro.
267      rewrite > true_to_pi_p_Sn
268       [rewrite > true_to_pi_p_Sn
269         [rewrite > true_to_pi_p_Sn
270           [rewrite > H.autobatch
271           |assumption
272           ]
273         |assumption
274         ]
275       |assumption
276       ]
277     |intro.
278      rewrite > false_to_pi_p_Sn
279       [rewrite > false_to_pi_p_Sn
280         [rewrite > false_to_pi_p_Sn;assumption
281         |assumption
282         ]
283       |assumption
284       ]
285     ]
286   ]
287 qed.
288
289 theorem pi_p_SO: \forall n,p. 
290 pi_p n p (\lambda i.S O) = S O.
291 intros.elim n
292   [reflexivity
293   |simplify.elim (p n1)
294     [simplify.rewrite < plus_n_O.assumption
295     |simplify.assumption
296     ]
297   ]
298 qed.
299
300 theorem exp_pi_p: \forall n,m,p,f. 
301 pi_p n p (\lambda x.exp (f x) m) = exp (pi_p n p f) m.
302 intros.
303 elim m
304   [simplify.apply pi_p_SO
305   |simplify.
306    rewrite > times_pi_p.
307    rewrite < H.
308    reflexivity
309   ]
310 qed.
311
312 theorem exp_times_pi_p: \forall n,m,k,p,f. 
313 pi_p n p (\lambda x.exp k (m*(f x))) = 
314 exp (pi_p n p (\lambda x.exp k (f x))) m.
315 intros.
316 apply (trans_eq ? ? (pi_p n p (\lambda x.(exp (exp k (f x)) m))))
317   [apply eq_pi_p;intros
318     [reflexivity
319     |apply sym_eq.rewrite > sym_times.
320      apply exp_exp_times
321     ]
322   |apply exp_pi_p
323   ]
324 qed.
325
326
327 theorem pi_p_knm:
328 \forall g: nat \to nat.
329 \forall h2:nat \to nat \to nat.
330 \forall h11,h12:nat \to nat. 
331 \forall k,n,m.
332 \forall p1,p21:nat \to bool.
333 \forall p22:nat \to nat \to bool.
334 (\forall x. x < k \to p1 x = true \to 
335 p21 (h11 x) = true \land p22 (h11 x) (h12 x) = true
336 \land h2 (h11 x) (h12 x) = x 
337 \land (h11 x) < n \land (h12 x) < m) \to
338 (\forall i,j. i < n \to j < m \to p21 i = true \to p22 i j = true \to 
339 p1 (h2 i j) = true \land 
340 h11 (h2 i j) = i \land h12 (h2 i j) = j
341 \land h2 i j < k) \to
342 pi_p k p1 g =
343 pi_p n p21 (\lambda x:nat.pi_p m (p22 x) (\lambda y. g (h2 x y))).
344 intros.
345 unfold pi_p.unfold pi_p.
346 apply (iter_p_gen_knm nat (S O) times sym_times assoc_times ? ? ? h11 h12)
347   [intros.apply sym_eq.apply times_n_SO.
348   |assumption
349   |assumption
350   ]
351 qed.
352
353 theorem pi_p_pi_p: 
354 \forall g: nat \to nat \to nat.
355 \forall h11,h12,h21,h22: nat \to nat \to nat. 
356 \forall n1,m1,n2,m2.
357 \forall p11,p21:nat \to bool.
358 \forall p12,p22:nat \to nat \to bool.
359 (\forall i,j. i < n2 \to j < m2 \to p21 i = true \to p22 i j = true \to 
360 p11 (h11 i j) = true \land p12 (h11 i j) (h12 i j) = true
361 \land h21 (h11 i j) (h12 i j) = i \land h22 (h11 i j) (h12 i j) = j
362 \land h11 i j < n1 \land h12 i j < m1) \to
363 (\forall i,j. i < n1 \to j < m1 \to p11 i = true \to p12 i j = true \to 
364 p21 (h21 i j) = true \land p22 (h21 i j) (h22 i j) = true
365 \land h11 (h21 i j) (h22 i j) = i \land h12 (h21 i j) (h22 i j) = j
366 \land (h21 i j) < n2 \land (h22 i j) < m2) \to
367 pi_p n1 p11 
368      (\lambda x:nat .pi_p m1 (p12 x) (\lambda y. g x y)) =
369 pi_p n2 p21 
370     (\lambda x:nat .pi_p m2 (p22 x)  (\lambda y. g (h11 x y) (h12 x y))).
371 intros.
372 unfold pi_p.unfold pi_p.
373 apply (iter_p_gen_2_eq ? ? ? sym_times assoc_times ? ? ? ? h21 h22)
374   [intros.apply sym_eq.apply times_n_SO.
375   |assumption
376   |assumption
377   ]
378 qed.
379
380 theorem pi_p_pi_p1: 
381 \forall g: nat \to nat \to nat.
382 \forall n,m.
383 \forall p11,p21:nat \to bool.
384 \forall p12,p22:nat \to nat \to bool.
385 (\forall x,y. x < n \to y < m \to 
386  (p11 x \land p12 x y) = (p21 y \land p22 y x)) \to
387 pi_p n p11 (\lambda x:nat.pi_p m (p12 x) (\lambda y. g x y)) =
388 pi_p m p21 (\lambda y:nat.pi_p n (p22 y) (\lambda x. g x y)).
389 intros.
390 unfold pi_p.unfold pi_p.
391 apply (iter_p_gen_iter_p_gen ? ? ? sym_times assoc_times)
392   [intros.apply sym_eq.apply times_n_SO.
393   |assumption
394   ]
395 qed.