src/HOL/ex/Primes.ML
changeset 4809 595f905cc348
parent 4728 b72dd6b2ba56
child 4812 d65372e425e5
equal deleted inserted replaced
4808:995bc5bd8319 4809:595f905cc348
    31 
    31 
    32 goal thy "gcd(m,0) = m";
    32 goal thy "gcd(m,0) = m";
    33 by (rtac (gcd_eq RS trans) 1);
    33 by (rtac (gcd_eq RS trans) 1);
    34 by (Simp_tac 1);
    34 by (Simp_tac 1);
    35 qed "gcd_0";
    35 qed "gcd_0";
       
    36 Addsimps [gcd_0];
    36 
    37 
    37 goal thy "!!m. 0<n ==> gcd(m,n) = gcd (n, m mod n)";
    38 goal thy "!!m. 0<n ==> gcd(m,n) = gcd (n, m mod n)";
    38 by (rtac (gcd_eq RS trans) 1);
    39 by (rtac (gcd_eq RS trans) 1);
    39 by (Asm_simp_tac 1);
    40 by (Asm_simp_tac 1);
    40 by (Blast_tac 1);
    41 by (Blast_tac 1);
    41 qed "gcd_less_0";
    42 qed "gcd_non_0";
    42 Addsimps [gcd_0, gcd_less_0];
       
    43 
    43 
    44 goal thy "gcd(m,0) dvd m";
    44 goal thy "gcd(m,1) = 1";
    45 by (Simp_tac 1);
    45 by (simp_tac (simpset() addsimps [gcd_non_0]) 1);
    46 qed "gcd_0_dvd_m";
    46 qed "gcd_1";
    47 
    47 Addsimps [gcd_1];
    48 goal thy "gcd(m,0) dvd 0";
       
    49 by (Simp_tac 1);
       
    50 qed "gcd_0_dvd_0";
       
    51 
    48 
    52 (*gcd(m,n) divides m and n.  The conjunctions don't seem provable separately*)
    49 (*gcd(m,n) divides m and n.  The conjunctions don't seem provable separately*)
    53 goal thy "(gcd(m,n) dvd m) & (gcd(m,n) dvd n)";
    50 goal thy "(gcd(m,n) dvd m) & (gcd(m,n) dvd n)";
    54 by (res_inst_tac [("u","m"),("v","n")] gcd_induct 1);
    51 by (res_inst_tac [("u","m"),("v","n")] gcd_induct 1);
    55 by (case_tac "n=0" 1);
    52 by (case_tac "n=0" 1);
    56 by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [mod_less_divisor])));
    53 by (ALLGOALS 
       
    54     (asm_full_simp_tac (simpset() addsimps [gcd_non_0, mod_less_divisor])));
    57 by (blast_tac (claset() addDs [dvd_mod_imp_dvd]) 1);
    55 by (blast_tac (claset() addDs [dvd_mod_imp_dvd]) 1);
    58 qed "gcd_divides_both";
    56 qed "gcd_dvd_both";
       
    57 
       
    58 bind_thm ("gcd_dvd1", gcd_dvd_both RS conjunct1);
       
    59 bind_thm ("gcd_dvd2", gcd_dvd_both RS conjunct2);
       
    60 
    59 
    61 
    60 (*Maximality: for all m,n,f naturals, 
    62 (*Maximality: for all m,n,f naturals, 
    61                 if f divides m and f divides n then f divides gcd(m,n)*)
    63                 if f divides m and f divides n then f divides gcd(m,n)*)
    62 goal thy "!!k. (f dvd m) --> (f dvd n) --> f dvd gcd(m,n)";
    64 goal thy "!!k. (f dvd m) --> (f dvd n) --> f dvd gcd(m,n)";
    63 by (res_inst_tac [("u","m"),("v","n")] gcd_induct 1);
    65 by (res_inst_tac [("u","m"),("v","n")] gcd_induct 1);
    64 by (case_tac "n=0" 1);
    66 by (case_tac "n=0" 1);
    65 by (ALLGOALS(asm_full_simp_tac(simpset() addsimps[dvd_mod,mod_less_divisor])));
    67 by (ALLGOALS
       
    68     (asm_full_simp_tac (simpset() addsimps[gcd_non_0, dvd_mod,
       
    69 					   mod_less_divisor])));
    66 qed_spec_mp "gcd_greatest";
    70 qed_spec_mp "gcd_greatest";
    67 
    71 
    68 (*Function gcd yields the Greatest Common Divisor*)
    72 (*Function gcd yields the Greatest Common Divisor*)
    69 goalw thy [is_gcd_def] "is_gcd (gcd(m,n)) m n";
    73 goalw thy [is_gcd_def] "is_gcd (gcd(m,n)) m n";
    70 by (asm_simp_tac (simpset() addsimps [gcd_greatest, gcd_divides_both]) 1);
    74 by (asm_simp_tac (simpset() addsimps [gcd_greatest, gcd_dvd_both]) 1);
    71 qed "is_gcd";
    75 qed "is_gcd";
    72 
    76 
    73 (*uniqueness of GCDs*)
    77 (*uniqueness of GCDs*)
    74 goalw thy [is_gcd_def] "is_gcd m a b & is_gcd n a b --> m=n";
    78 goalw thy [is_gcd_def] "!!m n. [| is_gcd m a b; is_gcd n a b |] ==> m=n";
    75 by (blast_tac (claset() addIs [dvd_anti_sym]) 1);
    79 by (blast_tac (claset() addIs [dvd_anti_sym]) 1);
    76 qed "is_gcd_unique";
    80 qed "is_gcd_unique";
       
    81 
       
    82 (*USED??*)
       
    83 goalw thy [is_gcd_def]
       
    84     "!!m n. [| is_gcd m a b; k dvd a; k dvd b |] ==> k dvd m";
       
    85 by (Blast_tac 1);
       
    86 qed "is_gcd_dvd";
       
    87 
       
    88 (** Commutativity **)
       
    89 
       
    90 goalw thy [is_gcd_def] "is_gcd k m n = is_gcd k n m";
       
    91 by (Blast_tac 1);
       
    92 qed "is_gcd_commute";
       
    93 
       
    94 goal thy "gcd(m,n) = gcd(n,m)";
       
    95 br is_gcd_unique 1;
       
    96 br is_gcd 2;
       
    97 by (asm_simp_tac (simpset() addsimps [is_gcd, is_gcd_commute]) 1);
       
    98 qed "gcd_commute";
       
    99 
       
   100 goal thy "gcd(gcd(k,m),n) = gcd(k,gcd(m,n))";
       
   101 br is_gcd_unique 1;
       
   102 br is_gcd 2;
       
   103 bw is_gcd_def;
       
   104 by (blast_tac (claset() addSIs [gcd_dvd1, gcd_dvd2]
       
   105    	                addIs  [gcd_greatest, dvd_trans]) 1);
       
   106 qed "gcd_assoc";
       
   107 
       
   108 goal thy "gcd(0,m) = m";
       
   109 by (stac gcd_commute 1);
       
   110 by (rtac gcd_0 1);
       
   111 qed "gcd_0_left";
       
   112 
       
   113 goal thy "gcd(1,m) = 1";
       
   114 by (stac gcd_commute 1);
       
   115 by (rtac gcd_1 1);
       
   116 qed "gcd_1_left";
       
   117 Addsimps [gcd_0_left, gcd_1_left];
       
   118 
       
   119 
       
   120 (** Multiplication laws **)
    77 
   121 
    78 (*Davenport, page 27*)
   122 (*Davenport, page 27*)
    79 goal thy "k * gcd(m,n) = gcd(k*m, k*n)";
   123 goal thy "k * gcd(m,n) = gcd(k*m, k*n)";
    80 by (res_inst_tac [("u","m"),("v","n")] gcd_induct 1);
   124 by (res_inst_tac [("u","m"),("v","n")] gcd_induct 1);
    81 by (case_tac "k=0" 1);
   125 by (case_tac "k=0" 1);
    82  by(Asm_full_simp_tac 1);
   126  by(Asm_full_simp_tac 1);
    83 by (case_tac "n=0" 1);
   127 by (case_tac "n=0" 1);
    84  by(Asm_full_simp_tac 1);
   128  by(Asm_full_simp_tac 1);
    85 by(asm_full_simp_tac (simpset() addsimps [mod_geq, mod_mult_distrib2]) 1);
   129 by (asm_full_simp_tac
       
   130     (simpset() addsimps [mod_geq, gcd_non_0, mod_mult_distrib2]) 1);
    86 qed "gcd_mult_distrib2";
   131 qed "gcd_mult_distrib2";
       
   132 
       
   133 goal thy "gcd(m,m) = m";
       
   134 by (cut_inst_tac [("k","m"),("m","1"),("n","1")] gcd_mult_distrib2 1);
       
   135 by (Asm_full_simp_tac 1);
       
   136 qed "gcd_self";
       
   137 Addsimps [gcd_self];
       
   138 
       
   139 goal thy "gcd(k, k*n) = k";
       
   140 by (cut_inst_tac [("k","k"),("m","1"),("n","n")] gcd_mult_distrib2 1);
       
   141 by (Asm_full_simp_tac 1);
       
   142 qed "gcd_mult";
       
   143 Addsimps [gcd_mult];
       
   144 
       
   145 goal thy "!!k. [| gcd(k,n)=1; k dvd (m*n) |] ==> k dvd m";
       
   146 by (subgoal_tac "m = gcd(m*k, m*n)" 1);
       
   147 by (etac ssubst 1 THEN rtac gcd_greatest 1);
       
   148 by (ALLGOALS (asm_simp_tac (simpset() addsimps [gcd_mult_distrib2 RS sym])));
       
   149 qed "relprime_dvd_mult";
       
   150 
       
   151 goalw thy [prime_def] "!!p. [| p: prime;  ~ p dvd n |] ==> gcd (p, n) = 1";
       
   152 by (cut_inst_tac [("m","p"),("n","n")] gcd_dvd_both 1);
       
   153 by (fast_tac (claset() addSss (simpset())) 1);
       
   154 qed "prime_imp_relprime";
    87 
   155 
    88 (*This theorem leads immediately to a proof of the uniqueness of factorization.
   156 (*This theorem leads immediately to a proof of the uniqueness of factorization.
    89   If p divides a product of primes then it is one of those primes.*)
   157   If p divides a product of primes then it is one of those primes.*)
    90 goalw thy [prime_def] "!!p. [| p: prime; p dvd (m*n) |] ==> p dvd m | p dvd n";
   158 goal thy "!!p. [| p: prime; p dvd (m*n) |] ==> p dvd m | p dvd n";
    91 by (Clarify_tac 1);
   159 by (blast_tac (claset() addIs [relprime_dvd_mult, prime_imp_relprime]) 1);
    92 by (subgoal_tac "m = gcd(m*p, m*n)" 1);
       
    93 by (etac ssubst 1);
       
    94 by (rtac gcd_greatest 1);
       
    95 by (ALLGOALS (asm_simp_tac (simpset() addsimps [gcd_mult_distrib2 RS sym])));
       
    96 (*Now deduce  gcd(p,n)=1  to finish the proof*)
       
    97 by (cut_inst_tac [("m","p"),("n","n")] gcd_divides_both 1);
       
    98 by (fast_tac (claset() addSss (simpset())) 1);
       
    99 qed "prime_dvd_mult";
   160 qed "prime_dvd_mult";
       
   161 
       
   162 
       
   163 (** Addition laws **)
       
   164 
       
   165 goal thy "gcd(m, m+n) = gcd(m,n)";
       
   166 by (res_inst_tac [("n1", "m+n")] (gcd_commute RS ssubst) 1);
       
   167 by (rtac (gcd_eq RS trans) 1);
       
   168 auto();
       
   169 by (asm_simp_tac (simpset() addsimps [mod_add_cancel1]) 1);
       
   170 by (stac gcd_commute 1);
       
   171 by (stac gcd_non_0 1);
       
   172 by Safe_tac;
       
   173 qed "gcd_add";
       
   174 
       
   175 goal thy "gcd(m, n+m) = gcd(m,n)";
       
   176 by (asm_simp_tac (simpset() addsimps [add_commute, gcd_add]) 1);
       
   177 qed "gcd_add2";
       
   178 
       
   179 
       
   180 (** More multiplication laws **)
       
   181 
       
   182 goal thy "gcd(m,n) dvd gcd(k*m, n)";
       
   183 by (blast_tac (claset() addIs [gcd_greatest, dvd_trans, 
       
   184                                gcd_dvd1, gcd_dvd2]) 1);
       
   185 qed "gcd_dvd_gcd_mult";
       
   186 
       
   187 goal thy "!!n. gcd(k,n) = 1 ==> gcd(k*m, n) = gcd(m,n)";
       
   188 br dvd_anti_sym 1;
       
   189 br gcd_dvd_gcd_mult 2;
       
   190 br ([relprime_dvd_mult, gcd_dvd2] MRS gcd_greatest) 1;
       
   191 by (stac mult_commute 2);
       
   192 br gcd_dvd1 2;
       
   193 by (stac gcd_commute 1);
       
   194 by (asm_simp_tac (simpset() addsimps [gcd_assoc RS sym]) 1);
       
   195 qed "gcd_mult_cancel";