src/HOL/Computational_Algebra/Primes.thy
changeset 65417 fc41a5650fb1
parent 65025 8c32ce2a01c6
child 65435 378175f44328
equal deleted inserted replaced
65416:f707dbcf11e3 65417:fc41a5650fb1
       
     1 (*  Authors:    Christophe Tabacznyj, Lawrence C. Paulson, Amine Chaieb,
       
     2                 Thomas M. Rasmussen, Jeremy Avigad, Tobias Nipkow, 
       
     3                 Manuel Eberl
       
     4 
       
     5 
       
     6 This file deals with properties of primes. Definitions and lemmas are
       
     7 proved uniformly for the natural numbers and integers.
       
     8 
       
     9 This file combines and revises a number of prior developments.
       
    10 
       
    11 The original theories "GCD" and "Primes" were by Christophe Tabacznyj
       
    12 and Lawrence C. Paulson, based on @{cite davenport92}. They introduced
       
    13 gcd, lcm, and prime for the natural numbers.
       
    14 
       
    15 The original theory "IntPrimes" was by Thomas M. Rasmussen, and
       
    16 extended gcd, lcm, primes to the integers. Amine Chaieb provided
       
    17 another extension of the notions to the integers, and added a number
       
    18 of results to "Primes" and "GCD". IntPrimes also defined and developed
       
    19 the congruence relations on the integers. The notion was extended to
       
    20 the natural numbers by Chaieb.
       
    21 
       
    22 Jeremy Avigad combined all of these, made everything uniform for the
       
    23 natural numbers and the integers, and added a number of new theorems.
       
    24 
       
    25 Tobias Nipkow cleaned up a lot.
       
    26 
       
    27 Florian Haftmann and Manuel Eberl put primality and prime factorisation
       
    28 onto an algebraic foundation and thus generalised these concepts to 
       
    29 other rings, such as polynomials. (see also the Factorial_Ring theory).
       
    30 
       
    31 There were also previous formalisations of unique factorisation by 
       
    32 Thomas Marthedal Rasmussen, Jeremy Avigad, and David Gray.
       
    33 *)
       
    34 
       
    35 
       
    36 section \<open>Primes\<close>
       
    37 
       
    38 theory Primes
       
    39 imports "~~/src/HOL/Binomial" Euclidean_Algorithm
       
    40 begin
       
    41 
       
    42 (* As a simp or intro rule,
       
    43 
       
    44      prime p \<Longrightarrow> p > 0
       
    45 
       
    46    wreaks havoc here. When the premise includes \<forall>x \<in># M. prime x, it
       
    47    leads to the backchaining
       
    48 
       
    49      x > 0
       
    50      prime x
       
    51      x \<in># M   which is, unfortunately,
       
    52      count M x > 0
       
    53 *)
       
    54 
       
    55 declare [[coercion int]]
       
    56 declare [[coercion_enabled]]
       
    57 
       
    58 lemma prime_elem_nat_iff:
       
    59   "prime_elem (n :: nat) \<longleftrightarrow> (1 < n \<and> (\<forall>m. m dvd n \<longrightarrow> m = 1 \<or> m = n))"
       
    60 proof safe
       
    61   assume *: "prime_elem n"
       
    62   from * have "n \<noteq> 0" "n \<noteq> 1" by (intro notI, simp del: One_nat_def)+
       
    63   thus "n > 1" by (cases n) simp_all
       
    64   fix m assume m: "m dvd n" "m \<noteq> n"
       
    65   from * \<open>m dvd n\<close> have "n dvd m \<or> is_unit m"
       
    66     by (intro irreducibleD' prime_elem_imp_irreducible)
       
    67   with m show "m = 1" by (auto dest: dvd_antisym)
       
    68 next
       
    69   assume "n > 1" "\<forall>m. m dvd n \<longrightarrow> m = 1 \<or> m = n"
       
    70   thus "prime_elem n"
       
    71     by (auto simp: prime_elem_iff_irreducible irreducible_altdef)
       
    72 qed
       
    73 
       
    74 lemma prime_nat_iff_prime_elem: "prime (n :: nat) \<longleftrightarrow> prime_elem n"
       
    75   by (simp add: prime_def)
       
    76 
       
    77 lemma prime_nat_iff:
       
    78   "prime (n :: nat) \<longleftrightarrow> (1 < n \<and> (\<forall>m. m dvd n \<longrightarrow> m = 1 \<or> m = n))"
       
    79   by (simp add: prime_nat_iff_prime_elem prime_elem_nat_iff)
       
    80 
       
    81 lemma prime_elem_int_nat_transfer: "prime_elem (n::int) \<longleftrightarrow> prime_elem (nat (abs n))"
       
    82 proof
       
    83   assume "prime_elem n"
       
    84   thus "prime_elem (nat (abs n))" by (auto simp: prime_elem_def nat_dvd_iff)
       
    85 next
       
    86   assume "prime_elem (nat (abs n))"
       
    87   thus "prime_elem n"
       
    88     by (auto simp: dvd_int_unfold_dvd_nat prime_elem_def abs_mult nat_mult_distrib)
       
    89 qed
       
    90 
       
    91 lemma prime_elem_nat_int_transfer [simp]: "prime_elem (int n) \<longleftrightarrow> prime_elem n"
       
    92   by (auto simp: prime_elem_int_nat_transfer)
       
    93 
       
    94 lemma prime_nat_int_transfer [simp]: "prime (int n) \<longleftrightarrow> prime n"
       
    95   by (auto simp: prime_elem_int_nat_transfer prime_def)
       
    96 
       
    97 lemma prime_int_nat_transfer: "prime (n::int) \<longleftrightarrow> n \<ge> 0 \<and> prime (nat n)"
       
    98   by (auto simp: prime_elem_int_nat_transfer prime_def)
       
    99 
       
   100 lemma prime_int_iff:
       
   101   "prime (n::int) \<longleftrightarrow> (1 < n \<and> (\<forall>m. m \<ge> 0 \<and> m dvd n \<longrightarrow> m = 1 \<or> m = n))"
       
   102 proof (intro iffI conjI allI impI; (elim conjE)?)
       
   103   assume *: "prime n"
       
   104   hence irred: "irreducible n" by (simp add: prime_elem_imp_irreducible)
       
   105   from * have "n \<ge> 0" "n \<noteq> 0" "n \<noteq> 1" 
       
   106     by (auto simp: prime_def zabs_def not_less split: if_splits)
       
   107   thus "n > 1" by presburger
       
   108   fix m assume "m dvd n" \<open>m \<ge> 0\<close>
       
   109   with irred have "m dvd 1 \<or> n dvd m" by (auto simp: irreducible_altdef)
       
   110   with \<open>m dvd n\<close> \<open>m \<ge> 0\<close> \<open>n > 1\<close> show "m = 1 \<or> m = n"
       
   111     using associated_iff_dvd[of m n] by auto
       
   112 next
       
   113   assume n: "1 < n" "\<forall>m. m \<ge> 0 \<and> m dvd n \<longrightarrow> m = 1 \<or> m = n"
       
   114   hence "nat n > 1" by simp
       
   115   moreover have "\<forall>m. m dvd nat n \<longrightarrow> m = 1 \<or> m = nat n"
       
   116   proof (intro allI impI)
       
   117     fix m assume "m dvd nat n"
       
   118     with \<open>n > 1\<close> have "int m dvd n" by (auto simp: int_dvd_iff)
       
   119     with n(2) have "int m = 1 \<or> int m = n" by auto
       
   120     thus "m = 1 \<or> m = nat n" by auto
       
   121   qed
       
   122   ultimately show "prime n" 
       
   123     unfolding prime_int_nat_transfer prime_nat_iff by auto
       
   124 qed
       
   125 
       
   126 
       
   127 lemma prime_nat_not_dvd:
       
   128   assumes "prime p" "p > n" "n \<noteq> (1::nat)"
       
   129   shows   "\<not>n dvd p"
       
   130 proof
       
   131   assume "n dvd p"
       
   132   from assms(1) have "irreducible p" by (simp add: prime_elem_imp_irreducible)
       
   133   from irreducibleD'[OF this \<open>n dvd p\<close>] \<open>n dvd p\<close> \<open>p > n\<close> assms show False
       
   134     by (cases "n = 0") (auto dest!: dvd_imp_le)
       
   135 qed
       
   136 
       
   137 lemma prime_int_not_dvd:
       
   138   assumes "prime p" "p > n" "n > (1::int)"
       
   139   shows   "\<not>n dvd p"
       
   140 proof
       
   141   assume "n dvd p"
       
   142   from assms(1) have "irreducible p" by (simp add: prime_elem_imp_irreducible)
       
   143   from irreducibleD'[OF this \<open>n dvd p\<close>] \<open>n dvd p\<close> \<open>p > n\<close> assms show False
       
   144     by (auto dest!: zdvd_imp_le)
       
   145 qed
       
   146 
       
   147 lemma prime_odd_nat: "prime p \<Longrightarrow> p > (2::nat) \<Longrightarrow> odd p"
       
   148   by (intro prime_nat_not_dvd) auto
       
   149 
       
   150 lemma prime_odd_int: "prime p \<Longrightarrow> p > (2::int) \<Longrightarrow> odd p"
       
   151   by (intro prime_int_not_dvd) auto
       
   152 
       
   153 lemma prime_ge_0_int: "prime p \<Longrightarrow> p \<ge> (0::int)"
       
   154   unfolding prime_int_iff by auto
       
   155 
       
   156 lemma prime_gt_0_nat: "prime p \<Longrightarrow> p > (0::nat)"
       
   157   unfolding prime_nat_iff by auto
       
   158 
       
   159 lemma prime_gt_0_int: "prime p \<Longrightarrow> p > (0::int)"
       
   160   unfolding prime_int_iff by auto
       
   161 
       
   162 lemma prime_ge_1_nat: "prime p \<Longrightarrow> p \<ge> (1::nat)"
       
   163   unfolding prime_nat_iff by auto
       
   164 
       
   165 lemma prime_ge_Suc_0_nat: "prime p \<Longrightarrow> p \<ge> Suc 0"
       
   166   unfolding prime_nat_iff by auto
       
   167 
       
   168 lemma prime_ge_1_int: "prime p \<Longrightarrow> p \<ge> (1::int)"
       
   169   unfolding prime_int_iff by auto
       
   170 
       
   171 lemma prime_gt_1_nat: "prime p \<Longrightarrow> p > (1::nat)"
       
   172   unfolding prime_nat_iff by auto
       
   173 
       
   174 lemma prime_gt_Suc_0_nat: "prime p \<Longrightarrow> p > Suc 0"
       
   175   unfolding prime_nat_iff by auto
       
   176 
       
   177 lemma prime_gt_1_int: "prime p \<Longrightarrow> p > (1::int)"
       
   178   unfolding prime_int_iff by auto
       
   179 
       
   180 lemma prime_ge_2_nat: "prime p \<Longrightarrow> p \<ge> (2::nat)"
       
   181   unfolding prime_nat_iff by auto
       
   182 
       
   183 lemma prime_ge_2_int: "prime p \<Longrightarrow> p \<ge> (2::int)"
       
   184   unfolding prime_int_iff by auto
       
   185 
       
   186 lemma prime_int_altdef:
       
   187   "prime p = (1 < p \<and> (\<forall>m::int. m \<ge> 0 \<longrightarrow> m dvd p \<longrightarrow>
       
   188     m = 1 \<or> m = p))"
       
   189   unfolding prime_int_iff by blast
       
   190 
       
   191 lemma not_prime_eq_prod_nat:
       
   192   assumes "m > 1" "\<not>prime (m::nat)"
       
   193   shows   "\<exists>n k. n = m * k \<and> 1 < m \<and> m < n \<and> 1 < k \<and> k < n"
       
   194   using assms irreducible_altdef[of m]
       
   195   by (auto simp: prime_elem_iff_irreducible prime_def irreducible_altdef)
       
   196 
       
   197     
       
   198 subsection\<open>Largest exponent of a prime factor\<close>
       
   199 text\<open>Possibly duplicates other material, but avoid the complexities of multisets.\<close>
       
   200   
       
   201 lemma prime_power_cancel_less:
       
   202   assumes "prime p" and eq: "m * (p ^ k) = m' * (p ^ k')" and less: "k < k'" and "\<not> p dvd m"
       
   203   shows False
       
   204 proof -
       
   205   obtain l where l: "k' = k + l" and "l > 0"
       
   206     using less less_imp_add_positive by auto
       
   207   have "m = m * (p ^ k) div (p ^ k)"
       
   208     using \<open>prime p\<close> by simp
       
   209   also have "\<dots> = m' * (p ^ k') div (p ^ k)"
       
   210     using eq by simp
       
   211   also have "\<dots> = m' * (p ^ l) * (p ^ k) div (p ^ k)"
       
   212     by (simp add: l mult.commute mult.left_commute power_add)
       
   213   also have "... = m' * (p ^ l)"
       
   214     using \<open>prime p\<close> by simp
       
   215   finally have "p dvd m"
       
   216     using \<open>l > 0\<close> by simp
       
   217   with assms show False
       
   218     by simp
       
   219 qed
       
   220 
       
   221 lemma prime_power_cancel:
       
   222   assumes "prime p" and eq: "m * (p ^ k) = m' * (p ^ k')" and "\<not> p dvd m" "\<not> p dvd m'"
       
   223   shows "k = k'"
       
   224   using prime_power_cancel_less [OF \<open>prime p\<close>] assms
       
   225   by (metis linorder_neqE_nat)
       
   226 
       
   227 lemma prime_power_cancel2:
       
   228   assumes "prime p" "m * (p ^ k) = m' * (p ^ k')" "\<not> p dvd m" "\<not> p dvd m'"
       
   229   obtains "m = m'" "k = k'"
       
   230   using prime_power_cancel [OF assms] assms by auto
       
   231 
       
   232 lemma prime_power_canonical:
       
   233   fixes m::nat
       
   234   assumes "prime p" "m > 0"
       
   235   shows "\<exists>k n. \<not> p dvd n \<and> m = n * p^k"
       
   236 using \<open>m > 0\<close>
       
   237 proof (induction m rule: less_induct)
       
   238   case (less m)
       
   239   show ?case
       
   240   proof (cases "p dvd m")
       
   241     case True
       
   242     then obtain m' where m': "m = p * m'"
       
   243       using dvdE by blast
       
   244     with \<open>prime p\<close> have "0 < m'" "m' < m"
       
   245       using less.prems prime_nat_iff by auto
       
   246     with m' less show ?thesis
       
   247       by (metis power_Suc mult.left_commute)
       
   248   next
       
   249     case False
       
   250     then show ?thesis
       
   251       by (metis mult.right_neutral power_0)
       
   252   qed
       
   253 qed
       
   254 
       
   255 
       
   256 subsubsection \<open>Make prime naively executable\<close>
       
   257 
       
   258 lemma Suc_0_not_prime_nat [simp]: "~prime (Suc 0)"
       
   259   unfolding One_nat_def [symmetric] by (rule not_prime_1)
       
   260 
       
   261 lemma prime_nat_iff':
       
   262   "prime (p :: nat) \<longleftrightarrow> p > 1 \<and> (\<forall>n \<in> {2..<p}. ~ n dvd p)"
       
   263 proof safe
       
   264   assume "p > 1" and *: "\<forall>n\<in>{2..<p}. \<not>n dvd p"
       
   265   show "prime p" unfolding prime_nat_iff
       
   266   proof (intro conjI allI impI)
       
   267     fix m assume "m dvd p"
       
   268     with \<open>p > 1\<close> have "m \<noteq> 0" by (intro notI) auto
       
   269     hence "m \<ge> 1" by simp
       
   270     moreover from \<open>m dvd p\<close> and * have "m \<notin> {2..<p}" by blast
       
   271     with \<open>m dvd p\<close> and \<open>p > 1\<close> have "m \<le> 1 \<or> m = p" by (auto dest: dvd_imp_le)
       
   272     ultimately show "m = 1 \<or> m = p" by simp
       
   273   qed fact+
       
   274 qed (auto simp: prime_nat_iff)
       
   275 
       
   276 lemma prime_int_iff':
       
   277   "prime (p :: int) \<longleftrightarrow> p > 1 \<and> (\<forall>n \<in> {2..<p}. ~ n dvd p)" (is "?lhs = ?rhs")
       
   278 proof
       
   279   assume "?lhs"
       
   280   thus "?rhs"
       
   281       by (auto simp: prime_int_nat_transfer dvd_int_unfold_dvd_nat prime_nat_iff')
       
   282 next
       
   283   assume "?rhs"
       
   284   thus "?lhs"
       
   285     by (auto simp: prime_int_nat_transfer zdvd_int prime_nat_iff')
       
   286 qed
       
   287 
       
   288 lemma prime_int_numeral_eq [simp]:
       
   289   "prime (numeral m :: int) \<longleftrightarrow> prime (numeral m :: nat)"
       
   290   by (simp add: prime_int_nat_transfer)
       
   291 
       
   292 lemma two_is_prime_nat [simp]: "prime (2::nat)"
       
   293   by (simp add: prime_nat_iff')
       
   294 
       
   295 lemma prime_nat_numeral_eq [simp]:
       
   296   "prime (numeral m :: nat) \<longleftrightarrow>
       
   297     (1::nat) < numeral m \<and>
       
   298     (\<forall>n::nat \<in> set [2..<numeral m]. \<not> n dvd numeral m)"
       
   299   by (simp only: prime_nat_iff' set_upt)  \<comment> \<open>TODO Sieve Of Erathosthenes might speed this up\<close>
       
   300 
       
   301 
       
   302 text\<open>A bit of regression testing:\<close>
       
   303 
       
   304 lemma "prime(97::nat)" by simp
       
   305 lemma "prime(97::int)" by simp
       
   306 
       
   307 lemma prime_factor_nat: 
       
   308   "n \<noteq> (1::nat) \<Longrightarrow> \<exists>p. prime p \<and> p dvd n"
       
   309   using prime_divisor_exists[of n]
       
   310   by (cases "n = 0") (auto intro: exI[of _ "2::nat"])
       
   311 
       
   312 
       
   313 subsection \<open>Infinitely many primes\<close>
       
   314 
       
   315 lemma next_prime_bound: "\<exists>p::nat. prime p \<and> n < p \<and> p \<le> fact n + 1"
       
   316 proof-
       
   317   have f1: "fact n + 1 \<noteq> (1::nat)" using fact_ge_1 [of n, where 'a=nat] by arith
       
   318   from prime_factor_nat [OF f1]
       
   319   obtain p :: nat where "prime p" and "p dvd fact n + 1" by auto
       
   320   then have "p \<le> fact n + 1" apply (intro dvd_imp_le) apply auto done
       
   321   { assume "p \<le> n"
       
   322     from \<open>prime p\<close> have "p \<ge> 1"
       
   323       by (cases p, simp_all)
       
   324     with \<open>p <= n\<close> have "p dvd fact n"
       
   325       by (intro dvd_fact)
       
   326     with \<open>p dvd fact n + 1\<close> have "p dvd fact n + 1 - fact n"
       
   327       by (rule dvd_diff_nat)
       
   328     then have "p dvd 1" by simp
       
   329     then have "p <= 1" by auto
       
   330     moreover from \<open>prime p\<close> have "p > 1"
       
   331       using prime_nat_iff by blast
       
   332     ultimately have False by auto}
       
   333   then have "n < p" by presburger
       
   334   with \<open>prime p\<close> and \<open>p <= fact n + 1\<close> show ?thesis by auto
       
   335 qed
       
   336 
       
   337 lemma bigger_prime: "\<exists>p. prime p \<and> p > (n::nat)"
       
   338   using next_prime_bound by auto
       
   339 
       
   340 lemma primes_infinite: "\<not> (finite {(p::nat). prime p})"
       
   341 proof
       
   342   assume "finite {(p::nat). prime p}"
       
   343   with Max_ge have "(EX b. (ALL x : {(p::nat). prime p}. x <= b))"
       
   344     by auto
       
   345   then obtain b where "ALL (x::nat). prime x \<longrightarrow> x <= b"
       
   346     by auto
       
   347   with bigger_prime [of b] show False
       
   348     by auto
       
   349 qed
       
   350 
       
   351 subsection\<open>Powers of Primes\<close>
       
   352 
       
   353 text\<open>Versions for type nat only\<close>
       
   354 
       
   355 lemma prime_product:
       
   356   fixes p::nat
       
   357   assumes "prime (p * q)"
       
   358   shows "p = 1 \<or> q = 1"
       
   359 proof -
       
   360   from assms have
       
   361     "1 < p * q" and P: "\<And>m. m dvd p * q \<Longrightarrow> m = 1 \<or> m = p * q"
       
   362     unfolding prime_nat_iff by auto
       
   363   from \<open>1 < p * q\<close> have "p \<noteq> 0" by (cases p) auto
       
   364   then have Q: "p = p * q \<longleftrightarrow> q = 1" by auto
       
   365   have "p dvd p * q" by simp
       
   366   then have "p = 1 \<or> p = p * q" by (rule P)
       
   367   then show ?thesis by (simp add: Q)
       
   368 qed
       
   369 
       
   370 (* TODO: Generalise? *)
       
   371 lemma prime_power_mult_nat:
       
   372   fixes p::nat
       
   373   assumes p: "prime p" and xy: "x * y = p ^ k"
       
   374   shows "\<exists>i j. x = p ^i \<and> y = p^ j"
       
   375 using xy
       
   376 proof(induct k arbitrary: x y)
       
   377   case 0 thus ?case apply simp by (rule exI[where x="0"], simp)
       
   378 next
       
   379   case (Suc k x y)
       
   380   from Suc.prems have pxy: "p dvd x*y" by auto
       
   381   from prime_dvd_multD [OF p pxy] have pxyc: "p dvd x \<or> p dvd y" .
       
   382   from p have p0: "p \<noteq> 0" by - (rule ccontr, simp)
       
   383   {assume px: "p dvd x"
       
   384     then obtain d where d: "x = p*d" unfolding dvd_def by blast
       
   385     from Suc.prems d  have "p*d*y = p^Suc k" by simp
       
   386     hence th: "d*y = p^k" using p0 by simp
       
   387     from Suc.hyps[OF th] obtain i j where ij: "d = p^i" "y = p^j" by blast
       
   388     with d have "x = p^Suc i" by simp
       
   389     with ij(2) have ?case by blast}
       
   390   moreover
       
   391   {assume px: "p dvd y"
       
   392     then obtain d where d: "y = p*d" unfolding dvd_def by blast
       
   393     from Suc.prems d  have "p*d*x = p^Suc k" by (simp add: mult.commute)
       
   394     hence th: "d*x = p^k" using p0 by simp
       
   395     from Suc.hyps[OF th] obtain i j where ij: "d = p^i" "x = p^j" by blast
       
   396     with d have "y = p^Suc i" by simp
       
   397     with ij(2) have ?case by blast}
       
   398   ultimately show ?case  using pxyc by blast
       
   399 qed
       
   400 
       
   401 lemma prime_power_exp_nat:
       
   402   fixes p::nat
       
   403   assumes p: "prime p" and n: "n \<noteq> 0"
       
   404     and xn: "x^n = p^k" shows "\<exists>i. x = p^i"
       
   405   using n xn
       
   406 proof(induct n arbitrary: k)
       
   407   case 0 thus ?case by simp
       
   408 next
       
   409   case (Suc n k) hence th: "x*x^n = p^k" by simp
       
   410   {assume "n = 0" with Suc have ?case by simp (rule exI[where x="k"], simp)}
       
   411   moreover
       
   412   {assume n: "n \<noteq> 0"
       
   413     from prime_power_mult_nat[OF p th]
       
   414     obtain i j where ij: "x = p^i" "x^n = p^j"by blast
       
   415     from Suc.hyps[OF n ij(2)] have ?case .}
       
   416   ultimately show ?case by blast
       
   417 qed
       
   418 
       
   419 lemma divides_primepow_nat:
       
   420   fixes p::nat
       
   421   assumes p: "prime p"
       
   422   shows "d dvd p^k \<longleftrightarrow> (\<exists> i. i \<le> k \<and> d = p ^i)"
       
   423 proof
       
   424   assume H: "d dvd p^k" then obtain e where e: "d*e = p^k"
       
   425     unfolding dvd_def  apply (auto simp add: mult.commute) by blast
       
   426   from prime_power_mult_nat[OF p e] obtain i j where ij: "d = p^i" "e=p^j" by blast
       
   427   from e ij have "p^(i + j) = p^k" by (simp add: power_add)
       
   428   hence "i + j = k" using p prime_gt_1_nat power_inject_exp[of p "i+j" k] by simp
       
   429   hence "i \<le> k" by arith
       
   430   with ij(1) show "\<exists>i\<le>k. d = p ^ i" by blast
       
   431 next
       
   432   {fix i assume H: "i \<le> k" "d = p^i"
       
   433     then obtain j where j: "k = i + j"
       
   434       by (metis le_add_diff_inverse)
       
   435     hence "p^k = p^j*d" using H(2) by (simp add: power_add)
       
   436     hence "d dvd p^k" unfolding dvd_def by auto}
       
   437   thus "\<exists>i\<le>k. d = p ^ i \<Longrightarrow> d dvd p ^ k" by blast
       
   438 qed
       
   439 
       
   440 
       
   441 subsection \<open>Chinese Remainder Theorem Variants\<close>
       
   442 
       
   443 lemma bezout_gcd_nat:
       
   444   fixes a::nat shows "\<exists>x y. a * x - b * y = gcd a b \<or> b * x - a * y = gcd a b"
       
   445   using bezout_nat[of a b]
       
   446 by (metis bezout_nat diff_add_inverse gcd_add_mult gcd.commute
       
   447   gcd_nat.right_neutral mult_0)
       
   448 
       
   449 lemma gcd_bezout_sum_nat:
       
   450   fixes a::nat
       
   451   assumes "a * x + b * y = d"
       
   452   shows "gcd a b dvd d"
       
   453 proof-
       
   454   let ?g = "gcd a b"
       
   455     have dv: "?g dvd a*x" "?g dvd b * y"
       
   456       by simp_all
       
   457     from dvd_add[OF dv] assms
       
   458     show ?thesis by auto
       
   459 qed
       
   460 
       
   461 
       
   462 text \<open>A binary form of the Chinese Remainder Theorem.\<close>
       
   463 
       
   464 (* TODO: Generalise? *)
       
   465 lemma chinese_remainder:
       
   466   fixes a::nat  assumes ab: "coprime a b" and a: "a \<noteq> 0" and b: "b \<noteq> 0"
       
   467   shows "\<exists>x q1 q2. x = u + q1 * a \<and> x = v + q2 * b"
       
   468 proof-
       
   469   from bezout_add_strong_nat[OF a, of b] bezout_add_strong_nat[OF b, of a]
       
   470   obtain d1 x1 y1 d2 x2 y2 where dxy1: "d1 dvd a" "d1 dvd b" "a * x1 = b * y1 + d1"
       
   471     and dxy2: "d2 dvd b" "d2 dvd a" "b * x2 = a * y2 + d2" by blast
       
   472   then have d12: "d1 = 1" "d2 =1"
       
   473     by (metis ab coprime_nat)+
       
   474   let ?x = "v * a * x1 + u * b * x2"
       
   475   let ?q1 = "v * x1 + u * y2"
       
   476   let ?q2 = "v * y1 + u * x2"
       
   477   from dxy2(3)[simplified d12] dxy1(3)[simplified d12]
       
   478   have "?x = u + ?q1 * a" "?x = v + ?q2 * b"
       
   479     by algebra+
       
   480   thus ?thesis by blast
       
   481 qed
       
   482 
       
   483 text \<open>Primality\<close>
       
   484 
       
   485 lemma coprime_bezout_strong:
       
   486   fixes a::nat assumes "coprime a b"  "b \<noteq> 1"
       
   487   shows "\<exists>x y. a * x = b * y + 1"
       
   488 by (metis assms bezout_nat gcd_nat.left_neutral)
       
   489 
       
   490 lemma bezout_prime:
       
   491   assumes p: "prime p" and pa: "\<not> p dvd a"
       
   492   shows "\<exists>x y. a*x = Suc (p*y)"
       
   493 proof -
       
   494   have ap: "coprime a p"
       
   495     by (metis gcd.commute p pa prime_imp_coprime)
       
   496   moreover from p have "p \<noteq> 1" by auto
       
   497   ultimately have "\<exists>x y. a * x = p * y + 1"
       
   498     by (rule coprime_bezout_strong)
       
   499   then show ?thesis by simp    
       
   500 qed
       
   501 (* END TODO *)
       
   502 
       
   503 
       
   504 
       
   505 subsection \<open>Multiplicity and primality for natural numbers and integers\<close>
       
   506 
       
   507 lemma prime_factors_gt_0_nat:
       
   508   "p \<in> prime_factors x \<Longrightarrow> p > (0::nat)"
       
   509   by (simp add: in_prime_factors_imp_prime prime_gt_0_nat)
       
   510 
       
   511 lemma prime_factors_gt_0_int:
       
   512   "p \<in> prime_factors x \<Longrightarrow> p > (0::int)"
       
   513   by (simp add: in_prime_factors_imp_prime prime_gt_0_int)
       
   514 
       
   515 lemma prime_factors_ge_0_int [elim]: (* FIXME !? *)
       
   516   fixes n :: int
       
   517   shows "p \<in> prime_factors n \<Longrightarrow> p \<ge> 0"
       
   518   by (drule prime_factors_gt_0_int) simp
       
   519   
       
   520 lemma prod_mset_prime_factorization_int:
       
   521   fixes n :: int
       
   522   assumes "n > 0"
       
   523   shows   "prod_mset (prime_factorization n) = n"
       
   524   using assms by (simp add: prod_mset_prime_factorization)
       
   525 
       
   526 lemma prime_factorization_exists_nat:
       
   527   "n > 0 \<Longrightarrow> (\<exists>M. (\<forall>p::nat \<in> set_mset M. prime p) \<and> n = (\<Prod>i \<in># M. i))"
       
   528   using prime_factorization_exists[of n] by (auto simp: prime_def)
       
   529 
       
   530 lemma prod_mset_prime_factorization_nat [simp]: 
       
   531   "(n::nat) > 0 \<Longrightarrow> prod_mset (prime_factorization n) = n"
       
   532   by (subst prod_mset_prime_factorization) simp_all
       
   533 
       
   534 lemma prime_factorization_nat:
       
   535     "n > (0::nat) \<Longrightarrow> n = (\<Prod>p \<in> prime_factors n. p ^ multiplicity p n)"
       
   536   by (simp add: prod_prime_factors)
       
   537 
       
   538 lemma prime_factorization_int:
       
   539     "n > (0::int) \<Longrightarrow> n = (\<Prod>p \<in> prime_factors n. p ^ multiplicity p n)"
       
   540   by (simp add: prod_prime_factors)
       
   541 
       
   542 lemma prime_factorization_unique_nat:
       
   543   fixes f :: "nat \<Rightarrow> _"
       
   544   assumes S_eq: "S = {p. 0 < f p}"
       
   545     and "finite S"
       
   546     and S: "\<forall>p\<in>S. prime p" "n = (\<Prod>p\<in>S. p ^ f p)"
       
   547   shows "S = prime_factors n \<and> (\<forall>p. prime p \<longrightarrow> f p = multiplicity p n)"
       
   548   using assms by (intro prime_factorization_unique'') auto
       
   549 
       
   550 lemma prime_factorization_unique_int:
       
   551   fixes f :: "int \<Rightarrow> _"
       
   552   assumes S_eq: "S = {p. 0 < f p}"
       
   553     and "finite S"
       
   554     and S: "\<forall>p\<in>S. prime p" "abs n = (\<Prod>p\<in>S. p ^ f p)"
       
   555   shows "S = prime_factors n \<and> (\<forall>p. prime p \<longrightarrow> f p = multiplicity p n)"
       
   556   using assms by (intro prime_factorization_unique'') auto
       
   557 
       
   558 lemma prime_factors_characterization_nat:
       
   559   "S = {p. 0 < f (p::nat)} \<Longrightarrow>
       
   560     finite S \<Longrightarrow> \<forall>p\<in>S. prime p \<Longrightarrow> n = (\<Prod>p\<in>S. p ^ f p) \<Longrightarrow> prime_factors n = S"
       
   561   by (rule prime_factorization_unique_nat [THEN conjunct1, symmetric])
       
   562 
       
   563 lemma prime_factors_characterization'_nat:
       
   564   "finite {p. 0 < f (p::nat)} \<Longrightarrow>
       
   565     (\<forall>p. 0 < f p \<longrightarrow> prime p) \<Longrightarrow>
       
   566       prime_factors (\<Prod>p | 0 < f p. p ^ f p) = {p. 0 < f p}"
       
   567   by (rule prime_factors_characterization_nat) auto
       
   568 
       
   569 lemma prime_factors_characterization_int:
       
   570   "S = {p. 0 < f (p::int)} \<Longrightarrow> finite S \<Longrightarrow>
       
   571     \<forall>p\<in>S. prime p \<Longrightarrow> abs n = (\<Prod>p\<in>S. p ^ f p) \<Longrightarrow> prime_factors n = S"
       
   572   by (rule prime_factorization_unique_int [THEN conjunct1, symmetric])
       
   573 
       
   574 (* TODO Move *)
       
   575 lemma abs_prod: "abs (prod f A :: 'a :: linordered_idom) = prod (\<lambda>x. abs (f x)) A"
       
   576   by (cases "finite A", induction A rule: finite_induct) (simp_all add: abs_mult)
       
   577 
       
   578 lemma primes_characterization'_int [rule_format]:
       
   579   "finite {p. p \<ge> 0 \<and> 0 < f (p::int)} \<Longrightarrow> \<forall>p. 0 < f p \<longrightarrow> prime p \<Longrightarrow>
       
   580       prime_factors (\<Prod>p | p \<ge> 0 \<and> 0 < f p. p ^ f p) = {p. p \<ge> 0 \<and> 0 < f p}"
       
   581   by (rule prime_factors_characterization_int) (auto simp: abs_prod prime_ge_0_int)
       
   582 
       
   583 lemma multiplicity_characterization_nat:
       
   584   "S = {p. 0 < f (p::nat)} \<Longrightarrow> finite S \<Longrightarrow> \<forall>p\<in>S. prime p \<Longrightarrow> prime p \<Longrightarrow>
       
   585     n = (\<Prod>p\<in>S. p ^ f p) \<Longrightarrow> multiplicity p n = f p"
       
   586   by (frule prime_factorization_unique_nat [of S f n, THEN conjunct2, rule_format, symmetric]) auto
       
   587 
       
   588 lemma multiplicity_characterization'_nat: "finite {p. 0 < f (p::nat)} \<longrightarrow>
       
   589     (\<forall>p. 0 < f p \<longrightarrow> prime p) \<longrightarrow> prime p \<longrightarrow>
       
   590       multiplicity p (\<Prod>p | 0 < f p. p ^ f p) = f p"
       
   591   by (intro impI, rule multiplicity_characterization_nat) auto
       
   592 
       
   593 lemma multiplicity_characterization_int: "S = {p. 0 < f (p::int)} \<Longrightarrow>
       
   594     finite S \<Longrightarrow> \<forall>p\<in>S. prime p \<Longrightarrow> prime p \<Longrightarrow> n = (\<Prod>p\<in>S. p ^ f p) \<Longrightarrow> multiplicity p n = f p"
       
   595   by (frule prime_factorization_unique_int [of S f n, THEN conjunct2, rule_format, symmetric]) 
       
   596      (auto simp: abs_prod power_abs prime_ge_0_int intro!: prod.cong)
       
   597 
       
   598 lemma multiplicity_characterization'_int [rule_format]:
       
   599   "finite {p. p \<ge> 0 \<and> 0 < f (p::int)} \<Longrightarrow>
       
   600     (\<forall>p. 0 < f p \<longrightarrow> prime p) \<Longrightarrow> prime p \<Longrightarrow>
       
   601       multiplicity p (\<Prod>p | p \<ge> 0 \<and> 0 < f p. p ^ f p) = f p"
       
   602   by (rule multiplicity_characterization_int) (auto simp: prime_ge_0_int)
       
   603 
       
   604 lemma multiplicity_one_nat [simp]: "multiplicity p (Suc 0) = 0"
       
   605   unfolding One_nat_def [symmetric] by (rule multiplicity_one)
       
   606 
       
   607 lemma multiplicity_eq_nat:
       
   608   fixes x and y::nat
       
   609   assumes "x > 0" "y > 0" "\<And>p. prime p \<Longrightarrow> multiplicity p x = multiplicity p y"
       
   610   shows "x = y"
       
   611   using multiplicity_eq_imp_eq[of x y] assms by simp
       
   612 
       
   613 lemma multiplicity_eq_int:
       
   614   fixes x y :: int
       
   615   assumes "x > 0" "y > 0" "\<And>p. prime p \<Longrightarrow> multiplicity p x = multiplicity p y"
       
   616   shows "x = y"
       
   617   using multiplicity_eq_imp_eq[of x y] assms by simp
       
   618 
       
   619 lemma multiplicity_prod_prime_powers:
       
   620   assumes "finite S" "\<And>x. x \<in> S \<Longrightarrow> prime x" "prime p"
       
   621   shows   "multiplicity p (\<Prod>p \<in> S. p ^ f p) = (if p \<in> S then f p else 0)"
       
   622 proof -
       
   623   define g where "g = (\<lambda>x. if x \<in> S then f x else 0)"
       
   624   define A where "A = Abs_multiset g"
       
   625   have "{x. g x > 0} \<subseteq> S" by (auto simp: g_def)
       
   626   from finite_subset[OF this assms(1)] have [simp]: "g :  multiset"
       
   627     by (simp add: multiset_def)
       
   628   from assms have count_A: "count A x = g x" for x unfolding A_def
       
   629     by simp
       
   630   have set_mset_A: "set_mset A = {x\<in>S. f x > 0}"
       
   631     unfolding set_mset_def count_A by (auto simp: g_def)
       
   632   with assms have prime: "prime x" if "x \<in># A" for x using that by auto
       
   633   from set_mset_A assms have "(\<Prod>p \<in> S. p ^ f p) = (\<Prod>p \<in> S. p ^ g p) "
       
   634     by (intro prod.cong) (auto simp: g_def)
       
   635   also from set_mset_A assms have "\<dots> = (\<Prod>p \<in> set_mset A. p ^ g p)"
       
   636     by (intro prod.mono_neutral_right) (auto simp: g_def set_mset_A)
       
   637   also have "\<dots> = prod_mset A"
       
   638     by (auto simp: prod_mset_multiplicity count_A set_mset_A intro!: prod.cong)
       
   639   also from assms have "multiplicity p \<dots> = sum_mset (image_mset (multiplicity p) A)"
       
   640     by (subst prime_elem_multiplicity_prod_mset_distrib) (auto dest: prime)
       
   641   also from assms have "image_mset (multiplicity p) A = image_mset (\<lambda>x. if x = p then 1 else 0) A"
       
   642     by (intro image_mset_cong) (auto simp: prime_multiplicity_other dest: prime)
       
   643   also have "sum_mset \<dots> = (if p \<in> S then f p else 0)" by (simp add: sum_mset_delta count_A g_def)
       
   644   finally show ?thesis .
       
   645 qed
       
   646 
       
   647 lemma prime_factorization_prod_mset:
       
   648   assumes "0 \<notin># A"
       
   649   shows "prime_factorization (prod_mset A) = \<Union>#(image_mset prime_factorization A)"
       
   650   using assms by (induct A) (auto simp add: prime_factorization_mult)
       
   651 
       
   652 lemma prime_factors_prod:
       
   653   assumes "finite A" and "0 \<notin> f ` A"
       
   654   shows "prime_factors (prod f A) = UNION A (prime_factors \<circ> f)"
       
   655   using assms by (simp add: prod_unfold_prod_mset prime_factorization_prod_mset)
       
   656 
       
   657 lemma prime_factors_fact:
       
   658   "prime_factors (fact n) = {p \<in> {2..n}. prime p}" (is "?M = ?N")
       
   659 proof (rule set_eqI)
       
   660   fix p
       
   661   { fix m :: nat
       
   662     assume "p \<in> prime_factors m"
       
   663     then have "prime p" and "p dvd m" by auto
       
   664     moreover assume "m > 0" 
       
   665     ultimately have "2 \<le> p" and "p \<le> m"
       
   666       by (auto intro: prime_ge_2_nat dest: dvd_imp_le)
       
   667     moreover assume "m \<le> n"
       
   668     ultimately have "2 \<le> p" and "p \<le> n"
       
   669       by (auto intro: order_trans)
       
   670   } note * = this
       
   671   show "p \<in> ?M \<longleftrightarrow> p \<in> ?N"
       
   672     by (auto simp add: fact_prod prime_factors_prod Suc_le_eq dest!: prime_prime_factors intro: *)
       
   673 qed
       
   674 
       
   675 lemma prime_dvd_fact_iff:
       
   676   assumes "prime p"
       
   677   shows "p dvd fact n \<longleftrightarrow> p \<le> n"
       
   678   using assms
       
   679   by (auto simp add: prime_factorization_subset_iff_dvd [symmetric]
       
   680     prime_factorization_prime prime_factors_fact prime_ge_2_nat)
       
   681 
       
   682 (* TODO Legacy names *)
       
   683 lemmas prime_imp_coprime_nat = prime_imp_coprime[where ?'a = nat]
       
   684 lemmas prime_imp_coprime_int = prime_imp_coprime[where ?'a = int]
       
   685 lemmas prime_dvd_mult_nat = prime_dvd_mult_iff[where ?'a = nat]
       
   686 lemmas prime_dvd_mult_int = prime_dvd_mult_iff[where ?'a = int]
       
   687 lemmas prime_dvd_mult_eq_nat = prime_dvd_mult_iff[where ?'a = nat]
       
   688 lemmas prime_dvd_mult_eq_int = prime_dvd_mult_iff[where ?'a = int]
       
   689 lemmas prime_dvd_power_nat = prime_dvd_power[where ?'a = nat]
       
   690 lemmas prime_dvd_power_int = prime_dvd_power[where ?'a = int]
       
   691 lemmas prime_dvd_power_nat_iff = prime_dvd_power_iff[where ?'a = nat]
       
   692 lemmas prime_dvd_power_int_iff = prime_dvd_power_iff[where ?'a = int]
       
   693 lemmas prime_imp_power_coprime_nat = prime_imp_power_coprime[where ?'a = nat]
       
   694 lemmas prime_imp_power_coprime_int = prime_imp_power_coprime[where ?'a = int]
       
   695 lemmas primes_coprime_nat = primes_coprime[where ?'a = nat]
       
   696 lemmas primes_coprime_int = primes_coprime[where ?'a = nat]
       
   697 lemmas prime_divprod_pow_nat = prime_elem_divprod_pow[where ?'a = nat]
       
   698 lemmas prime_exp = prime_elem_power_iff[where ?'a = nat]
       
   699 
       
   700 text \<open>Code generation\<close>
       
   701   
       
   702 context
       
   703 begin
       
   704 
       
   705 qualified definition prime_nat :: "nat \<Rightarrow> bool"
       
   706   where [simp, code_abbrev]: "prime_nat = prime"
       
   707 
       
   708 lemma prime_nat_naive [code]:
       
   709   "prime_nat p \<longleftrightarrow> p > 1 \<and> (\<forall>n \<in>{1<..<p}. \<not> n dvd p)"
       
   710   by (auto simp add: prime_nat_iff')
       
   711 
       
   712 qualified definition prime_int :: "int \<Rightarrow> bool"
       
   713   where [simp, code_abbrev]: "prime_int = prime"
       
   714 
       
   715 lemma prime_int_naive [code]:
       
   716   "prime_int p \<longleftrightarrow> p > 1 \<and> (\<forall>n \<in>{1<..<p}. \<not> n dvd p)"
       
   717   by (auto simp add: prime_int_iff')
       
   718 
       
   719 lemma "prime(997::nat)" by eval
       
   720 
       
   721 lemma "prime(997::int)" by eval
       
   722   
       
   723 end
       
   724 
       
   725 end