| author | wenzelm | 
| Wed, 26 Jan 2000 21:10:27 +0100 | |
| changeset 8145 | cdd5386eb6fe | 
| parent 6073 | fba734ba6894 | 
| child 8187 | 535d6253f930 | 
| 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 | |
| 3377 
afa1fedef73f
New results including the basis for unique factorization
 paulson parents: 
3359diff
changeset | 7 | |
| 
afa1fedef73f
New results including the basis for unique factorization
 paulson parents: 
3359diff
changeset | 8 | See H. Davenport, "The Higher Arithmetic". 6th edition. (CUP, 1992) | 
| 1804 | 9 | *) | 
| 10 | ||
| 11 | eta_contract:=false; | |
| 12 | ||
| 13 | (************************************************) | |
| 14 | (** Greatest Common Divisor **) | |
| 15 | (************************************************) | |
| 16 | ||
| 3495 
04739732b13e
New comments on how to deal with unproved termination conditions
 paulson parents: 
3457diff
changeset | 17 | (*** Euclid's Algorithm ***) | 
| 3242 
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
 paulson parents: 
2102diff
changeset | 18 | |
| 1804 | 19 | |
| 3495 
04739732b13e
New comments on how to deal with unproved termination conditions
 paulson parents: 
3457diff
changeset | 20 | (** Prove the termination condition and remove it from the recursion equations | 
| 
04739732b13e
New comments on how to deal with unproved termination conditions
 paulson parents: 
3457diff
changeset | 21 | and induction rule **) | 
| 
04739732b13e
New comments on how to deal with unproved termination conditions
 paulson parents: 
3457diff
changeset | 22 | |
| 3242 
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
 paulson parents: 
2102diff
changeset | 23 | Tfl.tgoalw thy [] gcd.rules; | 
| 4356 | 24 | by (simp_tac (simpset() addsimps [mod_less_divisor]) 1); | 
| 3242 
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
 paulson parents: 
2102diff
changeset | 25 | val tc = result(); | 
| 1804 | 26 | |
| 3242 
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
 paulson parents: 
2102diff
changeset | 27 | val gcd_eq = tc RS hd gcd.rules; | 
| 4812 
d65372e425e5
expandshort;  new gcd_induct with inbuilt case analysis
 paulson parents: 
4809diff
changeset | 28 | |
| 
d65372e425e5
expandshort;  new gcd_induct with inbuilt case analysis
 paulson parents: 
4809diff
changeset | 29 | val prems = goal thy | 
| 
d65372e425e5
expandshort;  new gcd_induct with inbuilt case analysis
 paulson parents: 
4809diff
changeset | 30 | "[| !!m. P m 0; \ | 
| 
d65372e425e5
expandshort;  new gcd_induct with inbuilt case analysis
 paulson parents: 
4809diff
changeset | 31 | \ !!m n. [| 0<n; P n (m mod n) |] ==> P m n \ | 
| 
d65372e425e5
expandshort;  new gcd_induct with inbuilt case analysis
 paulson parents: 
4809diff
changeset | 32 | \ |] ==> P m n"; | 
| 
d65372e425e5
expandshort;  new gcd_induct with inbuilt case analysis
 paulson parents: 
4809diff
changeset | 33 | by (res_inst_tac [("u","m"),("v","n")] (tc RS gcd.induct) 1);
 | 
| 
d65372e425e5
expandshort;  new gcd_induct with inbuilt case analysis
 paulson parents: 
4809diff
changeset | 34 | by (case_tac "n=0" 1); | 
| 
d65372e425e5
expandshort;  new gcd_induct with inbuilt case analysis
 paulson parents: 
4809diff
changeset | 35 | by (asm_simp_tac (simpset() addsimps prems) 1); | 
| 
d65372e425e5
expandshort;  new gcd_induct with inbuilt case analysis
 paulson parents: 
4809diff
changeset | 36 | by Safe_tac; | 
| 
d65372e425e5
expandshort;  new gcd_induct with inbuilt case analysis
 paulson parents: 
4809diff
changeset | 37 | by (asm_simp_tac (simpset() addsimps prems) 1); | 
| 
d65372e425e5
expandshort;  new gcd_induct with inbuilt case analysis
 paulson parents: 
4809diff
changeset | 38 | qed "gcd_induct"; | 
| 
d65372e425e5
expandshort;  new gcd_induct with inbuilt case analysis
 paulson parents: 
4809diff
changeset | 39 | |
| 1804 | 40 | |
| 5069 | 41 | Goal "gcd(m,0) = m"; | 
| 3242 
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
 paulson parents: 
2102diff
changeset | 42 | by (rtac (gcd_eq RS trans) 1); | 
| 
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
 paulson parents: 
2102diff
changeset | 43 | by (Simp_tac 1); | 
| 
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
 paulson parents: 
