]> matita.cs.unibo.it Git - helm.git/blob - matita/library/nat/ord.ma
More restructuring in moebius.ma
[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 set "baseuri" "cic:/matita/nat/ord".
16
17 include "datatypes/constructors.ma".
18 include "nat/exp.ma".
19 include "nat/nth_prime.ma".
20 include "nat/relevant_equations.ma". (* required by autobatch paramod *)
21
22 let rec p_ord_aux p n m \def
23   match n \mod m with
24   [ O \Rightarrow 
25     match p with
26       [ O \Rightarrow pair nat nat O n
27       | (S p) \Rightarrow 
28        match (p_ord_aux p (n / m) m) with
29        [ (pair q r) \Rightarrow pair nat nat (S q) r] ]
30   | (S a) \Rightarrow pair nat nat O n].
31   
32 (* p_ord n m = <q,r> if m divides n q times, with remainder r *)
33 definition p_ord \def \lambda n,m:nat.p_ord_aux n n m.
34
35 theorem p_ord_aux_to_Prop: \forall p,n,m. O < m \to
36   match p_ord_aux p n m with
37   [ (pair q r) \Rightarrow n = m \sup q *r ].
38 intro.
39 elim p.simplify.
40 apply (nat_case (n \mod m)).
41 simplify.apply plus_n_O.
42 intros.
43 simplify.apply plus_n_O.
44 simplify. 
45 apply (nat_case1 (n1 \mod m)).intro.
46 simplify.
47 generalize in match (H (n1 / m) m).
48 elim (p_ord_aux n (n1 / m) m).
49 simplify.
50 rewrite > assoc_times.
51 rewrite < H3.rewrite > (plus_n_O (m*(n1 / m))).
52 rewrite < H2.
53 rewrite > sym_times.
54 rewrite < div_mod.reflexivity.
55 assumption.assumption.
56 intros.simplify.apply plus_n_O. 
57 qed.
58
59 theorem p_ord_aux_to_exp: \forall p,n,m,q,r. O < m \to
60   (pair nat nat q r) = p_ord_aux p n m \to n = m \sup q * r.
61 intros.
62 change with 
63 match (pair nat nat q r) with
64   [ (pair q r) \Rightarrow n = m \sup q * r ].
65 rewrite > H1.
66 apply p_ord_aux_to_Prop.
67 assumption.
68 qed.
69
70 (* questo va spostato in primes1.ma *)
71 theorem p_ord_exp: \forall n,m,i. O < m \to n \mod m \neq O \to
72 \forall p. i \le p \to p_ord_aux p (m \sup i * n) m = pair nat nat i n.
73 intros 5.
74 elim i.
75 simplify.
76 rewrite < plus_n_O.
77 apply (nat_case p).
78 simplify.
79 elim (n \mod m).simplify.reflexivity.simplify.reflexivity.
80 intro.
81 simplify.
82 cut (O < n \mod m \lor O = n \mod m).
83 elim Hcut.apply (lt_O_n_elim (n \mod m) H3).
84 intros. simplify.reflexivity.
85 apply False_ind.
86 apply H1.apply sym_eq.assumption.
87 apply le_to_or_lt_eq.apply le_O_n.
88 generalize in match H3.
89 apply (nat_case p).intro.apply False_ind.apply (not_le_Sn_O n1 H4).
90 intros.
91 simplify. fold simplify (m \sup (S n1)).
92 cut (((m \sup (S n1)*n) \mod m) = O).
93 rewrite > Hcut.
94 simplify.fold simplify (m \sup (S n1)).
95 cut ((m \sup (S n1) *n) / m = m \sup n1 *n).
96 rewrite > Hcut1.
97 rewrite > (H2 m1). simplify.reflexivity.
98 apply le_S_S_to_le.assumption.
99 (* div_exp *)
100 simplify.
101 rewrite > assoc_times.
102 apply (lt_O_n_elim m H).
103 intro.apply div_times.
104 (* mod_exp = O *)
105 apply divides_to_mod_O.
106 assumption.
107 simplify.rewrite > assoc_times.
108 apply (witness ? ? (m \sup n1 *n)).reflexivity.
109 qed.
110
111 theorem p_ord_aux_to_Prop1: \forall p,n,m. (S O) < m \to O < n \to n \le p \to
112   match p_ord_aux p n m with
113   [ (pair q r) \Rightarrow r \mod m \neq O].
114 intro.elim p.absurd (O < n).assumption.
115 apply le_to_not_lt.assumption.
116 simplify.
117 apply (nat_case1 (n1 \mod m)).intro.
118 generalize in match (H (n1 / m) m).
119 elim (p_ord_aux n (n1 / m) m).
120 apply H5.assumption.
121 apply eq_mod_O_to_lt_O_div.
122 apply (trans_lt ? (S O)).unfold lt.apply le_n.
123 assumption.assumption.assumption.
124 apply le_S_S_to_le.
125 apply (trans_le ? n1).change with (n1 / m < n1).
126 apply lt_div_n_m_n.assumption.assumption.assumption.
127 intros.simplify.
128 rewrite > H4.
129 unfold Not.intro.
130 apply (not_eq_O_S m1).
131 rewrite > H5.reflexivity.
132 qed.
133
134 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
135  pair nat nat q r = p_ord_aux p n m \to r \mod m \neq O.
136 intros.
137 change with 
138   match (pair nat nat q r) with
139   [ (pair q r) \Rightarrow r \mod m \neq O].
140 rewrite > H3. 
141 apply p_ord_aux_to_Prop1.
142 assumption.assumption.assumption.
143 qed.
144
145 theorem p_ord_exp1: \forall p,n,q,r. O < p \to \lnot p \divides r \to
146 n = p \sup q * r \to p_ord n p = pair nat nat q r.
147 intros.unfold p_ord.
148 rewrite > H2.
149 apply p_ord_exp
150  [assumption
151  |unfold.intro.apply H1.
152   apply mod_O_to_divides[assumption|assumption]
153  |apply (trans_le ? (p \sup q)).
154   cut ((S O) \lt p).
155   elim q.simplify.apply le_n_Sn.
156   simplify.
157   generalize in match H3.
158   apply (nat_case n1).simplify.
159   rewrite < times_n_SO.intro.assumption.
160   intros.
161   apply (trans_le ? (p*(S m))).
162   apply (trans_le ? ((S (S O))*(S m))).
163   simplify.rewrite > plus_n_Sm.
164   rewrite < plus_n_O.
165   apply le_plus_n.
166   apply le_times_l.
167   assumption.
168   apply le_times_r.assumption.
169 alias id "not_eq_to_le_to_lt" = "cic:/matita/algebra/finite_groups/not_eq_to_le_to_lt.con".
170 apply not_eq_to_le_to_lt.
171   unfold.intro.apply H1.
172   rewrite < H3.
173   apply (witness ? r r ?).simplify.apply plus_n_O.
174   assumption.
175   rewrite > times_n_SO in \vdash (? % ?).
176   apply le_times_r.
177   change with (O \lt r).
178   apply not_eq_to_le_to_lt.
179   unfold.intro.
180   apply H1.rewrite < H3.
181   apply (witness ? ? O ?).rewrite < times_n_O.reflexivity.
182   apply le_O_n.
183   ]
184 qed.
185
186 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 
187 \lnot p \divides r \land n = p \sup q * r.
188 intros.
189 unfold p_ord in H2.
190 split.unfold.intro.
191 apply (p_ord_aux_to_not_mod_O n n p q r).assumption.assumption.
192 apply le_n.symmetry.assumption.
193 apply divides_to_mod_O.apply (trans_lt ? (S O)).
194 unfold.apply le_n.assumption.assumption.
195 apply (p_ord_aux_to_exp n).apply (trans_lt ? (S O)).
196 unfold.apply le_n.assumption.symmetry.assumption.
197 qed.
198
199 theorem p_ord_times: \forall p,a,b,qa,ra,qb,rb. prime p 
200 \to O \lt a \to O \lt b 
201 \to p_ord a p = pair nat nat qa ra  
202 \to p_ord b p = pair nat nat qb rb
203 \to p_ord (a*b) p = pair nat nat (qa + qb) (ra*rb).
204 intros.
205 cut ((S O) \lt p).
206 elim (p_ord_to_exp1 ? ? ? ? Hcut H1 H3).
207 elim (p_ord_to_exp1 ? ? ? ? Hcut H2 H4).
208 apply p_ord_exp1.
209 apply (trans_lt ? (S O)).unfold.apply le_n.assumption.
210 unfold.intro.
211 elim (divides_times_to_divides ? ? ? H H9).
212 apply (absurd ? ? H10 H5).
213 apply (absurd ? ? H10 H7).
214 (* rewrite > H6.
215 rewrite > H8. *)
216 autobatch paramodulation.
217 unfold prime in H. elim H. assumption.
218 qed.
219
220 theorem fst_p_ord_times: \forall p,a,b. prime p 
221 \to O \lt a \to O \lt b 
222 \to fst ? ? (p_ord (a*b) p) = (fst ? ? (p_ord a p)) + (fst ? ? (p_ord b p)).
223 intros.
224 rewrite > (p_ord_times p a b (fst ? ? (p_ord a p)) (snd ? ? (p_ord a p))
225 (fst ? ? (p_ord b p)) (snd ? ? (p_ord b p)) H H1 H2).
226 simplify.reflexivity.
227 apply eq_pair_fst_snd.
228 apply eq_pair_fst_snd.
229 qed.
230
231 theorem p_ord_p : \forall p:nat. (S O) \lt p \to p_ord p p = pair ? ? (S O) (S O).
232 intros.
233 apply p_ord_exp1.
234 apply (trans_lt ? (S O)). unfold.apply le_n.assumption.
235 unfold.intro.
236 apply (absurd ? ? H).
237 apply le_to_not_lt.
238 apply divides_to_le.unfold.apply le_n.assumption.
239 rewrite < times_n_SO.
240 apply exp_n_SO.
241 qed.
242
243 (* p_ord and divides *)
244 theorem divides_to_p_ord: \forall p,a,b,c,d,n,m:nat. 
245 O < n \to O < m \to prime p 
246 \to divides n m \to p_ord n p = pair ? ? a b \to
247 p_ord m p = pair ? ? c d \to divides b d \land a \le c.
248 intros.
249 cut (S O < p)
250   [lapply (p_ord_to_exp1 ? ? ? ? Hcut H H4).
251    lapply (p_ord_to_exp1 ? ? ? ? Hcut H1 H5).
252    elim Hletin. clear Hletin.
253    elim Hletin1. clear Hletin1.
254    rewrite > H9 in H3.
255    split
256     [apply (gcd_SO_to_divides_times_to_divides (exp p c))
257       [elim (le_to_or_lt_eq ? ? (le_O_n b))
258         [assumption
259         |apply False_ind.
260          apply (lt_to_not_eq O ? H).
261          rewrite > H7.
262          rewrite < H10.
263          autobatch
264         ]
265       |elim c
266         [rewrite > sym_gcd.
267          apply gcd_SO_n
268         |simplify.
269          apply eq_gcd_times_SO
270           [apply lt_to_le.assumption
271           |apply lt_O_exp.apply lt_to_le.assumption
272           |rewrite > sym_gcd.
273            (* hint non trova prime_to_gcd_SO e
274               autobatch non chiude il goal *)
275            apply prime_to_gcd_SO
276             [assumption|assumption]
277           |assumption
278           ]
279         ]
280       |apply (trans_divides ? n)
281         [apply (witness ? ? (exp p a)).
282          rewrite > sym_times.
283          assumption
284         |assumption
285         ]
286       ]
287     |apply (le_exp_to_le p)
288       [assumption
289       |apply divides_to_le
290         [apply lt_O_exp.apply lt_to_le.assumption
291         |apply (gcd_SO_to_divides_times_to_divides d)
292           [apply lt_O_exp.apply lt_to_le.assumption
293           |elim a
294             [apply gcd_SO_n
295             |simplify.rewrite < sym_gcd.
296              apply eq_gcd_times_SO
297               [apply lt_to_le.assumption
298               |apply lt_O_exp.apply lt_to_le.assumption
299               |rewrite > sym_gcd.
300               (* hint non trova prime_to_gcd_SO e
301                  autobatch non chiude il goal *)
302                apply prime_to_gcd_SO
303                 [assumption|assumption]
304               |rewrite > sym_gcd. assumption
305               ]
306             ]
307           |apply (trans_divides ? n)
308             [apply (witness ? ? b).assumption
309             |rewrite > sym_times.assumption
310             ]
311           ]
312         ]
313       ]
314     ]
315   |elim H2.assumption
316   ]
317 qed.    
318
319 (* p_ord and primes *)
320 theorem not_divides_to_p_ord_O: \forall n,i.
321 Not (divides (nth_prime i) n) \to p_ord n (nth_prime i) = 
322 pair nat nat O n.
323 intros.
324 apply p_ord_exp1
325   [apply lt_O_nth_prime_n
326   |assumption
327   |autobatch
328   ]
329 qed.
330
331 theorem p_ord_O_to_not_divides: \forall n,i,r.
332 O < n \to
333 p_ord n (nth_prime i) = pair nat nat O r 
334 \to Not (divides (nth_prime i) n).
335 intros.
336 lapply (p_ord_to_exp1 ? ? ? ? ? ? H1)
337   [apply lt_SO_nth_prime_n
338   |assumption
339   |elim Hletin.
340    simplify in H3.
341    rewrite > H3.
342    rewrite < plus_n_O.
343    assumption
344   ]
345 qed.
346
347 theorem p_ord_to_not_eq_O : \forall n,p,q,r.
348   (S O) < n \to
349   p_ord n (nth_prime p) = pair nat nat q r \to
350   Not (r=O).
351 intros.
352 unfold.intro.
353 cut (O < n)
354   [lapply (p_ord_to_exp1 ? ? ? ? ? ? H1)
355     [apply lt_SO_nth_prime_n.
356     |assumption
357     |elim Hletin.
358      apply (lt_to_not_eq ? ? Hcut).
359      rewrite > H4.
360      rewrite > H2.
361      apply times_n_O
362     ]
363   |apply (trans_lt ? (S O))[apply lt_O_S|assumption]
364   ]
365 qed.
366
367 definition ord :nat \to nat \to nat \def
368 \lambda n,p. fst ? ? (p_ord n p).
369
370 definition ord_rem :nat \to nat \to nat \def
371 \lambda n,p. snd ? ? (p_ord n p).
372          
373 theorem divides_to_ord: \forall p,n,m:nat. 
374 O < n \to O < m \to prime p 
375 \to divides n m 
376 \to divides (ord_rem n p) (ord_rem m p) \land (ord n p) \le (ord m p).  
377 intros.
378 apply (divides_to_p_ord p ? ? ? ? n m H H1 H2 H3)
379   [unfold ord.unfold ord_rem.apply eq_pair_fst_snd
380   |unfold ord.unfold ord_rem.apply eq_pair_fst_snd
381   ]
382 qed.
383
384 theorem divides_to_divides_ord_rem: \forall p,n,m:nat. 
385 O < n \to O < m \to prime p \to divides n m \to
386 divides (ord_rem n p) (ord_rem m p).
387 intros.
388 elim (divides_to_ord p n m H H1 H2 H3).assumption.
389 qed.
390
391 theorem divides_to_le_ord: \forall p,n,m:nat. 
392 O < n \to O < m \to prime p \to divides n m \to
393 (ord n p) \le (ord m p).  
394 intros.
395 elim (divides_to_ord p n m H H1 H2 H3).assumption.
396 qed.
397
398 theorem exp_ord: \forall p,n. (S O) \lt p 
399 \to O \lt n \to n = p \sup (ord n p) * (ord_rem n p).
400 intros.
401 elim (p_ord_to_exp1 p n (ord n p) (ord_rem n p))
402   [assumption
403   |assumption
404   |assumption
405   |unfold ord.unfold ord_rem.
406    apply eq_pair_fst_snd
407   ]
408 qed.
409
410 theorem divides_ord_rem: \forall p,n. (S O) < p \to O < n 
411 \to divides (ord_rem n p) n. 
412 intros.
413 apply (witness ? ? (p \sup (ord n p))).
414 rewrite > sym_times. 
415 apply exp_ord[assumption|assumption]
416 qed.
417
418 theorem lt_O_ord_rem: \forall p,n. (S O) < p \to O < n \to O < ord_rem n p. 
419 intros.
420 elim (le_to_or_lt_eq O (ord_rem n p))
421   [assumption
422   |apply False_ind.
423    apply (lt_to_not_eq ? ? H1).
424    lapply (divides_ord_rem ? ? H H1).
425    rewrite < H2 in Hletin.
426    elim Hletin.
427    rewrite > H3.
428    reflexivity
429   |apply le_O_n
430   ]
431 qed.
432
433 (* p_ord_inv is the inverse of ord *)
434 definition p_ord_inv \def
435 \lambda p,m,x.
436   match p_ord x p with
437   [pair q r \Rightarrow r*m+q].
438   
439 theorem  eq_p_ord_inv: \forall p,m,x.
440 p_ord_inv p m x = (ord_rem x p)*m+(ord x p).
441 intros.unfold p_ord_inv. unfold ord_rem.
442 unfold ord.
443 elim (p_ord x p).
444 reflexivity.
445 qed.
446
447 theorem div_p_ord_inv: 
448 \forall p,m,x. ord x p < m \to p_ord_inv p m x / m = ord_rem x p.
449 intros.rewrite > eq_p_ord_inv.
450 apply div_plus_times.
451 assumption.
452 qed.
453
454 theorem mod_p_ord_inv: 
455 \forall p,m,x. ord x p < m \to p_ord_inv p m x \mod m = ord x p.
456 intros.rewrite > eq_p_ord_inv.
457 apply mod_plus_times.
458 assumption.
459 qed.