author | wenzelm |
Tue, 27 May 1997 15:45:07 +0200 | |
changeset 3362 | 0b268cff9344 |
parent 3359 | 88cd6a2c6ebe |
child 3377 | afa1fedef73f |
permissions | -rw-r--r-- |
1804 | 1 |
(* Title: HOL/ex/Primes.ML |
2 |
ID: $Id$ |
|
3 |
Author: Christophe Tabacznyj and Lawrence C Paulson |
|
4 |
Copyright 1996 University of Cambridge |
|
5 |
||
6 |
The "divides" relation, the greatest common divisor and Euclid's algorithm |
|
7 |
*) |
|
8 |
||
9 |
eta_contract:=false; |
|
10 |
||
11 |
open Primes; |
|
12 |
||
13 |
(************************************************) |
|
14 |
(** Divides Relation **) |
|
15 |
(************************************************) |
|
16 |
||
17 |
goalw thy [dvd_def] "m dvd 0"; |
|
18 |
by (fast_tac (!claset addIs [mult_0_right RS sym]) 1); |
|
19 |
qed "dvd_0_right"; |
|
20 |
Addsimps [dvd_0_right]; |
|
21 |
||
22 |
goalw thy [dvd_def] "!!m. 0 dvd m ==> m = 0"; |
|
23 |
by (fast_tac (!claset addss !simpset) 1); |
|
24 |
qed "dvd_0_left"; |
|
25 |
||
26 |
goalw thy [dvd_def] "m dvd m"; |
|
27 |
by (fast_tac (!claset addIs [mult_1_right RS sym]) 1); |
|
28 |
qed "dvd_refl"; |
|
29 |
Addsimps [dvd_refl]; |
|
30 |
||
31 |
goalw thy [dvd_def] "!!m n p. [| m dvd n; n dvd p |] ==> m dvd p"; |
|
32 |
by (fast_tac (!claset addIs [mult_assoc] ) 1); |
|
33 |
qed "dvd_trans"; |
|
34 |
||
35 |
goalw thy [dvd_def] "!!m n. [| m dvd n; n dvd m |] ==> m=n"; |
|
36 |
by (fast_tac (!claset addDs [mult_eq_self_implies_10] |
|
37 |
addss (!simpset addsimps [mult_assoc, mult_eq_1_iff])) 1); |
|
38 |
qed "dvd_anti_sym"; |
|
39 |
||
3359 | 40 |
goalw thy [dvd_def] "!!k. [| k dvd m; k dvd n |] ==> k dvd (m + n)"; |
41 |
by (blast_tac (!claset addIs [add_mult_distrib2 RS sym]) 1); |
|
42 |
qed "dvd_add"; |
|
43 |
||
44 |
goalw thy [dvd_def] "!!k. k dvd m ==> k dvd (q * m)"; |
|
45 |
by (blast_tac (!claset addIs [mult_left_commute]) 1); |
|
46 |
qed "dvd_mult"; |
|
47 |
||
48 |
(*Based on a theorem of F. Kammüller's. Doesn't really belong here...*) |
|
49 |
goal Primes.thy "!!C. finite C ==> finite (Union C) --> \ |
|
50 |
\ (! c : C. k dvd card c) --> \ |
|
51 |
\ (! c1: C. ! c2: C. c1 ~= c2 --> c1 Int c2 = {}) \ |
|
52 |
\ --> k dvd card(Union C)"; |
|
53 |
by (etac finite_induct 1); |
|
54 |
by (ALLGOALS Asm_simp_tac); |
|
55 |
by (strip_tac 1); |
|
56 |
by (REPEAT (etac conjE 1)); |
|
57 |
by (stac card_Un_disjoint 1); |
|
58 |
by (ALLGOALS |
|
59 |
(asm_full_simp_tac (!simpset |
|
60 |
addsimps [dvd_add, disjoint_eq_subset_Compl]))); |
|
61 |
by (thin_tac "?PP-->?QQ" 1); |
|
62 |
by (thin_tac "!c:F. ?PP(c)" 1); |
|
63 |
by (thin_tac "!c:F. ?PP(c) & ?QQ(c)" 1); |
|
64 |
by (Step_tac 1); |
|
65 |
by (ball_tac 1); |
|
66 |
by (Blast_tac 1); |
|
67 |
qed "dvd_partition"; |
|
68 |
||
1804 | 69 |
|
70 |
(************************************************) |
|
71 |
(** Greatest Common Divisor **) |
|
72 |
(************************************************) |
|
73 |
||
3242
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
74 |
(* Euclid's Algorithm *) |
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
75 |
|
1804 | 76 |
|
3242
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
77 |
Tfl.tgoalw thy [] gcd.rules; |
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
78 |
by (simp_tac (!simpset addsimps [mod_less_divisor,zero_less_eq]) 1); |
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
79 |
val tc = result(); |
1804 | 80 |
|
3242
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
81 |
val gcd_eq = tc RS hd gcd.rules; |
3270 | 82 |
val gcd_induct = tc RS gcd.induct; |
1804 | 83 |
|
3242
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
84 |
goal thy "gcd(m,0) = m"; |
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
85 |
by (rtac (gcd_eq RS trans) 1); |
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
86 |
by (Simp_tac 1); |
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
87 |
qed "gcd_0"; |
1804 | 88 |
|
3242
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
89 |
goal thy "!!m. 0<n ==> gcd(m,n) = gcd (n, m mod n)"; |
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
90 |
by (rtac (gcd_eq RS trans) 1); |
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
91 |
by (asm_simp_tac (!simpset setloop split_tac [expand_if]) 1); |
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
92 |
qed "gcd_less_0"; |
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
93 |
Addsimps [gcd_0, gcd_less_0]; |
1804 | 94 |
|
3242
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
95 |
goal thy "gcd(m,0) dvd m"; |
1804 | 96 |
by (Simp_tac 1); |
3242
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
97 |
qed "gcd_0_dvd_m"; |
1804 | 98 |
|
3242
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
99 |
goal thy "gcd(m,0) dvd 0"; |
1804 | 100 |
by (Simp_tac 1); |
3242
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
101 |
qed "gcd_0_dvd_0"; |
1804 | 102 |
|
3242
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
103 |
goal thy "!!k. [| k dvd (m mod n); k dvd n; n~=0 |] ==> k dvd m"; |
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
104 |
by (subgoal_tac "k dvd ((m div n)*n + m mod n)" 1); |
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
105 |
by (asm_simp_tac (!simpset addsimps [dvd_add, dvd_mult]) 2); |
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
106 |
by (asm_full_simp_tac (!simpset addsimps [mod_div_equality, zero_less_eq]) 1); |
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
107 |
qed "gcd_recursion"; |
1804 | 108 |
|
109 |
||
3242
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
110 |
(*gcd(m,n) divides m and n. The conjunctions don't seem provable separately*) |
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
111 |
goal thy "(gcd(m,n) dvd m) & (gcd(m,n) dvd n)"; |
3301
cdcc4d5602b6
Now the recdef induction rule variables are named u, v, ...
paulson
parents:
3270
diff
changeset
|
112 |
by (res_inst_tac [("u","m"),("v","n")] gcd_induct 1); |
2102 | 113 |
by (case_tac "n=0" 1); |
3242
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
114 |
by (ALLGOALS |
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
115 |
(asm_simp_tac (!simpset addsimps [mod_less_divisor,zero_less_eq]))); |
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
116 |
by (blast_tac (!claset addDs [gcd_recursion]) 1); |
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
117 |
qed "gcd_divides_both"; |
1804 | 118 |
|
119 |
||
3242
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
120 |
(* if f divides m and n then f divides gcd(m,n) *) |
1804 | 121 |
|
122 |
goalw thy [dvd_def] "!!m. [| f dvd m; f dvd n; 0<n |] ==> f dvd (m mod n)"; |
|
123 |
by (Step_tac 1); |
|
3242
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
124 |
by (full_simp_tac (!simpset addsimps [zero_less_mult_iff]) 1); |
1804 | 125 |
by (res_inst_tac |
126 |
[("x", "(((k div ka)*ka + k mod ka) - ((f*k) div (f*ka)) * ka)")] |
|
127 |
exI 1); |
|
3242
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
128 |
by (asm_simp_tac (!simpset addsimps [diff_mult_distrib2, |
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
129 |
mult_mod_distrib, add_mult_distrib2]) 1); |
1804 | 130 |
qed "dvd_mod"; |
131 |
||
132 |
||
3242
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
133 |
(*Maximality: for all m,n,f naturals, |
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
134 |
if f divides m and f divides n then f divides gcd(m,n)*) |
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
135 |
goal thy "!!k. (f dvd m) & (f dvd n) --> f dvd gcd(m,n)"; |
3301
cdcc4d5602b6
Now the recdef induction rule variables are named u, v, ...
paulson
parents:
3270
diff
changeset
|
136 |
by (res_inst_tac [("u","m"),("v","n")] gcd_induct 1); |
2102 | 137 |
by (case_tac "n=0" 1); |
3242
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
138 |
by (ALLGOALS |
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
139 |
(asm_simp_tac (!simpset addsimps [dvd_mod, mod_less_divisor, |
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
140 |
zero_less_eq]))); |
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
141 |
qed_spec_mp "gcd_greatest"; |
1804 | 142 |
|
3242
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
143 |
(* GCD PROOF : GCD exists and gcd fits the definition *) |
1804 | 144 |
|
3242
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
145 |
goalw thy [is_gcd_def] "is_gcd (gcd(m,n)) m n"; |
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
146 |
by (asm_simp_tac (!simpset addsimps [gcd_greatest, gcd_divides_both]) 1); |
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
147 |
qed "is_gcd"; |
1804 | 148 |
|
149 |
(* GCD is unique *) |
|
150 |
||
3242
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
151 |
goalw thy [is_gcd_def] "is_gcd m a b & is_gcd n a b --> m=n"; |
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
152 |
by (blast_tac (!claset addIs [dvd_anti_sym]) 1); |
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents:
2102
diff
changeset
|
153 |
qed "is_gcd_unique"; |
1804 | 154 |