| author | wenzelm | 
| Thu, 31 Aug 2000 01:42:23 +0200 | |
| changeset 9762 | 66f7eefb3967 | 
| parent 9435 | c3a13a7d4424 | 
| child 10309 | a7f961fb62c6 | 
| permissions | -rw-r--r-- | 
| 9435 
c3a13a7d4424
lemmas [arith_split] = abs_split (*belongs to theory RealAbs*);
 wenzelm parents: 
9013diff
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: 
9013diff
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: 
9013diff
changeset | 11 | (*belongs to theory RealAbs*) | 
| 
c3a13a7d4424
lemmas [arith_split] = abs_split (*belongs to theory RealAbs*);
 wenzelm parents: 
9013diff
changeset | 12 | lemmas [arith_split] = abs_split | 
| 
c3a13a7d4424
lemmas [arith_split] = abs_split (*belongs to theory RealAbs*);
 wenzelm parents: 
9013diff
changeset | 13 | |
| 
c3a13a7d4424
lemmas [arith_split] = abs_split (*belongs to theory RealAbs*);
 wenzelm parents: 
9013diff
changeset | 14 | |
| 
c3a13a7d4424
lemmas [arith_split] = abs_split (*belongs to theory RealAbs*);
 wenzelm parents: 
9013diff
changeset | 15 | instance real :: power | 
| 
c3a13a7d4424
lemmas [arith_split] = abs_split (*belongs to theory RealAbs*);
 wenzelm parents: 
9013diff
changeset | 16 | by intro_classes | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 17 | |
| 8856 | 18 | primrec (realpow) | 
| 9435 
c3a13a7d4424
lemmas [arith_split] = abs_split (*belongs to theory RealAbs*);
 wenzelm parents: 
9013diff
changeset | 19 | realpow_0: "r ^ 0 = #1" | 
| 
c3a13a7d4424
lemmas [arith_split] = abs_split (*belongs to theory RealAbs*);
 wenzelm parents: 
9013diff
changeset | 20 | realpow_Suc: "r ^ (Suc n) = (r::real) * (r ^ n)" | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 21 | |
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 22 | end |