author | nipkow |
Thu, 10 May 2001 17:28:40 +0200 | |
changeset 11295 | 66925f23ac7f |
parent 10309 | a7f961fb62c6 |
child 11701 | 3d51fbf81c17 |
permissions | -rw-r--r-- |
(* Title : HOL/Real/RealPow.thy ID : $Id$ Author : Jacques D. Fleuriot Copyright : 1998 University of Cambridge Description : Natural powers theory *) theory RealPow = RealAbs: (*belongs to theory RealAbs*) lemmas [arith_split] = abs_split instance real :: power .. primrec (realpow) realpow_0: "r ^ 0 = #1" realpow_Suc: "r ^ (Suc n) = (r::real) * (r ^ n)" end