2102diff
changeset | 44 | qed "gcd_0"; | 
| 4809 | 45 | Addsimps [gcd_0]; | 
| 1804 | 46 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 47 | Goal "0<n ==> gcd(m,n) = gcd (n, m mod n)"; | 
| 3242 
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
 paulson parents: 
2102diff
changeset | 48 | by (rtac (gcd_eq RS trans) 1); | 
| 4686 | 49 | by (Asm_simp_tac 1); | 
| 4809 | 50 | qed "gcd_non_0"; | 
| 1804 | 51 | |
| 5069 | 52 | Goal "gcd(m,1) = 1"; | 
| 4809 | 53 | by (simp_tac (simpset() addsimps [gcd_non_0]) 1); | 
| 54 | qed "gcd_1"; | |
| 55 | Addsimps [gcd_1]; | |
| 1804 | 56 | |
| 3242 
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
 paulson parents: 
2102diff
changeset | 57 | (*gcd(m,n) divides m and n. The conjunctions don't seem provable separately*) | 
| 5069 | 58 | Goal "(gcd(m,n) dvd m) & (gcd(m,n) dvd n)"; | 
| 4812 
d65372e425e5
expandshort;  new gcd_induct with inbuilt case analysis
 paulson parents: 
4809diff
changeset | 59 | by (res_inst_tac [("m","m"),("n","n")] gcd_induct 1);
 | 
| 4809 | 60 | by (ALLGOALS | 
| 61 | (asm_full_simp_tac (simpset() addsimps [gcd_non_0, mod_less_divisor]))); | |
| 4089 | 62 | by (blast_tac (claset() addDs [dvd_mod_imp_dvd]) 1); | 
| 4809 | 63 | qed "gcd_dvd_both"; | 
| 64 | ||
| 65 | bind_thm ("gcd_dvd1", gcd_dvd_both RS conjunct1);
 | |
| 66 | bind_thm ("gcd_dvd2", gcd_dvd_both RS conjunct2);
 | |
| 67 | ||
| 1804 | 68 | |
| 3242 
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
 paulson parents: 
2102diff
changeset | 69 | (*Maximality: for all m,n,f naturals, | 
| 
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
 paulson parents: 
2102diff
changeset | 70 | if f divides m and f divides n then f divides gcd(m,n)*) | 
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 71 | Goal "(f dvd m) --> (f dvd n) --> f dvd gcd(m,n)"; | 
| 4812 
d65372e425e5
expandshort;  new gcd_induct with inbuilt case analysis
 paulson parents: 
4809diff
changeset | 72 | by (res_inst_tac [("m","m"),("n","n")] gcd_induct 1);
 | 
| 4809 | 73 | by (ALLGOALS | 
| 74 | (asm_full_simp_tac (simpset() addsimps[gcd_non_0, dvd_mod, | |
| 75 | mod_less_divisor]))); | |
| 3242 
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
 paulson parents: 
2102diff
changeset | 76 | qed_spec_mp "gcd_greatest"; | 
| 1804 | 77 | |
| 3377 
afa1fedef73f
New results including the basis for unique factorization
 paulson parents: 
3359diff
changeset | 78 | (*Function gcd yields the Greatest Common Divisor*) | 
| 5069 | 79 | Goalw [is_gcd_def] "is_gcd (gcd(m,n)) m n"; | 
| 4809 | 80 | by (asm_simp_tac (simpset() addsimps [gcd_greatest, gcd_dvd_both]) 1); | 
| 3242 
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
 paulson parents: 
2102diff
changeset | 81 | qed "is_gcd"; | 
| 1804 | 82 | |
| 3377 
afa1fedef73f
New results including the basis for unique factorization
 paulson parents: 
3359diff
changeset | 83 | (*uniqueness of GCDs*) | 
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 84 | Goalw [is_gcd_def] "[| is_gcd m a b; is_gcd n a b |] ==> m=n"; | 
| 4089 | 85 | by (blast_tac (claset() addIs [dvd_anti_sym]) 1); | 
| 3242 
406ae5ced4e9
Renamed egcd and gcd; defined the gcd function using TFL
 paulson parents: 
2102diff
changeset | 86 | qed "is_gcd_unique"; | 
| 1804 | 87 | |
| 4809 | 88 | (*USED??*) | 
| 5069 | 89 | Goalw [is_gcd_def] | 
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 90 | "[| is_gcd m a b; k dvd a; k dvd b |] ==> k dvd m"; | 
| 4809 | 91 | by (Blast_tac 1); | 
| 92 | qed "is_gcd_dvd"; | |
| 93 | ||
| 94 | (** Commutativity **) | |
| 95 | ||
| 5069 | 96 | Goalw [is_gcd_def] "is_gcd k m n = is_gcd k n m"; | 
| 4809 | 97 | by (Blast_tac 1); | 
| 98 | qed "is_gcd_commute"; | |
| 99 | ||
| 5069 | 100 | Goal "gcd(m,n) = gcd(n,m)"; | 
| 4812 
d65372e425e5
expandshort;  new gcd_induct with inbuilt case analysis
 paulson parents: 
4809diff
changeset | 101 | by (rtac is_gcd_unique 1); | 
| 
d65372e425e5
expandshort;  new gcd_induct with inbuilt case analysis
 paulson parents: 
4809diff
changeset | 102 | by (rtac is_gcd 2); | 
| 4809 | 103 | by (asm_simp_tac (simpset() addsimps [is_gcd, is_gcd_commute]) 1); | 
| 104 | qed "gcd_commute"; | |
| 105 | ||
| 5069 | 106 | Goal "gcd(gcd(k,m),n) = gcd(k,gcd(m,n))"; | 
| 4812 
d65372e425e5
expandshort;  new gcd_induct with inbuilt case analysis
 paulson parents: 
4809diff
changeset | 107 | by (rtac is_gcd_unique 1); | 
| 
d65372e425e5
expandshort;  new gcd_induct with inbuilt case analysis
 paulson parents: 
4809diff
changeset | 108 | by (rtac is_gcd 2); | 
| 
d65372e425e5
expandshort;  new gcd_induct with inbuilt case analysis
 paulson parents: 
4809diff
changeset | 109 | by (rewtac is_gcd_def); | 
| 4809 | 110 | by (blast_tac (claset() addSIs [gcd_dvd1, gcd_dvd2] | 
| 111 | addIs [gcd_greatest, dvd_trans]) 1); | |
| 112 | qed "gcd_assoc"; | |
| 113 | ||
| 5069 | 114 | Goal "gcd(0,m) = m"; | 
| 4809 | 115 | by (stac gcd_commute 1); | 
| 116 | by (rtac gcd_0 1); | |
| 117 | qed "gcd_0_left"; | |
| 118 | ||
| 5069 | 119 | Goal "gcd(1,m) = 1"; | 
| 4809 | 120 | by (stac gcd_commute 1); | 
| 121 | by (rtac gcd_1 1); | |
| 122 | qed "gcd_1_left"; | |
| 123 | Addsimps [gcd_0_left, gcd_1_left]; | |
| 124 | ||
| 125 | ||
| 126 | (** Multiplication laws **) | |
| 127 | ||
| 3377 
afa1fedef73f
New results including the basis for unique factorization
 paulson parents: 
3359diff
changeset | 128 | (*Davenport, page 27*) | 
| 5069 | 129 | Goal "k * gcd(m,n) = gcd(k*m, k*n)"; | 
| 4812 
d65372e425e5
expandshort;  new gcd_induct with inbuilt case analysis
 paulson parents: 
4809diff
changeset | 130 | by (res_inst_tac [("m","m"),("n","n")] gcd_induct 1);
 | 
| 
d65372e425e5
expandshort;  new gcd_induct with inbuilt case analysis
 paulson parents: 
4809diff
changeset | 131 | by (Asm_full_simp_tac 1); | 
| 3377 
afa1fedef73f
New results including the basis for unique factorization
 paulson parents: 
3359diff
changeset | 132 | by (case_tac "k=0" 1); | 
| 4812 
d65372e425e5
expandshort;  new gcd_induct with inbuilt case analysis
 paulson parents: 
4809diff
changeset | 133 | by (Asm_full_simp_tac 1); | 
| 4809 | 134 | by (asm_full_simp_tac | 
| 135 | (simpset() addsimps [mod_geq, gcd_non_0, mod_mult_distrib2]) 1); | |
| 3377 
afa1fedef73f
New results including the basis for unique factorization
 paulson parents: 
3359diff
changeset | 136 | qed "gcd_mult_distrib2"; | 
| 
afa1fedef73f
New results including the basis for unique factorization
 paulson parents: 
3359diff
changeset | 137 | |
| 5069 | 138 | Goal "gcd(m,m) = m"; | 
| 4809 | 139 | by (cut_inst_tac [("k","m"),("m","1"),("n","1")] gcd_mult_distrib2 1);
 | 
| 140 | by (Asm_full_simp_tac 1); | |
| 141 | qed "gcd_self"; | |
| 142 | Addsimps [gcd_self]; | |
| 143 | ||
| 5069 | 144 | Goal "gcd(k, k*n) = k"; | 
| 4809 | 145 | by (cut_inst_tac [("k","k"),("m","1"),("n","n")] gcd_mult_distrib2 1);
 | 
| 146 | by (Asm_full_simp_tac 1); | |
| 147 | qed "gcd_mult"; | |
| 148 | Addsimps [gcd_mult]; | |
| 149 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 150 | Goal "[| gcd(k,n)=1; k dvd (m*n) |] ==> k dvd m"; | 
| 4809 | 151 | by (subgoal_tac "m = gcd(m*k, m*n)" 1); | 
| 152 | by (etac ssubst 1 THEN rtac gcd_greatest 1); | |
| 153 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [gcd_mult_distrib2 RS sym]))); | |
| 154 | qed "relprime_dvd_mult"; | |
| 155 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 156 | Goalw [prime_def] "[| p: prime; ~ p dvd n |] ==> gcd (p, n) = 1"; | 
| 4809 | 157 | by (cut_inst_tac [("m","p"),("n","n")] gcd_dvd_both 1);
 | 
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 158 | by Auto_tac; | 
| 4809 | 159 | qed "prime_imp_relprime"; | 
| 160 | ||
| 3377 
afa1fedef73f
New results including the basis for unique factorization
 paulson parents: 
