src/HOL/ex/Primes.thy
author wenzelm
Tue, 27 May 1997 15:45:07 +0200
changeset 3362 0b268cff9344
parent 3242 406ae5ced4e9
child 3376 0cc2eaa8b0f9
permissions -rw-r--r--
NJ 1.09.2x as factory default!
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1804
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
     1
(*  Title:      HOL/ex/Primes.thy
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
     2
    ID:         $Id$
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
     3
    Author:     Christophe Tabacznyj and Lawrence C Paulson
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
     4
    Copyright   1996  University of Cambridge
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
     5
3242
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 1804
diff changeset
     6
The "divides" relation, the Greatest Common Divisor and Euclid's algorithm
1804
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
     7
*)
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
     8
3242
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 1804
diff changeset
     9
Primes = Arith + WF_Rel +
1804
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
    10
consts
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
    11
  dvd     :: [nat,nat]=>bool              (infixl 70) 
3242
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 1804
diff changeset
    12
  is_gcd  :: [nat,nat,nat]=>bool          (*gcd as a relation*)
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 1804
diff changeset
    13
  gcd     :: "nat*nat=>nat"               (*Euclid's algorithm *)
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 1804
diff changeset
    14
  coprime :: [nat,nat]=>bool
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 1804
diff changeset
    15
  prime   :: nat=>bool
1804
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
    16
  
3242
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 1804
diff changeset
    17
recdef gcd "measure ((%(x,y).y) ::nat*nat=>nat)"
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 1804
diff changeset
    18
    "gcd (m, n) = (if n=0 then m else gcd(n, m mod n))"
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 1804
diff changeset
    19
1804
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
    20
defs
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
    21
  dvd_def     "m dvd n == EX k. n = m*k"
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
    22
3242
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 1804
diff changeset
    23
  is_gcd_def  "is_gcd p m n == p dvd m  &  p dvd n  &
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 1804
diff changeset
    24
                               (ALL d. d dvd m & d dvd n --> d dvd p)"
1804
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
    25
3242
406ae5ced4e9 Renamed egcd and gcd; defined the gcd function using TFL
paulson
parents: 1804
diff changeset
    26
  coprime_def "coprime m n == gcd(m,n) = 1"
1804
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
    27
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
    28
  prime_def   "prime(n) == (1<n) & (ALL m. 1<m & m<n --> ~(m dvd n))"
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
    29
cfa0052d5fe9 New example of greatest common divisor
paulson
parents:
diff changeset
    30
end