1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 set "baseuri" "cic:/matita/nat/gcd_properties1".
19 (* this file contains some important properites of gcd in N *)
21 (* an alternative characterization for gcd *)
22 theorem gcd1: \forall a,b,c:nat.
23 c \divides a \to c \divides b \to
24 (\forall d:nat. d \divides a \to d \divides b \to d \divides c) \to (gcd a b) = c.
27 [ apply (antisymmetric_divides (gcd a b) c)
28 [ apply (witness (gcd a b) c n2).
30 | apply divides_d_gcd;
40 theorem eq_gcd_times_times_times_gcd: \forall a,b,c:nat.
41 (gcd (c*a) (c*b)) = c*(gcd a b).
52 | apply divides_gcd_n.
60 apply (divides_d_times_gcd)
70 theorem associative_nat_gcd: associative nat gcd.
71 change with (\forall a,b,c:nat. (gcd (gcd a b) c) = (gcd a (gcd b c))).
75 [ apply (trans_divides ? (gcd b c) ?)
81 | apply (trans_divides ? (gcd b c) ?)
86 cut (d \divides a \land d \divides b)
88 cut (d \divides (gcd b c))
89 [ apply (divides_d_gcd (gcd b c) a d Hcut1 H2)
90 | apply (divides_d_gcd c b d H1 H3)
93 [ apply (trans_divides d (gcd a b) a H).
95 | apply (trans_divides d (gcd a b) b H).
103 theorem eq_gcd_div_div_div_gcd: \forall a,b,m:nat.
104 O \lt m \to m \divides a \to m \divides b \to
105 (gcd (a/m) (b/m)) = (gcd a b) / m.
107 apply (inj_times_r1 m H).
108 rewrite > (sym_times m ((gcd a b)/m)).
109 rewrite > (divides_to_div m (gcd a b))
110 [ rewrite < eq_gcd_times_times_times_gcd.
111 rewrite > (sym_times m (a/m)).
112 rewrite > (sym_times m (b/m)).
113 rewrite > (divides_to_div m a H1).
114 rewrite > (divides_to_div m b H2).
116 | apply divides_d_gcd;
123 theorem divides_times_to_divides_div_gcd: \forall a,b,c:nat.
124 a \divides (b*c) \to (a/(gcd a b)) \divides c.
129 [ (*It's an impossible situation*)
136 apply (divides_n_n O)
137 | apply (lt_times_eq_O b c)
140 | apply antisymmetric_divides
152 [ cut (O \lt (gcd a b))
153 [ apply (gcd_SO_to_divides_times_to_divides (b/(gcd a b)) (a/(gcd a b)) c)
154 [ apply (O_lt_times_to_O_lt (a/(gcd a b)) (gcd a b)).
155 rewrite > (divides_to_div (gcd a b) a)
157 | apply divides_gcd_n
159 | rewrite < (div_n_n (gcd a b)) in \vdash (? ? ? %)
160 [ apply eq_gcd_div_div_div_gcd
162 | apply divides_gcd_n
163 | apply divides_gcd_m
167 | apply (witness ? ? n2).
168 apply (inj_times_r1 (gcd a b) Hcut1).
169 rewrite < assoc_times.
170 rewrite < sym_times in \vdash (? ? (? % ?) ?).
171 rewrite > (divides_to_div (gcd a b) b)
172 [ rewrite < assoc_times in \vdash (? ? ? %).
173 rewrite < sym_times in \vdash (? ? ? (? % ?)).
174 rewrite > (divides_to_div (gcd a b) a)
176 | apply divides_gcd_n
178 | apply divides_gcd_m
191 theorem gcd_plus_times_gcd: \forall a,b,d,m:nat.
192 (gcd (a+m*b) b) = (gcd a b).
196 [ apply divides_gcd_n
197 | apply (trans_divides ? b ?)
198 [ apply divides_gcd_m
199 | rewrite > sym_times.
200 apply (witness b (b*m) m).
204 | apply divides_gcd_m
208 | rewrite > (minus_plus_m_m a (m*b)).
211 | apply (trans_divides ? b ?)
213 | rewrite > sym_times.
214 apply (witness b (b*m) m).
224 theorem gcd_SO_to_divides_to_divides_to_divides_times: \forall c,e,f:nat.
225 (gcd e f) = (S O) \to e \divides c \to f \divides c \to
228 apply (nat_case1 c); intros
234 rewrite > (sym_times e f).
235 apply (divides_times)
236 [ apply (divides_n_n)
237 | rewrite > H5 in H1.
238 apply (gcd_SO_to_divides_times_to_divides f e n)
239 [ rewrite < H5 in H1.
241 apply (divides_to_lt_O e (S m))
251 (* the following theorem shows that gcd is a multiplicative function in
252 the following sense: if a1 and a2 are relatively prime, then
253 gcd(a1·a2, b) = gcd(a1, b)·gcd(a2, b).
255 theorem gcd_SO_to_eq_gcd_times_times_gcd_gcd: \forall a,b,c:nat.
256 (gcd a b) = (S O) \to (gcd (a*b) c) = (gcd a c) * (gcd b c).
259 [ apply divides_times;
261 | apply (gcd_SO_to_divides_to_divides_to_divides_times c (gcd a c) (gcd b c))
269 apply (divides_d_gcd b a d Hcut1 Hcut)
270 | apply (trans_divides d (gcd b c) b)
272 | apply (divides_gcd_n)
275 | apply (trans_divides d (gcd a c) a)
277 | apply (divides_gcd_n)
281 | apply (divides_gcd_m)
282 | apply (divides_gcd_m)
285 rewrite < (eq_gcd_times_times_times_gcd b c (gcd a c)).
286 rewrite > (sym_times (gcd a c) b).
287 rewrite > (sym_times (gcd a c) c).
288 rewrite < (eq_gcd_times_times_times_gcd a c b).
289 rewrite < (eq_gcd_times_times_times_gcd a c c).
290 apply (divides_d_gcd)
291 [ apply (divides_d_gcd)
292 [ rewrite > (times_n_SO d).
293 apply (divides_times)
297 | rewrite > (times_n_SO d).
298 apply (divides_times)
303 | apply (divides_d_gcd)
304 [ rewrite > (times_n_SO d).
305 rewrite > (sym_times d (S O)).
306 apply (divides_times)
307 [ apply (divides_SO_n)
310 | rewrite < (sym_times a b).
318 theorem eq_gcd_gcd_minus: \forall a,b:nat.
319 a \lt b \to (gcd a b) = (gcd (b - a) b).
323 [ apply (divides_minus (gcd a b) b a)
324 [ apply divides_gcd_m
325 | apply divides_gcd_n
327 | apply divides_gcd_m
332 [ cut (b - (d*n2) = a)
333 [ rewrite > H4 in Hcut1.
334 rewrite < (distr_times_minus d n n2) in Hcut1.
337 | apply (witness d a (n - n2)).
341 | apply (plus_to_minus ? ? ? Hcut)
343 | rewrite > sym_plus.
344 apply (minus_to_plus)