3359diff
changeset | 161 | (*This theorem leads immediately to a proof of the uniqueness of factorization. | 
| 
afa1fedef73f
New results including the basis for unique factorization
 paulson parents: 
3359diff
changeset | 162 | If p divides a product of primes then it is one of those primes.*) | 
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 163 | Goal "[| p: prime; p dvd (m*n) |] ==> p dvd m | p dvd n"; | 
| 4809 | 164 | by (blast_tac (claset() addIs [relprime_dvd_mult, prime_imp_relprime]) 1); | 
| 3377 
afa1fedef73f
New results including the basis for unique factorization
 paulson parents: 
3359diff
changeset | 165 | qed "prime_dvd_mult"; | 
| 4809 | 166 | |
| 167 | ||
| 168 | (** Addition laws **) | |
| 169 | ||
| 5069 | 170 | Goal "gcd(m, m+n) = gcd(m,n)"; | 
| 4809 | 171 | by (res_inst_tac [("n1", "m+n")] (gcd_commute RS ssubst) 1);
 | 
| 172 | by (rtac (gcd_eq RS trans) 1); | |
| 4812 
d65372e425e5
expandshort;  new gcd_induct with inbuilt case analysis
 paulson parents: 
4809diff
changeset | 173 | by Auto_tac; | 
| 
d65372e425e5
expandshort;  new gcd_induct with inbuilt case analysis
 paulson parents: 
