src/ZF/ex/Primes.thy
author paulson
Thu, 13 Jun 1996 16:22:37 +0200
changeset 1793 09fff2f0d727
parent 1792 75c54074cd8c
child 9548 15bee2731e43
permissions -rw-r--r--
New example of GCDs and divides relation
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1792
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
     1
(*  Title:      ZF/ex/Primes.thy
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
     2
    ID:         $Id$
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
     3
    Author:     Christophe Tabacznyj and Lawrence C Paulson
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
     4
    Copyright   1996  University of Cambridge
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
     5
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
     6
The "divides" relation, the greatest common divisor and Euclid's algorithm
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
     7
*)
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
     8
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
     9
Primes = Arith +
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    10
consts
1793
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1792
diff changeset
    11
  dvd     :: [i,i]=>o              (infixl 70) 
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1792
diff changeset
    12
  gcd     :: [i,i,i]=>o            (* great common divisor *)
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1792
diff changeset
    13
  egcd    :: [i,i]=>i              (* gcd by Euclid's algorithm *)
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1792
diff changeset
    14
  coprime :: [i,i]=>o              (* coprime definition *)
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1792
diff changeset
    15
  prime   :: i=>o                  (* prime definition *)
1792
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    16
  
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    17
defs
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    18
  dvd_def     "m dvd n == m:nat & n:nat & (EX k:nat. n = m#*k)"
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    19
1793
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1792
diff changeset
    20
  gcd_def     "gcd(p,m,n) == ((p dvd m) & (p dvd n))   &
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1792
diff changeset
    21
               (ALL d. (d dvd m) & (d dvd n) --> d dvd p)"
1792
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    22
1793
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1792
diff changeset
    23
  egcd_def    "egcd(m,n) ==   
09fff2f0d727 New example of GCDs and divides relation
paulson
parents: 1792
diff changeset
    24
               transrec(n, %n f. lam m:nat. if(n=0, m, f`(m mod n)`n)) ` m"
1792
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    25
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    26
  coprime_def "coprime(m,n) == egcd(m,n) = 1"
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    27
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    28
  prime_def   "prime(n) == (1<n) & (ALL m:nat. 1<m & m<n --> ~(m dvd n))"
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    29
75c54074cd8c The "divides" relation, the greatest common divisor and Euclid's algorithm
paulson
parents:
diff changeset
    30
end