src/HOL/Real/RComplete.thy
author wenzelm
Tue, 24 Aug 1999 11:50:58 +0200
changeset 7333 6cb15c6f1d9f
parent 7292 dff3470c5c62
child 9429 8ebc549e9326
permissions -rw-r--r--
isar: no_pos flag;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     1
(*  Title       : RComplete.thy
7219
4e3f386c2e37 inserted Id: lines
paulson
parents: 5078
diff changeset
     2
    ID          : $Id$
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     3
    Author      : Jacques D. Fleuriot
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     4
    Copyright   : 1998  University of Cambridge
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     5
    Description : Completeness theorems for positive
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     6
                  reals and reals 
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     7
*) 
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     8
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents: 7219
diff changeset
     9
RComplete = Lubs + RealBin
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    10