4809diff
changeset | 174 | by (asm_simp_tac (simpset() addsimps [mod_add_self1]) 1); | 
| 4809 | 175 | by (stac gcd_commute 1); | 
| 176 | by (stac gcd_non_0 1); | |
| 177 | by Safe_tac; | |
| 178 | qed "gcd_add"; | |
| 179 | ||
| 5069 | 180 | Goal "gcd(m, n+m) = gcd(m,n)"; | 
| 4809 | 181 | by (asm_simp_tac (simpset() addsimps [add_commute, gcd_add]) 1); | 
| 182 | qed "gcd_add2"; | |
| 183 | ||
| 184 | ||
| 185 | (** More multiplication laws **) | |
| 186 | ||
| 5069 | 187 | Goal "gcd(m,n) dvd gcd(k*m, n)"; | 
| 4809 | 188 | by (blast_tac (claset() addIs [gcd_greatest, dvd_trans, | 
| 189 | gcd_dvd1, gcd_dvd2]) 1); | |
| 190 | qed "gcd_dvd_gcd_mult"; | |
| 191 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 192 | Goal "gcd(k,n) = 1 ==> gcd(k*m, n) = gcd(m,n)"; | 
| 4812 
d65372e425e5
expandshort;  new gcd_induct with inbuilt case analysis
 paulson parents: 
4809diff
changeset | 193 | by (rtac dvd_anti_sym 1); | 
| 
d65372e425e5
expandshort;  new gcd_induct with inbuilt case analysis
 paulson parents: 
4809diff
changeset | 194 | by (rtac gcd_dvd_gcd_mult 2); | 
| 
d65372e425e5
expandshort;  new gcd_induct with inbuilt case analysis
 paulson parents: 
4809diff
changeset | 195 | by (rtac ([relprime_dvd_mult, gcd_dvd2] MRS gcd_greatest) 1); | 
| 4809 | 196 | by (stac mult_commute 2); | 
| 4812 
d65372e425e5
expandshort;  new gcd_induct with inbuilt case analysis
 paulson parents: 
4809diff
changeset | 197 | by (rtac gcd_dvd1 2); | 
| 4809 | 198 | by (stac gcd_commute 1); | 
| 199 | by (asm_simp_tac (simpset() addsimps [gcd_assoc RS sym]) 1); | |
| 200 | qed "gcd_mult_cancel"; |