src/HOL/Real/Lubs.thy
author paulson
Wed, 14 Jun 2000 17:47:18 +0200
changeset 9065 15f82c9aa331
parent 7562 8519d5019309
child 9279 fb4186e20148
permissions -rw-r--r--
full_rename_numerals -> rename_numerals; tidied
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     1
(*  Title       : Lubs.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 : Upper bounds, lubs definitions
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     6
                  suggested by James Margetson
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     7
*) 
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     8
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     9
7562
wenzelm
parents: 7219
diff changeset
    10
Lubs = Main +
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    11
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    12
consts
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    13
    
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    14
    "*<=" :: ['a set, 'a] => bool     (infixl 70)
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    15
    "<=*" :: ['a, 'a set] => bool     (infixl 70)
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    16
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    17
constdefs
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    18
    leastP      :: ['a =>bool,'a] => bool
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    19
    "leastP P x == (P x & x <=* Collect P)"
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    20
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    21
    isLub       :: ['a set, 'a set, 'a] => bool    
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    22
    "isLub R S x  == leastP (isUb R S) x"
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    23
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    24
    isUb        :: ['a set, 'a set, 'a] => bool     
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    25
    "isUb R S x   == S *<= x & x: R"
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    26
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    27
    ubs         :: ['a set, 'a set] => 'a set
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    28
    "ubs R S      == Collect (isUb R S)"
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    29
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    30
defs
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    31
    setle_def
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    32
    "S *<= x    == (ALL y: S. y <= x)"
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    33
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    34
    setge_def
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    35
    "x <=* S    == (ALL y: S. x <= y)"
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    36
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    37
end                    
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    38
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
    39