src/HOL/Library/Primes.thy
author nipkow
Fri, 25 May 2007 18:08:34 +0200
changeset 23100 1c84d7294d5b
parent 22665 cf152ff55d16
child 23839 d9fa0f457d9a
permissions -rw-r--r--
Added List_Comprehension
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$
a548865b1b6a moved Primes.thy from NumberTheory to Library
paulson
parents:
diff changeset
     3
    Author:     Christophe Tabacznyj and Lawrence C Paulson
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
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents: 19086
diff changeset
    10
imports GCD
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
19086
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 16762
diff changeset
    15
  "coprime m n = (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
19086
1b3780be6cc2 new-style definitions/abbreviations;
wenzelm
parents: 16762
diff changeset
    19
  "prime p = (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
16663
13e9c402308b prime is a predicate now.
nipkow
parents: 15628
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)
a548865b1b6a moved Primes.thy from NumberTheory to Library
paulson
parents:
diff changeset
    30
  apply (drule_tac x = "gcd (p, n)" in spec)
a548865b1b6a moved Primes.thy from NumberTheory to Library
paulson
parents:
diff changeset
    31
  apply auto
a548865b1b6a moved Primes.thy from NumberTheory to Library
paulson
parents:
diff changeset
    32
  apply (insert gcd_dvd2 [of p n])
a548865b1b6a moved Primes.thy from NumberTheory to Library
paulson
parents:
diff changeset
    33
  apply simp
a548865b1b6a moved Primes.thy from NumberTheory to Library
paulson
parents:
diff changeset
    34
  done
a548865b1b6a moved Primes.thy from NumberTheory to Library
paulson
parents:
diff changeset
    35
a548865b1b6a moved Primes.thy from NumberTheory to Library
paulson
parents:
diff changeset
    36
text {*
a548865b1b6a moved Primes.thy from NumberTheory to Library
paulson
parents:
diff changeset
    37
  This theorem leads immediately to a proof of the uniqueness of
a548865b1b6a moved Primes.thy from NumberTheory to Library
paulson
parents:
diff changeset
    38
  factorization.  If @{term p} divides a product of primes then it is
a548865b1b6a moved Primes.thy from NumberTheory to Library
paulson
parents:
diff changeset
    39
  one of those primes.
a548865b1b6a moved Primes.thy from NumberTheory to Library
paulson
parents:
diff changeset
    40
*}
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_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
    43
  by (blast intro: relprime_dvd_mult prime_imp_relprime)
11363
a548865b1b6a moved Primes.thy from NumberTheory to Library
paulson
parents:
diff changeset
    44
16663
13e9c402308b prime is a predicate now.
nipkow
parents: 15628
diff changeset
    45
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
    46
  by (auto dest: prime_dvd_mult)
1fce8f51034d prime_dvd_power_two;
wenzelm
parents: 12300
diff changeset
    47
16663
13e9c402308b prime is a predicate now.
nipkow
parents: 15628
diff changeset
    48
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
    49
  by (rule prime_dvd_square) (simp_all add: power2_eq_square)
11368
9c1995c73383 tuned Primes theory;
wenzelm
parents: 11363
diff changeset
    50
11363
a548865b1b6a moved Primes.thy from NumberTheory to Library
paulson
parents:
diff changeset
    51
end