src/HOL/Real/RealPow.thy
author wenzelm
Tue, 24 Aug 1999 11:50:58 +0200
changeset 7333 6cb15c6f1d9f
parent 7219 4e3f386c2e37
child 8856 435187ffc64e
permissions -rw-r--r--
isar: no_pos flag;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
     1
(*  Title       : RealPow.thy
7219
4e3f386c2e37 inserted Id: lines
paulson
parents: 7077
diff changeset
     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
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
     9
RealPow = WF_Rel + RealAbs + 
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    10
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    11
instance real :: {power}
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    12
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    13
primrec
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    14
     realpow_0   "r ^ 0       = 1r"
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    15
     realpow_Suc "r ^ (Suc n) = (r::real) * (r ^ n)"
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    16
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    17
end