| author | berghofe | 
| Mon, 24 Jan 2005 18:18:28 +0100 | |
| changeset 15464 | 02cc838b64ca | 
| parent 8791 | 50b650d19641 | 
| child 17278 | f27456a2a975 | 
| 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: 
1824diff
changeset | 6 | More arithmetic. Much of this duplicates ex/Primes. | 
| 1335 | 7 | *) | 
| 8 | ||
| 8791 | 9 | Arith2 = Main + | 
| 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: 
1824diff
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 | 
| 5183 | 19 | |
| 20 | primrec | |
| 21 | "fac 0 = Suc 0" | |
| 22 | "fac(Suc n) = (Suc n)*fac(n)" | |
| 1335 | 23 | |
| 24 | end |