author | wenzelm |
Fri, 08 Mar 2002 16:24:06 +0100 | |
changeset 13049 | ce180e5b7fa0 |
parent 11868 | 56db9f3a6b3e |
child 14259 | 79f7d3451b1e |
permissions | -rw-r--r-- |
9509 | 1 |
(* Title: IntPower.thy |
2 |
ID: $Id$ |
|
3 |
Author: Thomas M. Rasmussen |
|
4 |
Copyright 2000 University of Cambridge |
|
5 |
||
6 |
Integer powers |
|
7 |
*) |
|
8 |
||
9 |
IntPower = IntDiv + |
|
10 |
||
11 |
instance |
|
12 |
int :: {power} |
|
13 |
||
14 |
primrec |
|
11868
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
paulson
parents:
11701
diff
changeset
|
15 |
power_0 "p ^ 0 = 1" |
9509 | 16 |
power_Suc "p ^ (Suc n) = (p::int) * (p ^ n)" |
17 |
||
18 |
end |