]> matita.cs.unibo.it Git - helm.git/blob - matita/library/nat/ord.ma
matita 0.5.1 tagged
[helm.git] / matita / library / nat / ord.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 (*       \ /        Matita is distributed under the terms of the          *)
11 (*        v         GNU Lesser General Public License Version 2.1         *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 include "datatypes/constructors.ma".
16 include "nat/exp.ma".
17 include "nat/nth_prime.ma".
18 include "nat/gcd.ma". (* for some properties of divides *)
19 include "nat/relevant_equations.ma". (* required by autobatch paramod *)
20
21 let rec p_ord_aux p n m \def
22   match n \mod m with
23   [ O \Rightarrow 
24     match p with
25       [ O \Rightarrow pair nat nat O n
26       | (S p) \Rightarrow 
27        match (p_ord_aux p (n / m) m) with
28        [ (pair q r) \Rightarrow pair nat nat (S q) r] ]
29   | (S a) \Rightarrow pair nat nat O n].
30   
31 (* p_ord n m = <q,r> if m divides n q times, with remainder r *)
32 definition p_ord \def \lambda n,m:nat.p_ord_aux n n m.
33
34 theorem p_ord_aux_to_Prop: \forall p,n,m. O < m \to
35   match p_ord_aux p n m with
36   [ (pair q r) \Rightarrow n = m \sup q *r ].
37 intro.
38 elim p.simplify.
39 apply (nat_case (n \mod m)).
40 simplify.apply plus_n_O.
41 intros.
42 simplify.apply plus_n_O.
43 simplify. 
44 apply (nat_case1 (n1 \mod m)).intro.
45 simplify.
46 generalize in match (H (n1 / m) m).
47 elim (p_ord_aux n (n1 / m) m).
48 simplify.
49 rewrite > assoc_times.
50 rewrite < H3.rewrite > (plus_n_O (m*(n1 / m))).
51 rewrite < H2.
52 rewrite > sym_times.
53 rewrite < div_mod.reflexivity.
54 assumption.assumption.
55 intros.simplify.apply plus_n_O. 
56 qed.
57
58 theorem p_ord_aux_to_exp: \forall p,n,m,q,r. O < m \to
59   (pair nat nat q r) = p_ord_aux p n m \to n = m \sup q * r.
60 intros.
61 change with 
62 match (pair nat nat q r) with
63   [ (pair q r) \Rightarrow n = m \sup q * r ].
64 rewrite > H1.
65 apply p_ord_aux_to_Prop.
66 assumption.
67 qed.
68
69 (* questo va spostato in primes1.ma *)
70 theorem p_ord_exp: \forall n,m,i. O < m \to n \mod m \neq O \to
71 \forall p. i \le p \to p_ord_aux p (m \sup i * n) m = pair nat nat i n.
72 intros 5.
73 elim i.
74 simplify.
75 rewrite < plus_n_O.
76 apply (nat_case p).
77 simplify.
78 elim (n \mod m).simplify.reflexivity.simplify.reflexivity.
79 intro.
80 simplify.
81 cut (O < n \mod m \lor O = n \mod m).
82 elim Hcut.apply (lt_O_n_elim (n \mod m) H3).
83 intros. simplify.reflexivity.
84 apply False_ind.
85 apply H1.apply sym_eq.assumption.
86 apply le_to_or_lt_eq.apply le_O_n.
87 generalize in match H3.
88 apply (nat_case p).intro.apply False_ind.apply (not_le_Sn_O n1 H4).
89 intros.
90 simplify. fold simplify (m \sup (S n1)).
91 cut (((m \sup (S n1)*n) \mod m) = O).
92 rewrite > Hcut.
93 simplify.fold simplify (m \sup (S n1)).
94 cut ((m \sup (S n1) *n) / m = m \sup n1 *n).
95 rewrite > Hcut1.
96 rewrite > (H2 m1). simplify.reflexivity.
97 apply le_S_S_to_le.assumption.
98 (* div_exp *)
99 simplify.
100 rewrite > assoc_times.
101 apply (lt_O_n_elim m H).
102 intro.apply div_times.
103 (* mod_exp = O *)
104 apply divides_to_mod_O.
105 assumption.
106 simplify.rewrite > assoc_times.
107 apply (witness ? ? (m \sup n1 *n)).reflexivity.
108 qed.
109
110 theorem p_ord_aux_to_Prop1: \forall p,n,m. (S O) < m \to O < n \to n \le p \to
111   match p_ord_aux p n m with
112   [ (pair q r) \Rightarrow r \mod m \neq O].
113 intro.elim p.absurd (O < n).assumption.
114 apply le_to_not_lt.assumption.
115 simplify.
116 apply (nat_case1 (n1 \mod m)).intro.
117 generalize in match (H (n1 / m) m).
118 elim (p_ord_aux n (n1 / m) m).
119 apply H5.assumption.
120 apply eq_mod_O_to_lt_O_div.
121 apply (trans_lt ? (S O)).unfold lt.apply le_n.
122 assumption.assumption.assumption.
123 apply le_S_S_to_le.
124 apply (trans_le ? n1).change with (n1 / m < n1).
125 apply lt_div_n_m_n.assumption.assumption.assumption.
126 intros.simplify.
127 rewrite > H4.
128 unfold Not.intro.
129 apply (not_eq_O_S m1).
130 rewrite > H5.reflexivity.
131 qed.
132
133 theorem p_ord_aux_to_not_mod_O: \forall p,n,m,q,r. (S O) < m \to O < n \to n \le p \to
134  pair nat nat q r = p_ord_aux p n m \to r \mod m \neq O.
135 intros.
136 change with 
137   match (pair nat nat q r) with
138   [ (pair q r) \Rightarrow r \mod m \neq O].
139 rewrite > H3. 
140 apply p_ord_aux_to_Prop1.
141 assumption.assumption.assumption.
142 qed.
143
144 theorem p_ord_exp1: \forall p,n,q,r. O < p \to \lnot p \divides r \to
145 n = p \sup q * r \to p_ord n p = pair nat nat q r.
146 intros.unfold p_ord.
147 rewrite > H2.
148 apply p_ord_exp
149  [assumption
150  |unfold.intro.apply H1.
151   apply mod_O_to_divides[assumption|assumption]
152  |apply (trans_le ? (p \sup q)).
153   cut ((S O) \lt p).
154   elim q.simplify.apply le_n_Sn.
155   simplify.
156   generalize in match H3.
157   apply (nat_case n1).simplify.
158   rewrite < times_n_SO.intro.assumption.
159   intros.
160   apply (trans_le ? (p*(S m))).
161   apply (trans_le ? ((S (S O))*(S m))).
162   simplify.rewrite > plus_n_Sm.
163   rewrite < plus_n_O.
164   apply le_plus_n.
165   apply le_times_l.
166   assumption.
167   apply le_times_r.assumption.
168   apply not_eq_to_le_to_lt.
169   unfold.intro.apply H1.
170   rewrite < H3.
171   apply (witness ? r r ?).simplify.apply plus_n_O.
172   assumption.
173   rewrite > times_n_SO in \vdash (? % ?).
174   apply le_times_r.
175   change with (O \lt r).
176   apply not_eq_to_le_to_lt.
177   unfold.intro.
178   apply H1.rewrite < H3.
179   apply (witness ? ? O ?).rewrite < times_n_O.reflexivity.
180   apply le_O_n.
181   ]
182 qed.
183
184 theorem p_ord_to_exp1: \forall p,n,q,r. (S O) \lt p \to O \lt n \to p_ord n p = pair nat nat q r\to 
185 \lnot p \divides r \land n = p \sup q * r.
186 intros.
187 unfold p_ord in H2.
188 split.unfold.intro.
189 apply (p_ord_aux_to_not_mod_O n n p q r).assumption.assumption.
190 apply le_n.symmetry.assumption.
191 apply divides_to_mod_O.apply (trans_lt ? (S O)).
192 unfold.apply le_n.assumption.assumption.
193 apply (p_ord_aux_to_exp n).apply (trans_lt ? (S O)).
194 unfold.apply le_n.assumption.symmetry.assumption.
195 qed.
196
197 theorem p_ord_times: \forall p,a,b,qa,ra,qb,rb. prime p 
198 \to O \lt a \to O \lt b 
199 \to p_ord a p = pair nat nat qa ra  
200 \to p_ord b p = pair nat nat qb rb
201 \to p_ord (a*b) p = pair nat nat (qa + qb) (ra*rb).
202 intros.
203 cut ((S O) \lt p).
204 elim (p_ord_to_exp1 ? ? ? ? Hcut H1 H3).
205 elim (p_ord_to_exp1 ? ? ? ? Hcut H2 H4).
206 apply p_ord_exp1.
207 apply (trans_lt ? (S O)).unfold.apply le_n.assumption.
208 unfold.intro.
209 elim (divides_times_to_divides ? ? ? H H9).
210 apply (absurd ? ? H10 H5).
211 apply (absurd ? ? H10 H7).
212 (* rewrite > H6.
213 rewrite > H8. *)
214 autobatch paramodulation.
215 unfold prime in H. elim H. assumption.
216 qed.
217
218 theorem fst_p_ord_times: \forall p,a,b. prime p 
219 \to O \lt a \to O \lt b 
220 \to fst ? ? (p_ord (a*b) p) = (fst ? ? (p_ord a p)) + (fst ? ? (p_ord b p)).
221 intros.
222 rewrite > (p_ord_times p a b (fst ? ? (p_ord a p)) (snd ? ? (p_ord a p))
223 (fst ? ? (p_ord b p)) (snd ? ? (p_ord b p)) H H1 H2).
224 simplify.reflexivity.
225 apply eq_pair_fst_snd.
226 apply eq_pair_fst_snd.
227 qed.
228
229 theorem p_ord_p : \forall p:nat. (S O) \lt p \to p_ord p p = pair ? ? (S O) (S O).
230 intros.
231 apply p_ord_exp1.
232 apply (trans_lt ? (S O)). unfold.apply le_n.assumption.
233 unfold.intro.
234 apply (absurd ? ? H).
235 apply le_to_not_lt.
236 apply divides_to_le.unfold.apply le_n.assumption.
237 rewrite < times_n_SO.
238 apply exp_n_SO.
239 qed.
240
241 (* p_ord and divides *)
242 theorem divides_to_p_ord: \forall p,a,b,c,d,n,m:nat. 
243 O < n \to O < m \to prime p 
244 \to divides n m \to p_ord n p = pair ? ? a b \to
245 p_ord m p = pair ? ? c d \to divides b d \land a \le c.
246 intros.
247 cut (S O < p)
248   [lapply (p_ord_to_exp1 ? ? ? ? Hcut H H4).
249    lapply (p_ord_to_exp1 ? ? ? ? Hcut H1 H5).
250    elim Hletin. clear Hletin.
251    elim Hletin1. clear Hletin1.
252    rewrite > H9 in H3.
253    split
254     [apply (gcd_SO_to_divides_times_to_divides (exp p c))
255       [elim (le_to_or_lt_eq ? ? (le_O_n b))
256         [assumption
257         |apply False_ind.
258          apply (lt_to_not_eq O ? H).
259          rewrite > H7.
260          rewrite < H10.
261          autobatch
262         ]
263       |elim c
264         [rewrite > sym_gcd.
265          apply gcd_SO_n
266         |simplify.
267          apply eq_gcd_times_SO
268           [apply lt_to_le.assumption
269           |apply lt_O_exp.apply lt_to_le.assumption
270           |rewrite > sym_gcd.
271            (* hint non trova prime_to_gcd_SO e
272               autobatch non chiude il goal *)
273            apply prime_to_gcd_SO
274             [assumption|assumption]
275           |assumption
276           ]
277         ]
278       |apply (trans_divides ? n)
279         [apply (witness ? ? (exp p a)).
280          rewrite > sym_times.
281          assumption
282         |assumption
283         ]
284       ]
285     |apply (le_exp_to_le p)
286       [assumption
287       |apply divides_to_le
288         [apply lt_O_exp.apply lt_to_le.assumption
289         |apply (gcd_SO_to_divides_times_to_divides d)
290           [apply lt_O_exp.apply lt_to_le.assumption
291           |elim a
292             [apply gcd_SO_n
293             |simplify.rewrite < sym_gcd.
294              apply eq_gcd_times_SO
295               [apply lt_to_le.assumption
296               |apply lt_O_exp.apply lt_to_le.assumption
297               |rewrite > sym_gcd.
298               (* hint non trova prime_to_gcd_SO e
299                  autobatch non chiude il goal *)
300                apply prime_to_gcd_SO
301                 [assumption|assumption]
302               |rewrite > sym_gcd. assumption
303               ]
304             ]
305           |apply (trans_divides ? n)
306             [apply (witness ? ? b).assumption
307             |rewrite > sym_times.assumption
308             ]
309           ]
310         ]
311       ]
312     ]
313   |elim H2.assumption
314   ]
315 qed.    
316
317 (* p_ord and primes *)
318 theorem not_divides_to_p_ord_O: \forall n,i.
319 Not (divides (nth_prime i) n) \to p_ord n (nth_prime i) = 
320 pair nat nat O n.
321 intros.
322 apply p_ord_exp1
323   [apply lt_O_nth_prime_n
324   |assumption
325   |autobatch
326   ]
327 qed.
328
329 theorem p_ord_O_to_not_divides: \forall n,i,r.
330 O < n \to
331 p_ord n (nth_prime i) = pair nat nat O r 
332 \to Not (divides (nth_prime i) n).
333 intros.
334 lapply (p_ord_to_exp1 ? ? ? ? ? ? H1)
335   [apply lt_SO_nth_prime_n
336   |assumption
337   |elim Hletin.
338    simplify in H3.
339    rewrite > H3.
340    rewrite < plus_n_O.
341    assumption
342   ]
343 qed.
344
345 theorem p_ord_to_not_eq_O : \forall n,p,q,r.
346   (S O) < n \to
347   p_ord n (nth_prime p) = pair nat nat q r \to
348   Not (r=O).
349 intros.
350 unfold.intro.
351 cut (O < n)
352   [lapply (p_ord_to_exp1 ? ? ? ? ? ? H1)
353     [apply lt_SO_nth_prime_n.
354     |assumption
355     |elim Hletin.
356      apply (lt_to_not_eq ? ? Hcut).
357      rewrite > H4.
358      rewrite > H2.
359      apply times_n_O
360     ]
361   |apply (trans_lt ? (S O))[apply lt_O_S|assumption]
362   ]
363 qed.
364
365 definition ord :nat \to nat \to nat \def
366 \lambda n,p. fst ? ? (p_ord n p).
367
368 definition ord_rem :nat \to nat \to nat \def
369 \lambda n,p. snd ? ? (p_ord n p).
370          
371 theorem divides_to_ord: \forall p,n,m:nat. 
372 O < n \to O < m \to prime p 
373 \to divides n m 
374 \to divides (ord_rem n p) (ord_rem m p) \land (ord n p) \le (ord m p).  
375 intros.
376 apply (divides_to_p_ord p ? ? ? ? n m H H1 H2 H3)
377   [unfold ord.unfold ord_rem.apply eq_pair_fst_snd
378   |unfold ord.unfold ord_rem.apply eq_pair_fst_snd
379   ]
380 qed.
381
382 theorem divides_to_divides_ord_rem: \forall p,n,m:nat. 
383 O < n \to O < m \to prime p \to divides n m \to
384 divides (ord_rem n p) (ord_rem m p).
385 intros.
386 elim (divides_to_ord p n m H H1 H2 H3).assumption.
387 qed.
388
389 theorem divides_to_le_ord: \forall p,n,m:nat. 
390 O < n \to O < m \to prime p \to divides n m \to
391 (ord n p) \le (ord m p).  
392 intros.
393 elim (divides_to_ord p n m H H1 H2 H3).assumption.
394 qed.
395
396 theorem exp_ord: \forall p,n. (S O) \lt p 
397 \to O \lt n \to n = p \sup (ord n p) * (ord_rem n p).
398 intros.
399 elim (p_ord_to_exp1 p n (ord n p) (ord_rem n p))
400   [assumption
401   |assumption
402   |assumption
403   |unfold ord.unfold ord_rem.
404    apply eq_pair_fst_snd
405   ]
406 qed.
407
408 theorem divides_ord_rem: \forall p,n. (S O) < p \to O < n 
409 \to divides (ord_rem n p) n. 
410 intros.
411 apply (witness ? ? (p \sup (ord n p))).
412 rewrite > sym_times. 
413 apply exp_ord[assumption|assumption]
414 qed.
415
416 theorem lt_O_ord_rem: \forall p,n. (S O) < p \to O < n \to O < ord_rem n p. 
417 intros.
418 elim (le_to_or_lt_eq O (ord_rem n p))
419   [assumption
420   |apply False_ind.
421    apply (lt_to_not_eq ? ? H1).
422    lapply (divides_ord_rem ? ? H H1).
423    rewrite < H2 in Hletin.
424    elim Hletin.
425    rewrite > H3.
426    reflexivity
427   |apply le_O_n
428   ]
429 qed.
430
431 (* properties of ord e ord_rem *)
432
433 theorem ord_times: \forall p,m,n.O<m \to O<n \to prime p \to
434 ord (m*n) p = ord m p+ord n p.
435 intros.unfold ord.
436 rewrite > (p_ord_times ? ? ? (ord m p) (ord_rem m p) (ord n p) (ord_rem n p))
437   [reflexivity
438   |assumption
439   |assumption
440   |assumption
441   |unfold ord.unfold ord_rem.apply eq_pair_fst_snd
442   |unfold ord.unfold ord_rem.apply eq_pair_fst_snd
443   ]
444 qed.
445
446 theorem ord_exp: \forall p,m.S O < p \to
447 ord (exp p m) p = m.
448 intros.
449 unfold ord.
450 rewrite > (p_ord_exp1 p (exp p m) m (S O))
451   [reflexivity
452   |apply lt_to_le.assumption
453   |intro.apply (lt_to_not_le ? ? H).
454    apply divides_to_le
455     [apply lt_O_S
456     |assumption
457     ]
458   |apply times_n_SO
459   ]
460 qed.
461
462 theorem not_divides_to_ord_O: 
463 \forall p,m. prime p \to \lnot (divides p m) \to
464 ord m p = O.
465 intros.unfold ord.
466 rewrite > (p_ord_exp1 p m O m)
467   [reflexivity
468   |apply prime_to_lt_O.assumption
469   |assumption
470   |simplify.apply plus_n_O
471   ]
472 qed.
473
474 theorem ord_O_to_not_divides: 
475 \forall p,m. O < m \to prime p \to ord m p = O \to 
476 \lnot (divides p m).
477 intros.
478 lapply (p_ord_to_exp1 p m (ord m p) (ord_rem m p))
479   [elim Hletin.
480    rewrite > H4.
481    rewrite > H2.
482    rewrite > sym_times.
483    rewrite < times_n_SO.
484    assumption
485   |apply prime_to_lt_SO.assumption
486   |assumption
487   |unfold ord.unfold ord_rem.apply eq_pair_fst_snd
488   ]
489 qed.
490
491 theorem divides_to_not_ord_O: 
492 \forall p,m. O < m \to prime p \to divides p m \to
493 \lnot(ord m p = O).
494 intros.intro.
495 apply (ord_O_to_not_divides ? ? H H1 H3).
496 assumption.
497 qed.
498
499 theorem not_ord_O_to_divides: 
500 \forall p,m. O < m \to prime p \to \lnot (ord m p = O) \to 
501 divides p m.
502 intros.
503 rewrite > (exp_ord p) in ⊢ (? ? %)
504   [apply (trans_divides ? (exp p (ord m p)))
505     [generalize in match H2.
506      cases (ord m p);intro
507       [apply False_ind.apply H3.reflexivity
508       |simplify.autobatch
509       ]
510     |autobatch
511     ]
512   |apply prime_to_lt_SO.
513    assumption
514   |assumption
515   ]
516 qed.
517
518 theorem not_divides_ord_rem: \forall m,p.O< m \to S O < p \to
519 \lnot (divides p (ord_rem m p)).
520 intros.
521 elim (p_ord_to_exp1 p m (ord m p) (ord_rem m p))
522   [assumption
523   |assumption
524   |assumption
525   |unfold ord.unfold ord_rem.apply eq_pair_fst_snd
526   ]
527 qed.
528
529 theorem ord_ord_rem: \forall p,q,m. O < m \to 
530 prime p \to prime q \to
531 q < p \to ord (ord_rem m p) q = ord m q.
532 intros.
533 rewrite > (exp_ord p) in ⊢ (? ? ? (? % ?))
534   [rewrite > ord_times
535     [rewrite > not_divides_to_ord_O in ⊢ (? ? ? (? % ?))
536       [reflexivity
537       |assumption
538       |intro.
539        apply (lt_to_not_eq ? ? H3).
540        apply (divides_exp_to_eq ? ? ? H2 H1 H4)
541       ]
542     |apply lt_O_exp.
543      apply (ltn_to_ltO ? ? H3)
544     |apply lt_O_ord_rem
545       [elim H1.assumption
546       |assumption
547       
548       ]
549     |assumption
550     ]
551   |elim H1.assumption
552   |assumption
553   ]
554 qed.
555
556 theorem lt_ord_rem: \forall n,m. prime n \to O < m \to
557 divides n m \to ord_rem m n < m.
558 intros.
559 elim (le_to_or_lt_eq (ord_rem m n) m)
560   [assumption
561   |apply False_ind.
562    apply (ord_O_to_not_divides ? ? H1 H ? H2).
563    apply (inj_exp_r n)
564     [apply prime_to_lt_SO.assumption
565     |apply (inj_times_l1 m)
566       [assumption
567       |rewrite > sym_times in ⊢ (? ? ? %).
568        rewrite < times_n_SO.
569        rewrite < H3 in ⊢ (? ? (? ? %) ?).
570        apply sym_eq.
571        apply exp_ord
572         [apply prime_to_lt_SO.assumption
573         |assumption
574         ]
575       ]
576     ]
577   |apply divides_to_le
578     [assumption
579     |apply divides_ord_rem
580       [apply prime_to_lt_SO.assumption
581       |assumption
582       ]
583     ]
584   ]
585 qed.
586
587 (* p_ord_inv is used to encode the pair ord and rem into 
588    a single natural number. *)
589  
590 definition p_ord_inv \def
591 \lambda p,m,x.
592   match p_ord x p with
593   [pair q r \Rightarrow r*m+q].
594   
595 theorem  eq_p_ord_inv: \forall p,m,x.
596 p_ord_inv p m x = (ord_rem x p)*m+(ord x p).
597 intros.unfold p_ord_inv. unfold ord_rem.
598 unfold ord.
599 elim (p_ord x p).
600 reflexivity.
601 qed.
602
603 theorem div_p_ord_inv: 
604 \forall p,m,x. ord x p < m \to p_ord_inv p m x / m = ord_rem x p.
605 intros.rewrite > eq_p_ord_inv.
606 apply div_plus_times.
607 assumption.
608 qed.
609
610 theorem mod_p_ord_inv: 
611 \forall p,m,x. ord x p < m \to p_ord_inv p m x \mod m = ord x p.
612 intros.rewrite > eq_p_ord_inv.
613 apply mod_plus_times.
614 assumption.
615 qed.