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/nat/log".
17 include "datatypes/constructors.ma".
19 include "nat/lt_arith.ma".
20 include "nat/primes.ma".
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 ].
45 [ O \Rightarrow pair nat nat O n
46 | (S a) \Rightarrow pair nat nat O n] )
48 [ (pair q r) \Rightarrow n = m \sup q * r ].
49 apply (nat_case (n \mod m)).
50 simplify.apply plus_n_O.
52 simplify.apply plus_n_O.
57 match (p_ord_aux n (n1 / m) m) with
58 [ (pair q r) \Rightarrow pair nat nat (S q) r]
59 | (S a) \Rightarrow pair nat nat O n1] )
61 [ (pair q r) \Rightarrow n1 = m \sup q * r].
62 apply (nat_case1 (n1 \mod m)).intro.
65 match (p_ord_aux n (n1 / m) m) with
66 [ (pair q r) \Rightarrow pair nat nat (S q) r])
68 [ (pair q r) \Rightarrow n1 = m \sup q * r].
69 generalize in match (H (n1 / m) m).
70 elim (p_ord_aux n (n1 / m) m).
72 rewrite > assoc_times.
73 rewrite < H3.rewrite > (plus_n_O (m*(n1 / m))).
76 rewrite < div_mod.reflexivity.
77 assumption.assumption.
78 intros.simplify.apply plus_n_O.
81 theorem p_ord_aux_to_exp: \forall p,n,m,q,r. O < m \to
82 (pair nat nat q r) = p_ord_aux p n m \to n = m \sup q * r.
85 match (pair nat nat q r) with
86 [ (pair q r) \Rightarrow n = m \sup q * r ].
88 apply p_ord_aux_to_Prop.
91 (* questo va spostato in primes1.ma *)
92 theorem p_ord_exp: \forall n,m,i. O < m \to n \mod m \neq O \to
93 \forall p. i \le p \to p_ord_aux p (m \sup i * n) m = pair nat nat i n.
101 [ O \Rightarrow pair nat nat O n
102 | (S a) \Rightarrow pair nat nat O n]
104 elim (n \mod m).simplify.reflexivity.simplify.reflexivity.
109 match (p_ord_aux m1 (n / m) m) with
110 [ (pair q r) \Rightarrow pair nat nat (S q) r]
111 | (S a) \Rightarrow pair nat nat O n]
113 cut (O < n \mod m \lor O = n \mod m).
114 elim Hcut.apply (lt_O_n_elim (n \mod m) H3).
115 intros. simplify.reflexivity.
117 apply H1.apply sym_eq.assumption.
118 apply le_to_or_lt_eq.apply le_O_n.
119 generalize in match H3.
120 apply (nat_case p).intro.apply False_ind.apply (not_le_Sn_O n1 H4).
123 (match ((m \sup (S n1) *n) \mod m) with
125 match (p_ord_aux m1 ((m \sup (S n1) *n) / m) m) with
126 [ (pair q r) \Rightarrow pair nat nat (S q) r]
127 | (S a) \Rightarrow pair nat nat O (m \sup (S n1) *n)]
128 = pair nat nat (S n1) n).
129 cut (((m \sup (S n1)*n) \mod m) = O).
132 (match (p_ord_aux m1 ((m \sup (S n1)*n) / m) m) with
133 [ (pair q r) \Rightarrow pair nat nat (S q) r]
134 = pair nat nat (S n1) n).
135 cut ((m \sup (S n1) *n) / m = m \sup n1 *n).
137 rewrite > (H2 m1). simplify.reflexivity.
138 apply le_S_S_to_le.assumption.
140 change with ((m* m \sup n1 *n) / m = m \sup n1 * n).
141 rewrite > assoc_times.
142 apply (lt_O_n_elim m H).
143 intro.apply div_times.
145 apply divides_to_mod_O.
147 simplify.rewrite > assoc_times.
148 apply (witness ? ? (m \sup n1 *n)).reflexivity.
151 theorem p_ord_aux_to_Prop1: \forall p,n,m. (S O) < m \to O < n \to n \le p \to
152 match p_ord_aux p n m with
153 [ (pair q r) \Rightarrow r \mod m \neq O].
154 intro.elim p.absurd (O < n).assumption.
155 apply le_to_not_lt.assumption.
158 (match n1 \mod m with
160 match (p_ord_aux n(n1 / m) m) with
161 [ (pair q r) \Rightarrow pair nat nat (S q) r]
162 | (S a) \Rightarrow pair nat nat O n1])
164 [ (pair q r) \Rightarrow r \mod m \neq O].
165 apply (nat_case1 (n1 \mod m)).intro.
166 generalize in match (H (n1 / m) m).
167 elim (p_ord_aux n (n1 / m) m).
169 apply eq_mod_O_to_lt_O_div.
170 apply (trans_lt ? (S O)).unfold lt.apply le_n.
171 assumption.assumption.assumption.
173 apply (trans_le ? n1).change with (n1 / m < n1).
174 apply lt_div_n_m_n.assumption.assumption.assumption.
176 change with (n1 \mod m \neq O).
179 apply (not_eq_O_S m1).
180 rewrite > H5.reflexivity.
183 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
184 pair nat nat q r = p_ord_aux p n m \to r \mod m \neq O.
187 match (pair nat nat q r) with
188 [ (pair q r) \Rightarrow r \mod m \neq O].
190 apply p_ord_aux_to_Prop1.
191 assumption.assumption.assumption.