paulson@1804: (* Title: HOL/ex/Primes.ML paulson@1804: ID: $Id$ paulson@1804: Author: Christophe Tabacznyj and Lawrence C Paulson paulson@1804: Copyright 1996 University of Cambridge paulson@1804: paulson@1804: The "divides" relation, the greatest common divisor and Euclid's algorithm paulson@3377: paulson@3377: See H. Davenport, "The Higher Arithmetic". 6th edition. (CUP, 1992) paulson@1804: *) paulson@1804: paulson@1804: eta_contract:=false; paulson@1804: paulson@1804: open Primes; paulson@1804: paulson@1804: (************************************************) paulson@1804: (** Greatest Common Divisor **) paulson@1804: (************************************************) paulson@1804: paulson@3495: (*** Euclid's Algorithm ***) paulson@3242: paulson@1804: paulson@3495: (** Prove the termination condition and remove it from the recursion equations paulson@3495: and induction rule **) paulson@3495: paulson@3242: Tfl.tgoalw thy [] gcd.rules; wenzelm@4089: by (simp_tac (simpset() addsimps [mod_less_divisor, zero_less_eq]) 1); paulson@3242: val tc = result(); paulson@1804: paulson@3242: val gcd_eq = tc RS hd gcd.rules; paulson@3270: val gcd_induct = tc RS gcd.induct; paulson@1804: paulson@3242: goal thy "gcd(m,0) = m"; paulson@3242: by (rtac (gcd_eq RS trans) 1); paulson@3242: by (Simp_tac 1); paulson@3242: qed "gcd_0"; paulson@1804: paulson@3242: goal thy "!!m. 0 gcd(m,n) = gcd (n, m mod n)"; paulson@3242: by (rtac (gcd_eq RS trans) 1); wenzelm@4089: by (asm_simp_tac (simpset() addsplits [expand_if]) 1); paulson@3242: qed "gcd_less_0"; paulson@3242: Addsimps [gcd_0, gcd_less_0]; paulson@1804: paulson@3242: goal thy "gcd(m,0) dvd m"; paulson@1804: by (Simp_tac 1); paulson@3242: qed "gcd_0_dvd_m"; paulson@1804: paulson@3242: goal thy "gcd(m,0) dvd 0"; paulson@1804: by (Simp_tac 1); paulson@3242: qed "gcd_0_dvd_0"; paulson@1804: paulson@3242: (*gcd(m,n) divides m and n. The conjunctions don't seem provable separately*) paulson@3242: goal thy "(gcd(m,n) dvd m) & (gcd(m,n) dvd n)"; paulson@3301: by (res_inst_tac [("u","m"),("v","n")] gcd_induct 1); paulson@2102: by (case_tac "n=0" 1); paulson@3242: by (ALLGOALS wenzelm@4089: (asm_simp_tac (simpset() addsimps [mod_less_divisor,zero_less_eq]))); wenzelm@4089: by (blast_tac (claset() addDs [dvd_mod_imp_dvd]) 1); paulson@3242: qed "gcd_divides_both"; paulson@1804: paulson@3242: (*Maximality: for all m,n,f naturals, paulson@3242: if f divides m and f divides n then f divides gcd(m,n)*) paulson@3377: goal thy "!!k. (f dvd m) --> (f dvd n) --> f dvd gcd(m,n)"; paulson@3301: by (res_inst_tac [("u","m"),("v","n")] gcd_induct 1); paulson@2102: by (case_tac "n=0" 1); paulson@3242: by (ALLGOALS wenzelm@4089: (asm_simp_tac (simpset() addsimps [dvd_mod, mod_less_divisor, paulson@3242: zero_less_eq]))); paulson@3242: qed_spec_mp "gcd_greatest"; paulson@1804: paulson@3377: (*Function gcd yields the Greatest Common Divisor*) paulson@3242: goalw thy [is_gcd_def] "is_gcd (gcd(m,n)) m n"; wenzelm@4089: by (asm_simp_tac (simpset() addsimps [gcd_greatest, gcd_divides_both]) 1); paulson@3242: qed "is_gcd"; paulson@1804: paulson@3377: (*uniqueness of GCDs*) paulson@3242: goalw thy [is_gcd_def] "is_gcd m a b & is_gcd n a b --> m=n"; wenzelm@4089: by (blast_tac (claset() addIs [dvd_anti_sym]) 1); paulson@3242: qed "is_gcd_unique"; paulson@1804: paulson@3377: (*Davenport, page 27*) paulson@3377: goal thy "k * gcd(m,n) = gcd(k*m, k*n)"; paulson@3377: by (res_inst_tac [("u","m"),("v","n")] gcd_induct 1); paulson@3377: by (case_tac "k=0" 1); paulson@3377: by (case_tac "n=0" 2); paulson@3377: by (ALLGOALS wenzelm@4089: (asm_simp_tac (simpset() addsimps [mod_less_divisor, zero_less_eq, paulson@3377: mod_geq, mod_mult_distrib2]))); paulson@3377: qed "gcd_mult_distrib2"; paulson@3377: paulson@3377: (*This theorem leads immediately to a proof of the uniqueness of factorization. paulson@3377: If p divides a product of primes then it is one of those primes.*) paulson@3377: goalw thy [prime_def] "!!p. [| p: prime; p dvd (m*n) |] ==> p dvd m | p dvd n"; paulson@3718: by (Clarify_tac 1); paulson@3377: by (subgoal_tac "m = gcd(m*p, m*n)" 1); paulson@3457: by (etac ssubst 1); paulson@3457: by (rtac gcd_greatest 1); wenzelm@4089: by (ALLGOALS (asm_simp_tac (simpset() addsimps [gcd_mult_distrib2 RS sym]))); paulson@3377: (*Now deduce gcd(p,n)=1 to finish the proof*) paulson@3377: by (cut_inst_tac [("m","p"),("n","n")] gcd_divides_both 1); wenzelm@4089: by (fast_tac (claset() addSss (simpset())) 1); paulson@3377: qed "prime_dvd_mult";