author | kleing |
Tue, 13 May 2003 08:59:21 +0200 | |
changeset 14024 | 213dcc39358f |
parent 12018 | ec054019c910 |
child 14265 | 95b42e69436c |
permissions | -rw-r--r-- |
9435
c3a13a7d4424
lemmas [arith_split] = abs_split (*belongs to theory RealAbs*);
wenzelm
parents:
9013
diff
changeset
|
1 |
(* Title : HOL/Real/RealPow.thy |
7219 | 2 |
ID : $Id$ |
7077
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff
changeset
|
3 |
Author : Jacques D. Fleuriot |
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff
changeset
|
4 |
Copyright : 1998 University of Cambridge |
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff
changeset
|
5 |
Description : Natural powers theory |
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff
changeset
|
6 |
|
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff
changeset
|
7 |
*) |
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff
changeset
|
8 |
|
9435
c3a13a7d4424
lemmas [arith_split] = abs_split (*belongs to theory RealAbs*);
wenzelm
parents:
9013
diff
changeset
|
9 |
theory RealPow = RealAbs: |
7077
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff
changeset
|
10 |
|
9435
c3a13a7d4424
lemmas [arith_split] = abs_split (*belongs to theory RealAbs*);
wenzelm
parents:
9013
diff
changeset
|
11 |
(*belongs to theory RealAbs*) |
c3a13a7d4424
lemmas [arith_split] = abs_split (*belongs to theory RealAbs*);
wenzelm
parents:
9013
diff
changeset
|
12 |
lemmas [arith_split] = abs_split |
c3a13a7d4424
lemmas [arith_split] = abs_split (*belongs to theory RealAbs*);
wenzelm
parents:
9013
diff
changeset
|
13 |
|
c3a13a7d4424
lemmas [arith_split] = abs_split (*belongs to theory RealAbs*);
wenzelm
parents:
9013
diff
changeset
|
14 |
|
10309 | 15 |
instance real :: power .. |
7077
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff
changeset
|
16 |
|
8856 | 17 |
primrec (realpow) |
12018
ec054019c910
Numerals and simprocs for types real and hypreal. The abstract
paulson
parents:
11701
diff
changeset
|
18 |
realpow_0: "r ^ 0 = 1" |
9435
c3a13a7d4424
lemmas [arith_split] = abs_split (*belongs to theory RealAbs*);
wenzelm
parents:
9013
diff
changeset
|
19 |
realpow_Suc: "r ^ (Suc n) = (r::real) * (r ^ n)" |
7077
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff
changeset
|
20 |
|
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff
changeset
|
21 |
end |