src/HOL/Real/RealPow.thy
author wenzelm
Tue, 23 Oct 2001 22:52:31 +0200
changeset 11908 82f68fd05094
parent 11701 3d51fbf81c17
child 12018 ec054019c910
permissions -rw-r--r--
eliminated old numerals;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
9435
c3a13a7d4424 lemmas [arith_split] = abs_split (*belongs to theory RealAbs*);
wenzelm
parents: 9013
diff changeset
     1
(*  Title       : HOL/Real/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
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
a7f961fb62c6 intro_classes by default;
wenzelm
parents: 9435
diff changeset
    15
instance real :: power ..
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    16
8856
435187ffc64e fixed theory deps;
wenzelm
parents: 7219
diff changeset
    17
primrec (realpow)
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 10309
diff changeset
    18
     realpow_0:   "r ^ 0       = Numeral1"
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