src/HOL/Hoare/Arith2.thy
author paulson
Wed, 11 Mar 1998 11:03:43 +0100
changeset 4732 10af4886b33f
parent 4359 6f2986464280
child 5183 89f162de39cf
permissions -rw-r--r--
Arith.thy -> thy; proved a few new theorems

(*  Title:      HOL/Hoare/Arith2.thy
    ID:         $Id$
    Author:     Norbert Galm
    Copyright   1995 TUM

More arithmetic.  Much of this duplicates ex/Primes.
*)

Arith2 = Power +

constdefs
  cd      :: [nat, nat, nat] => bool
  "cd x m n  == x dvd m & x dvd n"

  gcd     :: [nat, nat] => nat
  "gcd m n     == @x.(cd x m n) & (!y.(cd y m n) --> y<=x)"

consts fac     :: nat => nat
primrec fac nat
"fac 0 = Suc 0"
"fac(Suc n) = (Suc n)*fac(n)"

end