src/HOL/Real/Hyperreal/Zorn.thy
author paulson
Thu, 06 Jul 2000 15:01:52 +0200
changeset 9266 1b917b8b1b38
parent 7564 90455fa8cebe
child 9279 fb4186e20148
permissions -rw-r--r--
removal of batch style, and tidying
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5979
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
     1
(*  Title       : Zorn.thy
7219
4e3f386c2e37 inserted Id: lines
paulson
parents: 5979
diff changeset
     2
    ID          : $Id$
5979
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
     3
    Author      : Jacques D. Fleuriot
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
     4
    Copyright   : 1998  University of Cambridge
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
     5
    Description : Zorn's Lemma -- See lcp's Zorn.thy in ZF
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
     6
*) 
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
     7
7564
wenzelm
parents: 7219
diff changeset
     8
Zorn = Main +  
5979
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
     9
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    10
constdefs
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    11
  chain     ::  'a set => 'a set set
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    12
    "chain S  == {F. F <= S & (ALL x: F. ALL y: F. x <= y | y <= x)}" 
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    13
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    14
  super     ::  ['a set,'a set] => (('a set) set) 
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    15
    "super S c == {d. d: chain(S) & c < d}"
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    16
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    17
  maxchain  ::  'a set => 'a set set
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    18
    "maxchain S == {c. c: chain S & super S c = {}}"
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    19
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    20
  succ      ::  ['a set,'a set] => 'a set
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    21
    "succ S c == if (c ~: chain S| c: maxchain S) 
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    22
                 then c else (@c'. c': (super S c))" 
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    23
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    24
consts 
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    25
  "TFin" :: 'a set => 'a set set
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    26
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    27
inductive "TFin(S)"
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    28
  intrs
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    29
    succI        "x : TFin S ==> succ S x : TFin S"
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    30
    Pow_UnionI   "Y : Pow(TFin S) ==> Union(Y) : TFin S"
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    31
           
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    32
  monos          Pow_mono
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    33
end
11cbf236ca16 Addition of Hyperreal theories Zorn and Filter
paulson
parents:
diff changeset
    34