author | wenzelm |
Fri, 29 May 1998 13:50:21 +0200 | |
changeset 4987 | 257aeccdefc3 |
parent 4359 | 6f2986464280 |
child 5183 | 89f162de39cf |
permissions | -rw-r--r-- |
1476 | 1 |
(* Title: HOL/Hoare/Arith2.thy |
1335 | 2 |
ID: $Id$ |
1476 | 3 |
Author: Norbert Galm |
1335 | 4 |
Copyright 1995 TUM |
5 |
||
3372
6e472c8f0011
Replacement of "divides" by "dvd" from Divides.thy, and updating of proofs
paulson
parents:
1824
diff
changeset
|
6 |
More arithmetic. Much of this duplicates ex/Primes. |
1335 | 7 |
*) |
8 |
||
4359 | 9 |
Arith2 = Power + |
1335 | 10 |
|
1558 | 11 |
constdefs |
1476 | 12 |
cd :: [nat, nat, nat] => bool |
3372
6e472c8f0011
Replacement of "divides" by "dvd" from Divides.thy, and updating of proofs
paulson
parents:
1824
diff
changeset
|
13 |
"cd x m n == x dvd m & x dvd n" |
1335 | 14 |
|
1558 | 15 |
gcd :: [nat, nat] => nat |
16 |
"gcd m n == @x.(cd x m n) & (!y.(cd y m n) --> y<=x)" |
|
1335 | 17 |
|
4359 | 18 |
consts fac :: nat => nat |
19 |
primrec fac nat |
|
20 |
"fac 0 = Suc 0" |
|
21 |
"fac(Suc n) = (Suc n)*fac(n)" |
|
1335 | 22 |
|
23 |
end |