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