src/HOL/Library/Primes.thy
changeset 26125 345465cc9e79
parent 25593 0b0df6c8646a
child 26144 98d23fc02585
equal deleted inserted replaced
26124:2514f0ade8bc 26125:345465cc9e79
     1 (*  Title:      HOL/Library/Primes.thy
     1 (*  Title:      HOL/Library/Primes.thy
     2     ID:         $Id$
     2     ID:         $Id$
     3     Author:     Christophe Tabacznyj and Lawrence C Paulson
     3     Author:     Amine Chaieb Christophe Tabacznyj and Lawrence C Paulson
     4     Copyright   1996  University of Cambridge
     4     Copyright   1996  University of Cambridge
     5 *)
     5 *)
     6 
     6 
     7 header {* Primality on nat *}
     7 header {* Primality on nat *}
     8 
     8 
     9 theory Primes
     9 theory Primes
    10 imports GCD ATP_Linkup
    10 imports GCD Parity
    11 begin
    11 begin
    12 
    12 
    13 definition
    13 definition
    14   coprime :: "nat => nat => bool" where
    14   coprime :: "nat => nat => bool" where
    15   "coprime m n = (gcd (m, n) = 1)"
    15   "coprime m n = (gcd (m, n) = 1)"
    43   by (auto dest: prime_dvd_mult)
    43   by (auto dest: prime_dvd_mult)
    44 
    44 
    45 lemma prime_dvd_power_two: "prime p ==> p dvd m\<twosuperior> ==> p dvd m"
    45 lemma prime_dvd_power_two: "prime p ==> p dvd m\<twosuperior> ==> p dvd m"
    46   by (rule prime_dvd_square) (simp_all add: power2_eq_square)
    46   by (rule prime_dvd_square) (simp_all add: power2_eq_square)
    47 
    47 
       
    48 
       
    49 lemma exp_eq_1:"(x::nat)^n = 1 \<longleftrightarrow> x = 1 \<or> n = 0" by (induct n, auto)
       
    50 lemma exp_mono_lt: "(x::nat) ^ (Suc n) < y ^ (Suc n) \<longleftrightarrow> x < y"
       
    51   using power_less_imp_less_base[of x "Suc n" y] power_strict_mono[of x y "Suc n"]
       
    52     by auto
       
    53 lemma exp_mono_le: "(x::nat) ^ (Suc n) \<le> y ^ (Suc n) \<longleftrightarrow> x \<le> y"
       
    54   by (simp only: linorder_not_less[symmetric] exp_mono_lt)
       
    55 
       
    56 lemma exp_mono_eq: "(x::nat) ^ Suc n = y ^ Suc n \<longleftrightarrow> x = y"
       
    57 using power_inject_base[of x n y] by auto
       
    58 
       
    59 
       
    60 lemma even_square: assumes e: "even (n::nat)" shows "\<exists>x. n ^ 2 = 4*x"
       
    61 proof-
       
    62   from e have "2 dvd n" by presburger
       
    63   then obtain k where k: "n = 2*k" using dvd_def by auto
       
    64   hence "n^2 = 4* (k^2)" by (simp add: power2_eq_square)
       
    65   thus ?thesis by blast
       
    66 qed
       
    67 
       
    68 lemma odd_square: assumes e: "odd (n::nat)" shows "\<exists>x. n ^ 2 = 4*x + 1"
       
    69 proof-
       
    70   from e have np: "n > 0" by presburger
       
    71   from e have "2 dvd (n - 1)" by presburger
       
    72   then obtain k where "n - 1 = 2*k" using dvd_def by auto
       
    73   hence k: "n = 2*k + 1"  using e by presburger 
       
    74   hence "n^2 = 4* (k^2 + k) + 1" by algebra   
       
    75   thus ?thesis by blast
       
    76 qed
       
    77 
       
    78 lemma diff_square: "(x::nat)^2 - y^2 = (x+y)*(x - y)" 
       
    79 proof-
       
    80   have "x \<le> y \<or> y \<le> x" by (rule nat_le_linear)
       
    81   moreover
       
    82   {assume le: "x \<le> y"
       
    83     hence "x ^2 \<le> y^2" by (simp only: numeral_2_eq_2 exp_mono_le Let_def)
       
    84     with le have ?thesis by simp }
       
    85   moreover
       
    86   {assume le: "y \<le> x"
       
    87     hence le2: "y ^2 \<le> x^2" by (simp only: numeral_2_eq_2 exp_mono_le Let_def)
       
    88     from le have "\<exists>z. y + z = x" by presburger
       
    89     then obtain z where z: "x = y + z" by blast 
       
    90     from le2 have "\<exists>z. x^2 = y^2 + z" by presburger
       
    91     then obtain z2 where z2: "x^2 = y^2 + z2"  by blast
       
    92     from z z2 have ?thesis apply simp by algebra }
       
    93   ultimately show ?thesis by blast  
       
    94 qed
       
    95 
       
    96 (* Elementary theory of divisibility                                         *)
       
    97 lemma divides_ge: "(a::nat) dvd b \<Longrightarrow> b = 0 \<or> a \<le> b" unfolding dvd_def by auto
       
    98 lemma divides_antisym: "(x::nat) dvd y \<and> y dvd x \<longleftrightarrow> x = y"
       
    99   using dvd_anti_sym[of x y] by auto
       
   100 
       
   101 lemma divides_add_revr: assumes da: "(d::nat) dvd a" and dab:"d dvd (a + b)"
       
   102   shows "d dvd b"
       
   103 proof-
       
   104   from da obtain k where k:"a = d*k" by (auto simp add: dvd_def)
       
   105   from dab obtain k' where k': "a + b = d*k'" by (auto simp add: dvd_def)
       
   106   from k k' have "b = d *(k' - k)" by (simp add : diff_mult_distrib2)
       
   107   thus ?thesis unfolding dvd_def by blast
       
   108 qed
       
   109 
       
   110 declare nat_mult_dvd_cancel_disj[presburger]
       
   111 lemma nat_mult_dvd_cancel_disj'[presburger]: 
       
   112   "(m\<Colon>nat)*k dvd n*k \<longleftrightarrow> k = 0 \<or> m dvd n" unfolding mult_commute[of m k] mult_commute[of n k] by presburger 
       
   113 
       
   114 lemma divides_mul_l: "(a::nat) dvd b ==> (c * a) dvd (c * b)"
       
   115   by presburger
       
   116 
       
   117 lemma divides_mul_r: "(a::nat) dvd b ==> (a * c) dvd (b * c)" by presburger
       
   118 lemma divides_cases: "(n::nat) dvd m ==> m = 0 \<or> m = n \<or> 2 * n <= m" 
       
   119   by (auto simp add: dvd_def)
       
   120 lemma divides_le: "m dvd n ==> m <= n \<or> n = (0::nat)" by (auto simp add: dvd_def)
       
   121 
       
   122 lemma divides_div_not: "(x::nat) = (q * n) + r \<Longrightarrow> 0 < r \<Longrightarrow> r < n ==> ~(n dvd x)"
       
   123 proof(auto simp add: dvd_def)
       
   124   fix k assume H: "0 < r" "r < n" "q * n + r = n * k"
       
   125   from H(3) have r: "r = n* (k -q)" by(simp add: diff_mult_distrib2 mult_commute)
       
   126   {assume "k - q = 0" with r H(1) have False by simp}
       
   127   moreover
       
   128   {assume "k - q \<noteq> 0" with r have "r \<ge> n" by auto
       
   129     with H(2) have False by simp}
       
   130   ultimately show False by blast
       
   131 qed
       
   132 lemma divides_exp: "(x::nat) dvd y ==> x ^ n dvd y ^ n"
       
   133   by (auto simp add: power_mult_distrib dvd_def)
       
   134 
       
   135 lemma divides_exp2: "n \<noteq> 0 \<Longrightarrow> (x::nat) ^ n dvd y \<Longrightarrow> x dvd y" 
       
   136   by (induct n ,auto simp add: dvd_def)
       
   137 
       
   138 fun fact :: "nat \<Rightarrow> nat" where
       
   139   "fact 0 = 1"
       
   140 | "fact (Suc n) = Suc n * fact n"	
       
   141 
       
   142 lemma fact_lt: "0 < fact n" by(induct n, simp_all)
       
   143 lemma fact_le: "fact n \<ge> 1" using fact_lt[of n] by simp 
       
   144 lemma fact_mono: assumes le: "m \<le> n" shows "fact m \<le> fact n"
       
   145 proof-
       
   146   from le have "\<exists>i. n = m+i" by presburger
       
   147   then obtain i where i: "n = m+i" by blast 
       
   148   have "fact m \<le> fact (m + i)"
       
   149   proof(induct m)
       
   150     case 0 thus ?case using fact_le[of i] by simp
       
   151   next
       
   152     case (Suc m)
       
   153     have "fact (Suc m) = Suc m * fact m" by simp
       
   154     have th1: "Suc m \<le> Suc (m + i)" by simp
       
   155     from mult_le_mono[of "Suc m" "Suc (m+i)" "fact m" "fact (m+i)", OF th1 Suc.hyps]
       
   156     show ?case by simp
       
   157   qed
       
   158   thus ?thesis using i by simp
       
   159 qed
       
   160 
       
   161 lemma divides_fact: "1 <= p \<Longrightarrow> p <= n ==> p dvd fact n"
       
   162 proof(induct n arbitrary: p)
       
   163   case 0 thus ?case by simp
       
   164 next
       
   165   case (Suc n p)
       
   166   from Suc.prems have "p = Suc n \<or> p \<le> n" by presburger 
       
   167   moreover
       
   168   {assume "p = Suc n" hence ?case  by (simp only: fact.simps dvd_triv_left)}
       
   169   moreover
       
   170   {assume "p \<le> n"
       
   171     with Suc.prems(1) Suc.hyps have th: "p dvd fact n" by simp
       
   172     from dvd_mult[OF th] have ?case by (simp only: fact.simps) }
       
   173   ultimately show ?case by blast
       
   174 qed
       
   175 
       
   176 declare dvd_triv_left[presburger]
       
   177 declare dvd_triv_right[presburger]
       
   178 lemma divides_rexp: 
       
   179   "x dvd y \<Longrightarrow> (x::nat) dvd (y^(Suc n))" by (simp add: dvd_mult2[of x y])
       
   180 
       
   181 (* The Bezout theorem is a bit ugly for N; it'd be easier for Z              *)
       
   182 lemma ind_euclid: 
       
   183   assumes c: " \<forall>a b. P (a::nat) b \<longleftrightarrow> P b a" and z: "\<forall>a. P a 0" 
       
   184   and add: "\<forall>a b. P a b \<longrightarrow> P a (a + b)" 
       
   185   shows "P a b"
       
   186 proof(induct n\<equiv>"a+b" arbitrary: a b rule: nat_less_induct)
       
   187   fix n a b
       
   188   assume H: "\<forall>m < n. \<forall>a b. m = a + b \<longrightarrow> P a b" "n = a + b"
       
   189   have "a = b \<or> a < b \<or> b < a" by arith
       
   190   moreover {assume eq: "a= b"
       
   191     from add[rule_format, OF z[rule_format, of a]] have "P a b" using eq by simp}
       
   192   moreover
       
   193   {assume lt: "a < b"
       
   194     hence "a + b - a < n \<or> a = 0"  using H(2) by arith
       
   195     moreover
       
   196     {assume "a =0" with z c have "P a b" by blast }
       
   197     moreover
       
   198     {assume ab: "a + b - a < n"
       
   199       have th0: "a + b - a = a + (b - a)" using lt by arith
       
   200       from add[rule_format, OF H(1)[rule_format, OF ab th0]]
       
   201       have "P a b" by (simp add: th0[symmetric])}
       
   202     ultimately have "P a b" by blast}
       
   203   moreover
       
   204   {assume lt: "a > b"
       
   205     hence "b + a - b < n \<or> b = 0"  using H(2) by arith
       
   206     moreover
       
   207     {assume "b =0" with z c have "P a b" by blast }
       
   208     moreover
       
   209     {assume ab: "b + a - b < n"
       
   210       have th0: "b + a - b = b + (a - b)" using lt by arith
       
   211       from add[rule_format, OF H(1)[rule_format, OF ab th0]]
       
   212       have "P b a" by (simp add: th0[symmetric])
       
   213       hence "P a b" using c by blast }
       
   214     ultimately have "P a b" by blast}
       
   215 ultimately  show "P a b" by blast
       
   216 qed
       
   217 
       
   218 lemma bezout_lemma: 
       
   219   assumes ex: "\<exists>(d::nat) x y. d dvd a \<and> d dvd b \<and> (a * x = b * y + d \<or> b * x = a * y + d)"
       
   220   shows "\<exists>d x y. d dvd a \<and> d dvd a + b \<and> (a * x = (a + b) * y + d \<or> (a + b) * x = a * y + d)"
       
   221 using ex
       
   222 apply clarsimp
       
   223 apply (rule_tac x="d" in exI, simp add: dvd_add)
       
   224 apply (case_tac "a * x = b * y + d" , simp_all)
       
   225 apply (rule_tac x="x + y" in exI)
       
   226 apply (rule_tac x="y" in exI)
       
   227 apply algebra
       
   228 apply (rule_tac x="x" in exI)
       
   229 apply (rule_tac x="x + y" in exI)
       
   230 apply algebra
       
   231 done
       
   232 
       
   233 lemma bezout_add: "\<exists>(d::nat) x y. d dvd a \<and> d dvd b \<and> (a * x = b * y + d \<or> b * x = a * y + d)"
       
   234 apply(induct a b rule: ind_euclid)
       
   235 apply blast
       
   236 apply clarify
       
   237 apply (rule_tac x="a" in exI, simp add: dvd_add)
       
   238 apply clarsimp
       
   239 apply (rule_tac x="d" in exI)
       
   240 apply (case_tac "a * x = b * y + d", simp_all add: dvd_add)
       
   241 apply (rule_tac x="x+y" in exI)
       
   242 apply (rule_tac x="y" in exI)
       
   243 apply algebra
       
   244 apply (rule_tac x="x" in exI)
       
   245 apply (rule_tac x="x+y" in exI)
       
   246 apply algebra
       
   247 done
       
   248 
       
   249 lemma bezout: "\<exists>(d::nat) x y. d dvd a \<and> d dvd b \<and> (a * x - b * y = d \<or> b * x - a * y = d)"
       
   250 using bezout_add[of a b]
       
   251 apply clarsimp
       
   252 apply (rule_tac x="d" in exI, simp)
       
   253 apply (rule_tac x="x" in exI)
       
   254 apply (rule_tac x="y" in exI)
       
   255 apply auto
       
   256 done
       
   257 
       
   258 (* We can get a stronger version with a nonzeroness assumption.              *)
       
   259 
       
   260 lemma bezout_add_strong: assumes nz: "a \<noteq> (0::nat)"
       
   261   shows "\<exists>d x y. d dvd a \<and> d dvd b \<and> a * x = b * y + d"
       
   262 proof-
       
   263   from nz have ap: "a > 0" by simp
       
   264  from bezout_add[of a b] 
       
   265  have "(\<exists>d x y. d dvd a \<and> d dvd b \<and> a * x = b * y + d) \<or> (\<exists>d x y. d dvd a \<and> d dvd b \<and> b * x = a * y + d)" by blast
       
   266  moreover
       
   267  {fix d x y assume H: "d dvd a" "d dvd b" "a * x = b * y + d"
       
   268    from H have ?thesis by blast }
       
   269  moreover
       
   270  {fix d x y assume H: "d dvd a" "d dvd b" "b * x = a * y + d"
       
   271    {assume b0: "b = 0" with H  have ?thesis by simp}
       
   272    moreover 
       
   273    {assume b: "b \<noteq> 0" hence bp: "b > 0" by simp
       
   274      from divides_le[OF H(2)] b have "d < b \<or> d = b" using le_less by blast
       
   275      moreover
       
   276      {assume db: "d=b"
       
   277        from prems have ?thesis apply simp
       
   278 	 apply (rule exI[where x = b], simp)
       
   279 	 apply (rule exI[where x = b])
       
   280 	by (rule exI[where x = "a - 1"], simp add: diff_mult_distrib2)}
       
   281     moreover
       
   282     {assume db: "d < b" 
       
   283 	{assume "x=0" hence ?thesis  using prems by simp }
       
   284 	moreover
       
   285 	{assume x0: "x \<noteq> 0" hence xp: "x > 0" by simp
       
   286 	  
       
   287 	  from db have "d \<le> b - 1" by simp
       
   288 	  hence "d*b \<le> b*(b - 1)" by simp
       
   289 	  with xp mult_mono[of "1" "x" "d*b" "b*(b - 1)"]
       
   290 	  have dble: "d*b \<le> x*b*(b - 1)" using bp by simp
       
   291 	  from H (3) have "d + (b - 1) * (b*x) = d + (b - 1) * (a*y + d)" by simp
       
   292 	  hence "d + (b - 1) * a * y + (b - 1) * d = d + (b - 1) * b * x"
       
   293 	    by (simp only: mult_assoc right_distrib)
       
   294 	  hence "a * ((b - 1) * y) + d * (b - 1 + 1) = d + x*b*(b - 1)" by algebra
       
   295 	  hence "a * ((b - 1) * y) = d + x*b*(b - 1) - d*b" using bp by simp
       
   296 	  hence "a * ((b - 1) * y) = d + (x*b*(b - 1) - d*b)"
       
   297 	    by (simp only: diff_add_assoc[OF dble, of d, symmetric])
       
   298 	  hence "a * ((b - 1) * y) = b*(x*(b - 1) - d) + d"
       
   299 	    by (simp only: diff_mult_distrib2 add_commute mult_ac)
       
   300 	  hence ?thesis using H(1,2)
       
   301 	    apply -
       
   302 	    apply (rule exI[where x=d], simp)
       
   303 	    apply (rule exI[where x="(b - 1) * y"])
       
   304 	    by (rule exI[where x="x*(b - 1) - d"], simp)}
       
   305 	ultimately have ?thesis by blast}
       
   306     ultimately have ?thesis by blast}
       
   307   ultimately have ?thesis by blast}
       
   308  ultimately show ?thesis by blast
       
   309 qed
       
   310 
       
   311 (* Greatest common divisor.                                                  *)
       
   312 lemma gcd_unique: "d dvd a\<and>d dvd b \<and> (\<forall>e. e dvd a \<and> e dvd b \<longrightarrow> e dvd d) \<longleftrightarrow> d = gcd(a,b)"
       
   313 proof(auto)
       
   314   assume H: "d dvd a" "d dvd b" "\<forall>e. e dvd a \<and> e dvd b \<longrightarrow> e dvd d"
       
   315   from H(3)[rule_format] gcd_dvd1[of a b] gcd_dvd2[of a b] 
       
   316   have th: "gcd (a,b) dvd d" by blast
       
   317   from dvd_anti_sym[OF th gcd_greatest[OF H(1,2)]]  show "d = gcd (a,b)" by blast 
       
   318 qed
       
   319 
       
   320 lemma gcd_eq: assumes H: "\<forall>d. d dvd x \<and> d dvd y \<longleftrightarrow> d dvd u \<and> d dvd v"
       
   321   shows "gcd (x,y) = gcd(u,v)"
       
   322 proof-
       
   323   from H have "\<forall>d. d dvd x \<and> d dvd y \<longleftrightarrow> d dvd gcd (u,v)" by simp
       
   324   with gcd_unique[of "gcd(u,v)" x y]  show ?thesis by auto
       
   325 qed
       
   326 
       
   327 lemma bezout_gcd: "\<exists>x y. a * x - b * y = gcd(a,b) \<or> b * x - a * y = gcd(a,b)"
       
   328 proof-
       
   329   let ?g = "gcd (a,b)"
       
   330   from bezout[of a b] obtain d x y where d: "d dvd a" "d dvd b" "a * x - b * y = d \<or> b * x - a * y = d" by blast
       
   331   from d(1,2) have "d dvd ?g" by simp
       
   332   then obtain k where k: "?g = d*k" unfolding dvd_def by blast
       
   333   from d(3) have "(a * x - b * y)*k = d*k \<or> (b * x - a * y)*k = d*k" by blast 
       
   334   hence "a * x * k - b * y*k = d*k \<or> b * x * k - a * y*k = d*k"
       
   335     by (simp only: diff_mult_distrib)
       
   336   hence "a * (x * k) - b * (y*k) = ?g \<or> b * (x * k) - a * (y*k) = ?g"
       
   337     by (simp add: k mult_assoc)
       
   338   thus ?thesis by blast
       
   339 qed
       
   340 
       
   341 lemma bezout_gcd_strong: assumes a: "a \<noteq> 0" 
       
   342   shows "\<exists>x y. a * x = b * y + gcd(a,b)"
       
   343 proof-
       
   344   let ?g = "gcd (a,b)"
       
   345   from bezout_add_strong[OF a, of b]
       
   346   obtain d x y where d: "d dvd a" "d dvd b" "a * x = b * y + d" by blast
       
   347   from d(1,2) have "d dvd ?g" by simp
       
   348   then obtain k where k: "?g = d*k" unfolding dvd_def by blast
       
   349   from d(3) have "a * x * k = (b * y + d) *k " by auto 
       
   350   hence "a * (x * k) = b * (y*k) + ?g" by (algebra add: k)
       
   351   thus ?thesis by blast
       
   352 qed
       
   353 
       
   354 lemma gcd_mult_distrib: "gcd(a * c, b * c) = c * gcd(a,b)"
       
   355 by(simp add: gcd_mult_distrib2 mult_commute)
       
   356 
       
   357 lemma gcd_bezout: "(\<exists>x y. a * x - b * y = d \<or> b * x - a * y = d) \<longleftrightarrow> gcd(a,b) dvd d"
       
   358   (is "?lhs \<longleftrightarrow> ?rhs")
       
   359 proof-
       
   360   let ?g = "gcd (a,b)"
       
   361   {assume H: ?rhs then obtain k where k: "d = ?g*k" unfolding dvd_def by blast
       
   362     from bezout_gcd[of a b] obtain x y where xy: "a * x - b * y = ?g \<or> b * x - a * y = ?g"
       
   363       by blast
       
   364     hence "(a * x - b * y)*k = ?g*k \<or> (b * x - a * y)*k = ?g*k" by auto
       
   365     hence "a * x*k - b * y*k = ?g*k \<or> b * x * k - a * y*k = ?g*k" 
       
   366       by (simp only: diff_mult_distrib)
       
   367     hence "a * (x*k) - b * (y*k) = d \<or> b * (x * k) - a * (y*k) = d" 
       
   368       by (simp add: k[symmetric] mult_assoc)
       
   369     hence ?lhs by blast}
       
   370   moreover
       
   371   {fix x y assume H: "a * x - b * y = d \<or> b * x - a * y = d"
       
   372     have dv: "?g dvd a*x" "?g dvd b * y" "?g dvd b*x" "?g dvd a * y" 
       
   373       using dvd_mult2[OF gcd_dvd1[of a b]] dvd_mult2[OF gcd_dvd2[of a b]] by simp_all
       
   374     from dvd_diff[OF dv(1,2)] dvd_diff[OF dv(3,4)] H
       
   375     have ?rhs by auto}
       
   376   ultimately show ?thesis by blast
       
   377 qed
       
   378 
       
   379 lemma gcd_bezout_sum: assumes H:"a * x + b * y = d" shows "gcd(a,b) dvd d"
       
   380 proof-
       
   381   let ?g = "gcd (a,b)"
       
   382     have dv: "?g dvd a*x" "?g dvd b * y" 
       
   383       using dvd_mult2[OF gcd_dvd1[of a b]] dvd_mult2[OF gcd_dvd2[of a b]] by simp_all
       
   384     from dvd_add[OF dv] H
       
   385     show ?thesis by auto
       
   386 qed
       
   387 
       
   388 lemma gcd_mult': "gcd(b,a * b) = b"
       
   389 by (simp add: gcd_mult mult_commute[of a b]) 
       
   390 
       
   391 lemma gcd_add: "gcd(a + b,b) = gcd(a,b)" "gcd(b + a,b) = gcd(a,b)" "gcd(a,a + b) = gcd(a,b)" "gcd(a,b + a) = gcd(a,b)"
       
   392 apply (simp_all add: gcd_add1)
       
   393 by (simp add: gcd_commute gcd_add1)
       
   394 
       
   395 lemma gcd_sub: "b <= a ==> gcd(a - b,b) = gcd(a,b)" "a <= b ==> gcd(a,b - a) = gcd(a,b)"
       
   396 proof-
       
   397   {fix a b assume H: "b \<le> (a::nat)"
       
   398     hence th: "a - b + b = a" by arith
       
   399     from gcd_add(1)[of "a - b" b] th  have "gcd(a - b,b) = gcd(a,b)" by simp}
       
   400   note th = this
       
   401 {
       
   402   assume ab: "b \<le> a"
       
   403   from th[OF ab] show "gcd (a - b, b) = gcd (a, b)" by blast
       
   404 next
       
   405   assume ab: "a \<le> b"
       
   406   from th[OF ab] show "gcd (a,b - a) = gcd (a, b)" 
       
   407     by (simp add: gcd_commute)}
       
   408 qed
       
   409 
       
   410 (* Coprimality                                                               *)
       
   411 
       
   412 lemma coprime: "coprime a b \<longleftrightarrow> (\<forall>d. d dvd a \<and> d dvd b \<longleftrightarrow> d = 1)"
       
   413 using gcd_unique[of 1 a b, simplified] by (auto simp add: coprime_def)
       
   414 lemma coprime_commute: "coprime a b \<longleftrightarrow> coprime b a" by (simp add: coprime_def gcd_commute)
       
   415 
       
   416 lemma coprime_bezout: "coprime a b \<longleftrightarrow> (\<exists>x y. a * x - b * y = 1 \<or> b * x - a * y = 1)"
       
   417 using coprime_def gcd_bezout by auto
       
   418 
       
   419 lemma coprime_divprod: "d dvd a * b  \<Longrightarrow> coprime d a \<Longrightarrow> d dvd b"
       
   420   using relprime_dvd_mult_iff[of d a b] by (auto simp add: coprime_def mult_commute)
       
   421 
       
   422 lemma coprime_1[simp]: "coprime a 1" by (simp add: coprime_def)
       
   423 lemma coprime_1'[simp]: "coprime 1 a" by (simp add: coprime_def)
       
   424 lemma coprime_Suc0[simp]: "coprime a (Suc 0)" by (simp add: coprime_def)
       
   425 lemma coprime_Suc0'[simp]: "coprime (Suc 0) a" by (simp add: coprime_def)
       
   426 
       
   427 lemma gcd_coprime: 
       
   428   assumes z: "gcd(a,b) \<noteq> 0" and a: "a = a' * gcd(a,b)" and b: "b = b' * gcd(a,b)" 
       
   429   shows    "coprime a' b'"
       
   430 proof-
       
   431   let ?g = "gcd(a,b)"
       
   432   {assume bz: "a = 0" from b bz z a have ?thesis by (simp add: gcd_zero coprime_def)}
       
   433   moreover 
       
   434   {assume az: "a\<noteq> 0" 
       
   435     from z have z': "?g > 0" by simp
       
   436     from bezout_gcd_strong[OF az, of b] 
       
   437     obtain x y where xy: "a*x = b*y + ?g" by blast
       
   438     from xy a b have "?g * a'*x = ?g * (b'*y + 1)" by (simp add: ring_simps)
       
   439     hence "?g * (a'*x) = ?g * (b'*y + 1)" by (simp add: mult_assoc)
       
   440     hence "a'*x = (b'*y + 1)"
       
   441       by (simp only: nat_mult_eq_cancel1[OF z']) 
       
   442     hence "a'*x - b'*y = 1" by simp
       
   443     with coprime_bezout[of a' b'] have ?thesis by auto}
       
   444   ultimately show ?thesis by blast
       
   445 qed
       
   446 lemma coprime_0: "coprime d 0 \<longleftrightarrow> d = 1" by (simp add: coprime_def)
       
   447 lemma coprime_mul: assumes da: "coprime d a" and db: "coprime d b"
       
   448   shows "coprime d (a * b)"
       
   449 proof-
       
   450   from da have th: "gcd(a, d) = 1" by (simp add: coprime_def gcd_commute)
       
   451   from gcd_mult_cancel[of a d b, OF th] db[unfolded coprime_def] have "gcd (d, a*b) = 1"
       
   452     by (simp add: gcd_commute)
       
   453   thus ?thesis unfolding coprime_def .
       
   454 qed
       
   455 lemma coprime_lmul2: assumes dab: "coprime d (a * b)" shows "coprime d b"
       
   456 using prems unfolding coprime_bezout
       
   457 apply clarsimp
       
   458 apply (case_tac "d * x - a * b * y = Suc 0 ", simp_all)
       
   459 apply (rule_tac x="x" in exI)
       
   460 apply (rule_tac x="a*y" in exI)
       
   461 apply (simp add: mult_ac)
       
   462 apply (rule_tac x="a*x" in exI)
       
   463 apply (rule_tac x="y" in exI)
       
   464 apply (simp add: mult_ac)
       
   465 done
       
   466 
       
   467 lemma coprime_rmul2: "coprime d (a * b) \<Longrightarrow> coprime d a"
       
   468 unfolding coprime_bezout
       
   469 apply clarsimp
       
   470 apply (case_tac "d * x - a * b * y = Suc 0 ", simp_all)
       
   471 apply (rule_tac x="x" in exI)
       
   472 apply (rule_tac x="b*y" in exI)
       
   473 apply (simp add: mult_ac)
       
   474 apply (rule_tac x="b*x" in exI)
       
   475 apply (rule_tac x="y" in exI)
       
   476 apply (simp add: mult_ac)
       
   477 done
       
   478 lemma coprime_mul_eq: "coprime d (a * b) \<longleftrightarrow> coprime d a \<and>  coprime d b"
       
   479   using coprime_rmul2[of d a b] coprime_lmul2[of d a b] coprime_mul[of d a b] 
       
   480   by blast
       
   481 
       
   482 lemma gcd_coprime_exists:
       
   483   assumes nz: "gcd(a,b) \<noteq> 0" 
       
   484   shows "\<exists>a' b'. a = a' * gcd(a,b) \<and> b = b' * gcd(a,b) \<and> coprime a' b'"
       
   485 proof-
       
   486   let ?g = "gcd (a,b)"
       
   487   from gcd_dvd1[of a b] gcd_dvd2[of a b] 
       
   488   obtain a' b' where "a = ?g*a'"  "b = ?g*b'" unfolding dvd_def by blast
       
   489   hence ab': "a = a'*?g" "b = b'*?g" by algebra+
       
   490   from ab' gcd_coprime[OF nz ab'] show ?thesis by blast
       
   491 qed
       
   492 
       
   493 lemma coprime_exp: "coprime d a ==> coprime d (a^n)" 
       
   494   by(induct n, simp_all add: coprime_mul)
       
   495 
       
   496 lemma coprime_exp_imp: "coprime a b ==> coprime (a ^n) (b ^n)"
       
   497   by (induct n, simp_all add: coprime_mul_eq coprime_commute coprime_exp)
       
   498 lemma coprime_refl[simp]: "coprime n n \<longleftrightarrow> n = 1" by (simp add: coprime_def)
       
   499 lemma coprime_plus1[simp]: "coprime (n + 1) n"
       
   500   apply (simp add: coprime_bezout)
       
   501   apply (rule exI[where x=1])
       
   502   apply (rule exI[where x=1])
       
   503   apply simp
       
   504   done
       
   505 lemma coprime_minus1: "n \<noteq> 0 ==> coprime (n - 1) n"
       
   506   using coprime_plus1[of "n - 1"] coprime_commute[of "n - 1" n] by auto
       
   507 
       
   508 lemma bezout_gcd_pow: "\<exists>x y. a ^n * x - b ^ n * y = gcd(a,b) ^ n \<or> b ^ n * x - a ^ n * y = gcd(a,b) ^ n"
       
   509 proof-
       
   510   let ?g = "gcd (a,b)"
       
   511   {assume z: "?g = 0" hence ?thesis 
       
   512       apply (cases n, simp)
       
   513       apply arith
       
   514       apply (simp only: z power_0_Suc)
       
   515       apply (rule exI[where x=0])
       
   516       apply (rule exI[where x=0])
       
   517       by simp}
       
   518   moreover
       
   519   {assume z: "?g \<noteq> 0"
       
   520     from gcd_dvd1[of a b] gcd_dvd2[of a b] obtain a' b' where
       
   521       ab': "a = a'*?g" "b = b'*?g" unfolding dvd_def by (auto simp add: mult_ac)
       
   522     hence ab'': "?g*a' = a" "?g * b' = b" by algebra+
       
   523     from coprime_exp_imp[OF gcd_coprime[OF z ab'], unfolded coprime_bezout, of n]
       
   524     obtain x y where "a'^n * x - b'^n * y = 1 \<or> b'^n * x - a'^n * y = 1"  by blast
       
   525     hence "?g^n * (a'^n * x - b'^n * y) = ?g^n \<or> ?g^n*(b'^n * x - a'^n * y) = ?g^n"
       
   526       using z by auto 
       
   527     then have "a^n * x - b^n * y = ?g^n \<or> b^n * x - a^n * y = ?g^n"
       
   528       using z ab'' by (simp only: power_mult_distrib[symmetric] 
       
   529 	diff_mult_distrib2 mult_assoc[symmetric])
       
   530     hence  ?thesis by blast }
       
   531   ultimately show ?thesis by blast
       
   532 qed
       
   533 lemma gcd_exp: "gcd (a^n, b^n) = gcd(a,b)^n"
       
   534 proof-
       
   535   let ?g = "gcd(a^n,b^n)"
       
   536   let ?gn = "gcd(a,b)^n"
       
   537   {fix e assume H: "e dvd a^n" "e dvd b^n"
       
   538     from bezout_gcd_pow[of a n b] obtain x y 
       
   539       where xy: "a ^ n * x - b ^ n * y = ?gn \<or> b ^ n * x - a ^ n * y = ?gn" by blast
       
   540     from dvd_diff [OF dvd_mult2[OF H(1), of x] dvd_mult2[OF H(2), of y]]
       
   541       dvd_diff [OF dvd_mult2[OF H(2), of x] dvd_mult2[OF H(1), of y]] xy
       
   542     have "e dvd ?gn" by (cases "a ^ n * x - b ^ n * y = gcd (a, b) ^ n", simp_all)}
       
   543   hence th:  "\<forall>e. e dvd a^n \<and> e dvd b^n \<longrightarrow> e dvd ?gn" by blast
       
   544   from divides_exp[OF gcd_dvd1[of a b], of n] divides_exp[OF gcd_dvd2[of a b], of n] th
       
   545     gcd_unique have "?gn = ?g" by blast thus ?thesis by simp 
       
   546 qed
       
   547 
       
   548 lemma coprime_exp2:  "coprime (a ^ Suc n) (b^ Suc n) \<longleftrightarrow> coprime a b"
       
   549 by (simp only: coprime_def gcd_exp exp_eq_1) simp
       
   550 
       
   551 lemma division_decomp: assumes dc: "(a::nat) dvd b * c"
       
   552   shows "\<exists>b' c'. a = b' * c' \<and> b' dvd b \<and> c' dvd c"
       
   553 proof-
       
   554   let ?g = "gcd (a,b)"
       
   555   {assume "?g = 0" with dc have ?thesis apply (simp add: gcd_zero)
       
   556       apply (rule exI[where x="0"])
       
   557       by (rule exI[where x="c"], simp)}
       
   558   moreover
       
   559   {assume z: "?g \<noteq> 0"
       
   560     from gcd_coprime_exists[OF z]
       
   561     obtain a' b' where ab': "a = a' * ?g" "b = b' * ?g" "coprime a' b'" by blast
       
   562     from gcd_dvd2[of a b] have thb: "?g dvd b" .
       
   563     from ab'(1) have "a' dvd a"  unfolding dvd_def by blast  
       
   564     with dc have th0: "a' dvd b*c" using dvd_trans[of a' a "b*c"] by simp
       
   565     from dc ab'(1,2) have "a'*?g dvd (b'*?g) *c" by auto
       
   566     hence "?g*a' dvd ?g * (b' * c)" by (simp add: mult_assoc)
       
   567     with z have th_1: "a' dvd b'*c" by simp
       
   568     from coprime_divprod[OF th_1 ab'(3)] have thc: "a' dvd c" . 
       
   569     from ab' have "a = ?g*a'" by algebra
       
   570     with thb thc have ?thesis by blast }
       
   571   ultimately show ?thesis by blast
       
   572 qed
       
   573 
       
   574 lemma nat_power_eq_0_iff: "(m::nat) ^ n = 0 \<longleftrightarrow> n \<noteq> 0 \<and> m = 0" by (induct n, auto)
       
   575 
       
   576 lemma divides_rev: assumes ab: "(a::nat) ^ n dvd b ^n" and n:"n \<noteq> 0" shows "a dvd b"
       
   577 proof-
       
   578   let ?g = "gcd (a,b)"
       
   579   from n obtain m where m: "n = Suc m" by (cases n, simp_all)
       
   580   {assume "?g = 0" with ab n have ?thesis by (simp add: gcd_zero)}
       
   581   moreover
       
   582   {assume z: "?g \<noteq> 0"
       
   583     hence zn: "?g ^ n \<noteq> 0" using n by (simp add: neq0_conv)
       
   584     from gcd_coprime_exists[OF z] 
       
   585     obtain a' b' where ab': "a = a' * ?g" "b = b' * ?g" "coprime a' b'" by blast
       
   586     from ab have "(a' * ?g) ^ n dvd (b' * ?g)^n" by (simp add: ab'(1,2)[symmetric])
       
   587     hence "?g^n*a'^n dvd ?g^n *b'^n" by (simp only: power_mult_distrib mult_commute)
       
   588     with zn z n have th0:"a'^n dvd b'^n" by (auto simp add: nat_power_eq_0_iff)
       
   589     have "a' dvd a'^n" by (simp add: m)
       
   590     with th0 have "a' dvd b'^n" using dvd_trans[of a' "a'^n" "b'^n"] by simp
       
   591     hence th1: "a' dvd b'^m * b'" by (simp add: m mult_commute)
       
   592     from coprime_divprod[OF th1 coprime_exp[OF ab'(3), of m]]
       
   593     have "a' dvd b'" .
       
   594     hence "a'*?g dvd b'*?g" by simp
       
   595     with ab'(1,2)  have ?thesis by simp }
       
   596   ultimately show ?thesis by blast
       
   597 qed
       
   598 
       
   599 lemma divides_mul: assumes mr: "m dvd r" and nr: "n dvd r" and mn:"coprime m n" 
       
   600   shows "m * n dvd r"
       
   601 proof-
       
   602   from mr nr obtain m' n' where m': "r = m*m'" and n': "r = n*n'"
       
   603     unfolding dvd_def by blast
       
   604   from mr n' have "m dvd n'*n" by (simp add: mult_commute)
       
   605   hence "m dvd n'" using relprime_dvd_mult_iff[OF mn[unfolded coprime_def]] by simp
       
   606   then obtain k where k: "n' = m*k" unfolding dvd_def by blast
       
   607   from n' k show ?thesis unfolding dvd_def by auto
       
   608 qed
       
   609 
       
   610 (* A binary form of the Chinese Remainder Theorem.                           *)
       
   611 
       
   612 lemma chinese_remainder: assumes ab: "coprime a b" and a:"a \<noteq> 0" and b:"b \<noteq> 0"
       
   613   shows "\<exists>x q1 q2. x = u + q1 * a \<and> x = v + q2 * b"
       
   614 proof-
       
   615   from bezout_add_strong[OF a, of b] bezout_add_strong[OF b, of a]
       
   616   obtain d1 x1 y1 d2 x2 y2 where dxy1: "d1 dvd a" "d1 dvd b" "a * x1 = b * y1 + d1" 
       
   617     and dxy2: "d2 dvd b" "d2 dvd a" "b * x2 = a * y2 + d2" by blast
       
   618   from gcd_unique[of 1 a b, simplified ab[unfolded coprime_def], simplified] 
       
   619     dxy1(1,2) dxy2(1,2) have d12: "d1 = 1" "d2 =1" by auto
       
   620   let ?x = "v * a * x1 + u * b * x2"
       
   621   let ?q1 = "v * x1 + u * y2"
       
   622   let ?q2 = "v * y1 + u * x2"
       
   623   from dxy2(3)[simplified d12] dxy1(3)[simplified d12] 
       
   624   have "?x = u + ?q1 * a" "?x = v + ?q2 * b" by algebra+ 
       
   625   thus ?thesis by blast
       
   626 qed
       
   627 
       
   628 (* Primality                                                                 *)
       
   629 (* A few useful theorems about primes                                        *)
       
   630 
       
   631 lemma prime_0[simp]: "~prime 0" by (simp add: prime_def)
       
   632 lemma prime_1[simp]: "~ prime 1"  by (simp add: prime_def)
       
   633 lemma prime_Suc0[simp]: "~ prime (Suc 0)"  by (simp add: prime_def)
       
   634 
       
   635 lemma prime_ge_2: "prime p ==> p \<ge> 2" by (simp add: prime_def)
       
   636 lemma prime_factor: assumes n: "n \<noteq> 1" shows "\<exists> p. prime p \<and> p dvd n"
       
   637 using n
       
   638 proof(induct n rule: nat_less_induct)
       
   639   fix n
       
   640   assume H: "\<forall>m<n. m \<noteq> 1 \<longrightarrow> (\<exists>p. prime p \<and> p dvd m)" "n \<noteq> 1"
       
   641   let ?ths = "\<exists>p. prime p \<and> p dvd n"
       
   642   {assume "n=0" hence ?ths using two_is_prime by auto}
       
   643   moreover
       
   644   {assume nz: "n\<noteq>0" 
       
   645     {assume "prime n" hence ?ths by - (rule exI[where x="n"], simp)}
       
   646     moreover
       
   647     {assume n: "\<not> prime n"
       
   648       with nz H(2) 
       
   649       obtain k where k:"k dvd n" "k \<noteq> 1" "k \<noteq> n" by (auto simp add: prime_def) 
       
   650       from dvd_imp_le[OF k(1)] nz k(3) have kn: "k < n" by simp
       
   651       from H(1)[rule_format, OF kn k(2)] obtain p where p: "prime p" "p dvd k" by blast
       
   652       from dvd_trans[OF p(2) k(1)] p(1) have ?ths by blast}
       
   653     ultimately have ?ths by blast}
       
   654   ultimately show ?ths by blast
       
   655 qed
       
   656 
       
   657 lemma prime_factor_lt: assumes p: "prime p" and n: "n \<noteq> 0" and npm:"n = p * m"
       
   658   shows "m < n"
       
   659 proof-
       
   660   {assume "m=0" with n have ?thesis by simp}
       
   661   moreover
       
   662   {assume m: "m \<noteq> 0"
       
   663     from npm have mn: "m dvd n" unfolding dvd_def by auto
       
   664     from npm m have "n \<noteq> m" using p by auto
       
   665     with dvd_imp_le[OF mn] n have ?thesis by simp}
       
   666   ultimately show ?thesis by blast
       
   667 qed
       
   668 
       
   669 lemma euclid_bound: "\<exists>p. prime p \<and> n < p \<and>  p <= Suc (fact n)"
       
   670 proof-
       
   671   have f1: "fact n + 1 \<noteq> 1" using fact_le[of n] by arith 
       
   672   from prime_factor[OF f1] obtain p where p: "prime p" "p dvd fact n + 1" by blast
       
   673   from dvd_imp_le[OF p(2)] have pfn: "p \<le> fact n + 1" by simp
       
   674   {assume np: "p \<le> n"
       
   675     from p(1) have p1: "p \<ge> 1" by (cases p, simp_all)
       
   676     from divides_fact[OF p1 np] have pfn': "p dvd fact n" .
       
   677     from divides_add_revr[OF pfn' p(2)] p(1) have False by simp}
       
   678   hence "n < p" by arith
       
   679   with p(1) pfn show ?thesis by auto
       
   680 qed
       
   681 
       
   682 lemma euclid: "\<exists>p. prime p \<and> p > n" using euclid_bound by auto
       
   683 lemma primes_infinite: "\<not> (finite {p. prime p})"
       
   684 proof (auto simp add: finite_conv_nat_seg_image)
       
   685   fix n f 
       
   686   assume H: "Collect prime = f ` {i. i < (n::nat)}"
       
   687   let ?P = "Collect prime"
       
   688   let ?m = "Max ?P"
       
   689   have P0: "?P \<noteq> {}" using two_is_prime by auto
       
   690   from H have fP: "finite ?P" using finite_conv_nat_seg_image by blast 
       
   691   from Max_in[OF fP P0]  have "?m \<in> ?P" . 
       
   692   from Max_ge[OF fP P0] have contr: "\<forall> p. prime p \<longrightarrow> p \<le> ?m" by blast
       
   693   from euclid[of ?m] obtain q where q: "prime q" "q > ?m" by blast
       
   694   with contr show False by auto
       
   695 qed
       
   696 
       
   697 lemma coprime_prime: assumes ab: "coprime a b"
       
   698   shows "~(prime p \<and> p dvd a \<and> p dvd b)"
       
   699 proof
       
   700   assume "prime p \<and> p dvd a \<and> p dvd b"
       
   701   thus False using ab gcd_greatest[of p a b] by (simp add: coprime_def)
       
   702 qed
       
   703 lemma coprime_prime_eq: "coprime a b \<longleftrightarrow> (\<forall>p. ~(prime p \<and> p dvd a \<and> p dvd b))" 
       
   704   (is "?lhs = ?rhs")
       
   705 proof-
       
   706   {assume "?lhs" with coprime_prime  have ?rhs by blast}
       
   707   moreover
       
   708   {assume r: "?rhs" and c: "\<not> ?lhs"
       
   709     then obtain g where g: "g\<noteq>1" "g dvd a" "g dvd b" unfolding coprime_def by blast
       
   710     from prime_factor[OF g(1)] obtain p where p: "prime p" "p dvd g" by blast
       
   711     from dvd_trans [OF p(2) g(2)] dvd_trans [OF p(2) g(3)] 
       
   712     have "p dvd a" "p dvd b" . with p(1) r have False by blast}
       
   713   ultimately show ?thesis by blast
       
   714 qed
       
   715 
       
   716 lemma prime_coprime: assumes p: "prime p" 
       
   717   shows "n = 1 \<or> p dvd n \<or> coprime p n"
       
   718 using p prime_imp_relprime[of p n] by (auto simp add: coprime_def)
       
   719 
       
   720 lemma prime_coprime_strong: "prime p \<Longrightarrow> p dvd n \<or> coprime p n"
       
   721   using prime_coprime[of p n] by auto
       
   722 
       
   723 declare  coprime_0[simp]
       
   724 
       
   725 lemma coprime_0'[simp]: "coprime 0 d \<longleftrightarrow> d = 1" by (simp add: coprime_commute[of 0 d])
       
   726 lemma coprime_bezout_strong: assumes ab: "coprime a b" and b: "b \<noteq> 1"
       
   727   shows "\<exists>x y. a * x = b * y + 1"
       
   728 proof-
       
   729   from ab b have az: "a \<noteq> 0" by - (rule ccontr, auto)
       
   730   from bezout_gcd_strong[OF az, of b] ab[unfolded coprime_def]
       
   731   show ?thesis by auto
       
   732 qed
       
   733 
       
   734 lemma bezout_prime: assumes p: "prime p"  and pa: "\<not> p dvd a"
       
   735   shows "\<exists>x y. a*x = p*y + 1"
       
   736 proof-
       
   737   from p have p1: "p \<noteq> 1" using prime_1 by blast 
       
   738   from prime_coprime[OF p, of a] p1 pa have ap: "coprime a p" 
       
   739     by (auto simp add: coprime_commute)
       
   740   from coprime_bezout_strong[OF ap p1] show ?thesis . 
       
   741 qed
       
   742 lemma prime_divprod: assumes p: "prime p" and pab: "p dvd a*b"
       
   743   shows "p dvd a \<or> p dvd b"
       
   744 proof-
       
   745   {assume "a=1" hence ?thesis using pab by simp }
       
   746   moreover
       
   747   {assume "p dvd a" hence ?thesis by blast}
       
   748   moreover
       
   749   {assume pa: "coprime p a" from coprime_divprod[OF pab pa]  have ?thesis .. }
       
   750   ultimately show ?thesis using prime_coprime[OF p, of a] by blast
       
   751 qed
       
   752 
       
   753 lemma prime_divprod_eq: assumes p: "prime p"
       
   754   shows "p dvd a*b \<longleftrightarrow> p dvd a \<or> p dvd b"
       
   755 using p prime_divprod dvd_mult dvd_mult2 by auto
       
   756 
       
   757 lemma prime_divexp: assumes p:"prime p" and px: "p dvd x^n"
       
   758   shows "p dvd x"
       
   759 using px
       
   760 proof(induct n)
       
   761   case 0 thus ?case by simp
       
   762 next
       
   763   case (Suc n) 
       
   764   hence th: "p dvd x*x^n" by simp
       
   765   {assume H: "p dvd x^n"
       
   766     from Suc.hyps[OF H] have ?case .}
       
   767   with prime_divprod[OF p th] show ?case by blast
       
   768 qed
       
   769 
       
   770 lemma prime_divexp_n: "prime p \<Longrightarrow> p dvd x^n \<Longrightarrow> p^n dvd x^n"
       
   771   using prime_divexp[of p x n] divides_exp[of p x n] by blast
       
   772 
       
   773 lemma coprime_prime_dvd_ex: assumes xy: "\<not>coprime x y"
       
   774   shows "\<exists>p. prime p \<and> p dvd x \<and> p dvd y"
       
   775 proof-
       
   776   from xy[unfolded coprime_def] obtain g where g: "g \<noteq> 1" "g dvd x" "g dvd y" 
       
   777     by blast
       
   778   from prime_factor[OF g(1)] obtain p where p: "prime p" "p dvd g" by blast
       
   779   from g(2,3) dvd_trans[OF p(2)] p(1) show ?thesis by auto
       
   780 qed
       
   781 lemma coprime_sos: assumes xy: "coprime x y" 
       
   782   shows "coprime (x * y) (x^2 + y^2)"
       
   783 proof-
       
   784   {assume c: "\<not> coprime (x * y) (x^2 + y^2)"
       
   785     from coprime_prime_dvd_ex[OF c] obtain p 
       
   786       where p: "prime p" "p dvd x*y" "p dvd x^2 + y^2" by blast
       
   787     {assume px: "p dvd x"
       
   788       from dvd_mult[OF px, of x] p(3) have "p dvd y^2" 
       
   789 	unfolding dvd_def 
       
   790 	apply (auto simp add: power2_eq_square)
       
   791 	apply (rule_tac x= "ka - k" in exI)
       
   792 	by (simp add: diff_mult_distrib2)
       
   793       with prime_divexp[OF p(1), of y 2] have py: "p dvd y" .
       
   794       from p(1) px py xy[unfolded coprime, rule_format, of p] prime_1  
       
   795       have False by simp }
       
   796     moreover
       
   797     {assume py: "p dvd y"
       
   798       from dvd_mult[OF py, of y] p(3) have "p dvd x^2" 
       
   799 	unfolding dvd_def 
       
   800 	apply (auto simp add: power2_eq_square)
       
   801 	apply (rule_tac x= "ka - k" in exI)
       
   802 	by (simp add: diff_mult_distrib2)
       
   803       with prime_divexp[OF p(1), of x 2] have px: "p dvd x" .
       
   804       from p(1) px py xy[unfolded coprime, rule_format, of p] prime_1  
       
   805       have False by simp }
       
   806     ultimately have False using prime_divprod[OF p(1,2)] by blast}
       
   807   thus ?thesis by blast
       
   808 qed
       
   809 
       
   810 lemma distinct_prime_coprime: "prime p \<Longrightarrow> prime q \<Longrightarrow> p \<noteq> q \<Longrightarrow> coprime p q"
       
   811   unfolding prime_def coprime_prime_eq by blast
       
   812 
       
   813 lemma prime_coprime_lt: assumes p: "prime p" and x: "0 < x" and xp: "x < p"
       
   814   shows "coprime x p"
       
   815 proof-
       
   816   {assume c: "\<not> coprime x p"
       
   817     then obtain g where g: "g \<noteq> 1" "g dvd x" "g dvd p" unfolding coprime_def by blast
       
   818   from dvd_imp_le[OF g(2)] x xp have gp: "g < p" by arith
       
   819   from g(2) x have "g \<noteq> 0" by - (rule ccontr, simp)
       
   820   with g gp p[unfolded prime_def] have False by blast}
       
   821 thus ?thesis by blast
       
   822 qed
       
   823 
       
   824 lemma even_dvd[simp]: "even (n::nat) \<longleftrightarrow> 2 dvd n" by presburger
       
   825 lemma prime_odd: "prime p \<Longrightarrow> p = 2 \<or> odd p" unfolding prime_def by auto
       
   826 
       
   827 (* One property of coprimality is easier to prove via prime factors.         *)
       
   828 
       
   829 lemma prime_divprod_pow: 
       
   830   assumes p: "prime p" and ab: "coprime a b" and pab: "p^n dvd a * b"
       
   831   shows "p^n dvd a \<or> p^n dvd b"
       
   832 proof-
       
   833   {assume "n = 0 \<or> a = 1 \<or> b = 1" with pab have ?thesis 
       
   834       apply (cases "n=0", simp_all)
       
   835       apply (cases "a=1", simp_all) done}
       
   836   moreover
       
   837   {assume n: "n \<noteq> 0" and a: "a\<noteq>1" and b: "b\<noteq>1" 
       
   838     then obtain m where m: "n = Suc m" by (cases n, auto)
       
   839     from divides_exp2[OF n pab] have pab': "p dvd a*b" .
       
   840     from prime_divprod[OF p pab'] 
       
   841     have "p dvd a \<or> p dvd b" .
       
   842     moreover
       
   843     {assume pa: "p dvd a"
       
   844       have pnba: "p^n dvd b*a" using pab by (simp add: mult_commute)
       
   845       from coprime_prime[OF ab, of p] p pa have "\<not> p dvd b" by blast
       
   846       with prime_coprime[OF p, of b] b 
       
   847       have cpb: "coprime b p" using coprime_commute by blast 
       
   848       from coprime_exp[OF cpb] have pnb: "coprime (p^n) b" 
       
   849 	by (simp add: coprime_commute)
       
   850       from coprime_divprod[OF pnba pnb] have ?thesis by blast }
       
   851     moreover
       
   852     {assume pb: "p dvd b"
       
   853       have pnba: "p^n dvd b*a" using pab by (simp add: mult_commute)
       
   854       from coprime_prime[OF ab, of p] p pb have "\<not> p dvd a" by blast
       
   855       with prime_coprime[OF p, of a] a
       
   856       have cpb: "coprime a p" using coprime_commute by blast 
       
   857       from coprime_exp[OF cpb] have pnb: "coprime (p^n) a" 
       
   858 	by (simp add: coprime_commute)
       
   859       from coprime_divprod[OF pab pnb] have ?thesis by blast }
       
   860     ultimately have ?thesis by blast}
       
   861   ultimately show ?thesis by blast
       
   862 qed
       
   863 
       
   864 lemma nat_mult_eq_one: "(n::nat) * m = 1 \<longleftrightarrow> n = 1 \<and> m = 1" (is "?lhs \<longleftrightarrow> ?rhs")
       
   865 proof
       
   866   assume H: "?lhs"
       
   867   hence "n dvd 1" "m dvd 1" unfolding dvd_def by (auto simp add: mult_commute)
       
   868   thus ?rhs by auto
       
   869 next
       
   870   assume ?rhs then show ?lhs by auto
       
   871 qed
       
   872   
       
   873 lemma power_Suc0[simp]: "Suc 0 ^ n = Suc 0" 
       
   874   unfolding One_nat_def[symmetric] power_one ..
       
   875 lemma coprime_pow: assumes ab: "coprime a b" and abcn: "a * b = c ^n"
       
   876   shows "\<exists>r s. a = r^n  \<and> b = s ^n"
       
   877   using ab abcn
       
   878 proof(induct c arbitrary: a b rule: nat_less_induct)
       
   879   fix c a b
       
   880   assume H: "\<forall>m<c. \<forall>a b. coprime a b \<longrightarrow> a * b = m ^ n \<longrightarrow> (\<exists>r s. a = r ^ n \<and> b = s ^ n)" "coprime a b" "a * b = c ^ n" 
       
   881   let ?ths = "\<exists>r s. a = r^n  \<and> b = s ^n"
       
   882   {assume n: "n = 0"
       
   883     with H(3) power_one have "a*b = 1" by simp
       
   884     hence "a = 1 \<and> b = 1" by simp
       
   885     hence ?ths 
       
   886       apply -
       
   887       apply (rule exI[where x=1])
       
   888       apply (rule exI[where x=1])
       
   889       using power_one[of  n]
       
   890       by simp}
       
   891   moreover
       
   892   {assume n: "n \<noteq> 0" then obtain m where m: "n = Suc m" by (cases n, auto)
       
   893     {assume c: "c = 0"
       
   894       with H(3) m H(2) have ?ths apply simp 
       
   895 	apply (cases "a=0", simp_all) 
       
   896 	apply (rule exI[where x="0"], simp)
       
   897 	apply (rule exI[where x="0"], simp)
       
   898 	done}
       
   899     moreover
       
   900     {assume "c=1" with H(3) power_one have "a*b = 1" by simp 
       
   901 	hence "a = 1 \<and> b = 1" by simp
       
   902 	hence ?ths 
       
   903       apply -
       
   904       apply (rule exI[where x=1])
       
   905       apply (rule exI[where x=1])
       
   906       using power_one[of  n]
       
   907       by simp}
       
   908   moreover
       
   909   {assume c: "c\<noteq>1" "c \<noteq> 0"
       
   910     from prime_factor[OF c(1)] obtain p where p: "prime p" "p dvd c" by blast
       
   911     from prime_divprod_pow[OF p(1) H(2), unfolded H(3), OF divides_exp[OF p(2), of n]] 
       
   912     have pnab: "p ^ n dvd a \<or> p^n dvd b" . 
       
   913     from p(2) obtain l where l: "c = p*l" unfolding dvd_def by blast
       
   914     have pn0: "p^n \<noteq> 0" using n prime_ge_2 [OF p(1)] by (simp add: neq0_conv)
       
   915     {assume pa: "p^n dvd a"
       
   916       then obtain k where k: "a = p^n * k" unfolding dvd_def by blast
       
   917       from l have "l dvd c" by auto
       
   918       with dvd_imp_le[of l c] c have "l \<le> c" by auto
       
   919       moreover {assume "l = c" with l c  have "p = 1" by simp with p have False by simp}
       
   920       ultimately have lc: "l < c" by arith
       
   921       from coprime_lmul2 [OF H(2)[unfolded k coprime_commute[of "p^n*k" b]]]
       
   922       have kb: "coprime k b" by (simp add: coprime_commute) 
       
   923       from H(3) l k pn0 have kbln: "k * b = l ^ n" 
       
   924 	by (auto simp add: power_mult_distrib)
       
   925       from H(1)[rule_format, OF lc kb kbln]
       
   926       obtain r s where rs: "k = r ^n" "b = s^n" by blast
       
   927       from k rs(1) have "a = (p*r)^n" by (simp add: power_mult_distrib)
       
   928       with rs(2) have ?ths by blast }
       
   929     moreover
       
   930     {assume pb: "p^n dvd b"
       
   931       then obtain k where k: "b = p^n * k" unfolding dvd_def by blast
       
   932       from l have "l dvd c" by auto
       
   933       with dvd_imp_le[of l c] c have "l \<le> c" by auto
       
   934       moreover {assume "l = c" with l c  have "p = 1" by simp with p have False by simp}
       
   935       ultimately have lc: "l < c" by arith
       
   936       from coprime_lmul2 [OF H(2)[unfolded k coprime_commute[of "p^n*k" a]]]
       
   937       have kb: "coprime k a" by (simp add: coprime_commute) 
       
   938       from H(3) l k pn0 n have kbln: "k * a = l ^ n" 
       
   939 	by (simp add: power_mult_distrib mult_commute)
       
   940       from H(1)[rule_format, OF lc kb kbln]
       
   941       obtain r s where rs: "k = r ^n" "a = s^n" by blast
       
   942       from k rs(1) have "b = (p*r)^n" by (simp add: power_mult_distrib)
       
   943       with rs(2) have ?ths by blast }
       
   944     ultimately have ?ths using pnab by blast}
       
   945   ultimately have ?ths by blast}
       
   946 ultimately show ?ths by blast
       
   947 qed
       
   948 
       
   949 (* More useful lemmas.                                                       *)
       
   950 lemma prime_product: 
       
   951   "prime (p*q) \<Longrightarrow> p = 1 \<or> q  = 1" unfolding prime_def by auto
       
   952 
       
   953 lemma prime_exp: "prime (p^n) \<longleftrightarrow> prime p \<and> n = 1"
       
   954 proof(induct n)
       
   955   case 0 thus ?case by simp
       
   956 next
       
   957   case (Suc n)
       
   958   {assume "p = 0" hence ?case by simp}
       
   959   moreover
       
   960   {assume "p=1" hence ?case by simp}
       
   961   moreover
       
   962   {assume p: "p \<noteq> 0" "p\<noteq>1"
       
   963     {assume pp: "prime (p^Suc n)"
       
   964       hence "p = 1 \<or> p^n = 1" using prime_product[of p "p^n"] by simp
       
   965       with p have n: "n = 0" 
       
   966 	by (simp only: exp_eq_1 ) simp
       
   967       with pp have "prime p \<and> Suc n = 1" by simp}
       
   968     moreover
       
   969     {assume n: "prime p \<and> Suc n = 1" hence "prime (p^Suc n)" by simp}
       
   970     ultimately have ?case by blast}
       
   971   ultimately show ?case by blast
       
   972 qed
       
   973 
       
   974 lemma prime_power_mult: 
       
   975   assumes p: "prime p" and xy: "x * y = p ^ k"
       
   976   shows "\<exists>i j. x = p ^i \<and> y = p^ j"
       
   977   using xy
       
   978 proof(induct k arbitrary: x y)
       
   979   case 0 thus ?case apply simp by (rule exI[where x="0"], simp)
       
   980 next
       
   981   case (Suc k x y)
       
   982   from Suc.prems have pxy: "p dvd x*y" by auto
       
   983   from prime_divprod[OF p pxy] have pxyc: "p dvd x \<or> p dvd y" .
       
   984   from p have p0: "p \<noteq> 0" by - (rule ccontr, simp) 
       
   985   {assume px: "p dvd x"
       
   986     then obtain d where d: "x = p*d" unfolding dvd_def by blast
       
   987     from Suc.prems d  have "p*d*y = p^Suc k" by simp
       
   988     hence th: "d*y = p^k" using p0 by simp
       
   989     from Suc.hyps[OF th] obtain i j where ij: "d = p^i" "y = p^j" by blast
       
   990     with d have "x = p^Suc i" by simp 
       
   991     with ij(2) have ?case by blast}
       
   992   moreover 
       
   993   {assume px: "p dvd y"
       
   994     then obtain d where d: "y = p*d" unfolding dvd_def by blast
       
   995     from Suc.prems d  have "p*d*x = p^Suc k" by (simp add: mult_commute)
       
   996     hence th: "d*x = p^k" using p0 by simp
       
   997     from Suc.hyps[OF th] obtain i j where ij: "d = p^i" "x = p^j" by blast
       
   998     with d have "y = p^Suc i" by simp 
       
   999     with ij(2) have ?case by blast}
       
  1000   ultimately show ?case  using pxyc by blast
       
  1001 qed
       
  1002 
       
  1003 lemma prime_power_exp: assumes p: "prime p" and n:"n \<noteq> 0" 
       
  1004   and xn: "x^n = p^k" shows "\<exists>i. x = p^i"
       
  1005   using n xn
       
  1006 proof(induct n arbitrary: k)
       
  1007   case 0 thus ?case by simp
       
  1008 next
       
  1009   case (Suc n k) hence th: "x*x^n = p^k" by simp
       
  1010   {assume "n = 0" with prems have ?case apply simp 
       
  1011       by (rule exI[where x="k"],simp)}
       
  1012   moreover
       
  1013   {assume n: "n \<noteq> 0"
       
  1014     from prime_power_mult[OF p th] 
       
  1015     obtain i j where ij: "x = p^i" "x^n = p^j"by blast
       
  1016     from Suc.hyps[OF n ij(2)] have ?case .}
       
  1017   ultimately show ?case by blast
       
  1018 qed
       
  1019 
       
  1020 lemma divides_primepow: assumes p: "prime p" 
       
  1021   shows "d dvd p^k \<longleftrightarrow> (\<exists> i. i \<le> k \<and> d = p ^i)"
       
  1022 proof
       
  1023   assume H: "d dvd p^k" then obtain e where e: "d*e = p^k" 
       
  1024     unfolding dvd_def  apply (auto simp add: mult_commute) by blast
       
  1025   from prime_power_mult[OF p e] obtain i j where ij: "d = p^i" "e=p^j" by blast
       
  1026   from prime_ge_2[OF p] have p1: "p > 1" by arith
       
  1027   from e ij have "p^(i + j) = p^k" by (simp add: power_add)
       
  1028   hence "i + j = k" using power_inject_exp[of p "i+j" k, OF p1] by simp 
       
  1029   hence "i \<le> k" by arith
       
  1030   with ij(1) show "\<exists>i\<le>k. d = p ^ i" by blast
       
  1031 next
       
  1032   {fix i assume H: "i \<le> k" "d = p^i"
       
  1033     hence "\<exists>j. k = i + j" by arith
       
  1034     then obtain j where j: "k = i + j" by blast
       
  1035     hence "p^k = p^j*d" using H(2) by (simp add: power_add)
       
  1036     hence "d dvd p^k" unfolding dvd_def by auto}
       
  1037   thus "\<exists>i\<le>k. d = p ^ i \<Longrightarrow> d dvd p ^ k" by blast
       
  1038 qed
       
  1039 
       
  1040 lemma coprime_divisors: "d dvd a \<Longrightarrow> e dvd b \<Longrightarrow> coprime a b \<Longrightarrow> coprime d e"
       
  1041   by (auto simp add: dvd_def coprime)
       
  1042 
    48 end
  1043 end