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 set "baseuri" "cic:/matita/library_autobatch/nat/ord".
17 include "datatypes/constructors.ma".
18 include "auto/nat/exp.ma".
19 include "auto/nat/gcd.ma".
20 include "auto/nat/relevant_equations.ma". (* required by autobatch paramod *)
22 (* this definition of log is based on pairs, with a remainder *)
24 let rec p_ord_aux p n m \def
28 [ O \Rightarrow pair nat nat O n
30 match (p_ord_aux p (n / m) m) with
31 [ (pair q r) \Rightarrow pair nat nat (S q) r] ]
32 | (S a) \Rightarrow pair nat nat O n].
34 (* p_ord n m = <q,r> if m divides n q times, with remainder r *)
35 definition p_ord \def \lambda n,m:nat.p_ord_aux n n m.
37 theorem p_ord_aux_to_Prop: \forall p,n,m. O < m \to
38 match p_ord_aux p n m with
39 [ (pair q r) \Rightarrow n = m \sup q *r ].
43 apply (nat_case (n \mod m))
51 apply (nat_case1 (n1 \mod m))
54 generalize in match (H (n1 / m) m).
55 elim (p_ord_aux n (n1 / m) m).
57 rewrite > assoc_times.
59 [ rewrite > (plus_n_O (m*(n1 / m))).
76 theorem p_ord_aux_to_exp: \forall p,n,m,q,r. O < m \to
77 (pair nat nat q r) = p_ord_aux p n m \to n = m \sup q * r.
80 match (pair nat nat q r) with
81 [ (pair q r) \Rightarrow n = m \sup q * r ].
83 apply p_ord_aux_to_Prop.
87 (* questo va spostato in primes1.ma *)
88 theorem p_ord_exp: \forall n,m,i. O < m \to n \mod m \neq O \to
89 \forall p. i \le p \to p_ord_aux p (m \sup i * n) m = pair nat nat i n.
96 elim (n \mod m);autobatch
104 cut (O < n \mod m \lor O = n \mod m)
106 [ apply (lt_O_n_elim (n \mod m) H3).
110 | apply False_ind.autobatch
116 (*apply le_to_or_lt_eq.
120 | generalize in match H3.
124 apply (not_le_Sn_O n1 H4)
127 fold simplify (m \sup (S n1)).
128 cut (((m \sup (S n1)*n) \mod m) = O)
131 fold simplify (m \sup (S n1)).
132 cut ((m \sup (S n1) *n) / m = m \sup n1 *n)
134 rewrite > (H2 m1);autobatch
137 | apply le_S_S_to_le.
142 rewrite > assoc_times.
143 apply (lt_O_n_elim m H).
148 apply divides_to_mod_O
151 (*rewrite > assoc_times.
152 apply (witness ? ? (m \sup n1 *n)).
160 theorem p_ord_aux_to_Prop1: \forall p,n,m. (S O) < m \to O < n \to n \le p \to
161 match p_ord_aux p n m with
162 [ (pair q r) \Rightarrow r \mod m \neq O].
165 [ absurd (O < n);autobatch
167 | apply le_to_not_lt.
171 apply (nat_case1 (n1 \mod m))
173 generalize in match (H (n1 / m) m).
174 elim (p_ord_aux n (n1 / m) m).
178 (*apply eq_mod_O_to_lt_O_div
179 [ apply (trans_lt ? (S O))
187 | apply le_S_S_to_le.autobatch
188 (*apply (trans_le ? n1)
189 [ change with (n1 / m < n1).
190 apply lt_div_n_m_n;assumption
199 apply (not_eq_O_S m1).
206 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
207 pair nat nat q r = p_ord_aux p n m \to r \mod m \neq O.
210 match (pair nat nat q r) with
211 [ (pair q r) \Rightarrow r \mod m \neq O].
213 apply p_ord_aux_to_Prop1;
217 axiom not_eq_to_le_to_lt: ∀n,m. n≠m → n≤m → n<m.
219 theorem p_ord_exp1: \forall p,n,q,r. O < p \to \lnot p \divides r \to
220 n = p \sup q * r \to p_ord n p = pair nat nat q r.
230 apply mod_O_to_divides
234 | apply (trans_le ? (p \sup q))
241 generalize in match H3.
244 rewrite < times_n_SO.
248 apply (trans_le ? (p*(S m)))
249 [ apply (trans_le ? ((S (S O))*(S m)))
262 | apply not_eq_to_le_to_lt
268 apply (witness ? r r ?).
274 | rewrite > times_n_SO in \vdash (? % ?).
276 change with (O \lt r).
277 apply not_eq_to_le_to_lt
280 (*apply H1.rewrite < H3.
281 apply (witness ? ? O ?).rewrite < times_n_O.
289 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
290 \lnot p \divides r \land n = p \sup q * r.
295 apply (p_ord_aux_to_not_mod_O n n p q r);autobatch
301 | apply divides_to_mod_O
302 [ apply (trans_lt ? (S O))
310 | apply (p_ord_aux_to_exp n);autobatch
311 (*[ apply (trans_lt ? (S O))
322 theorem p_ord_times: \forall p,a,b,qa,ra,qb,rb. prime p
323 \to O \lt a \to O \lt b
324 \to p_ord a p = pair nat nat qa ra
325 \to p_ord b p = pair nat nat qb rb
326 \to p_ord (a*b) p = pair nat nat (qa + qb) (ra*rb).
329 [ elim (p_ord_to_exp1 ? ? ? ? Hcut H1 H3).
330 elim (p_ord_to_exp1 ? ? ? ? Hcut H2 H4).
333 (*apply (trans_lt ? (S O))
340 elim (divides_times_to_divides ? ? ? H H9);autobatch
341 (*[ apply (absurd ? ? H10 H5)
342 | apply (absurd ? ? H10 H7)
346 autobatch paramodulation
354 theorem fst_p_ord_times: \forall p,a,b. prime p
355 \to O \lt a \to O \lt b
356 \to fst ? ? (p_ord (a*b) p) = (fst ? ? (p_ord a p)) + (fst ? ? (p_ord b p)).
358 rewrite > (p_ord_times p a b (fst ? ? (p_ord a p)) (snd ? ? (p_ord a p))
359 (fst ? ? (p_ord b p)) (snd ? ? (p_ord b p)) H H1 H2);autobatch.
362 | apply eq_pair_fst_snd
363 | apply eq_pair_fst_snd
367 theorem p_ord_p : \forall p:nat. (S O) \lt p \to p_ord p p = pair ? ? (S O) (S O).
371 (*apply (trans_lt ? (S O))
378 apply (absurd ? ? H).autobatch
379 (*apply le_to_not_lt.
386 (*rewrite < times_n_SO.