src/ZF/Cardinal.thy
author lcp
Thu, 25 Aug 1994 12:09:21 +0200
changeset 578 efc648d29dd0
parent 435 ca5356bd315a
child 753 ec86863e87c8
permissions -rw-r--r--
ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without the keyword "inductive" making the theory file fail ZF/Makefile: now has Inductive.thy,.ML ZF/Datatype,Finite,Zorn: depend upon Inductive ZF/intr_elim: now checks that the inductive name does not clash with existing theory names ZF/ind_section: deleted things replicated in Pure/section_utils.ML ZF/ROOT: now loads Pure/section_utils

(*  Title: 	ZF/Cardinal.thy
    ID:         $Id$
    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
    Copyright   1994  University of Cambridge

Cardinals in Zermelo-Fraenkel Set Theory 
*)

Cardinal = OrderType + Fixedpt + Nat + Sum + 
consts
  Least            :: "(i=>o) => i"    (binder "LEAST " 10)
  eqpoll, lepoll   :: "[i,i] => o"     (infixl 50)
  "cardinal"       :: "i=>i"           ("|_|")
  Card             :: "i=>o"

  swap       :: "[i,i,i]=>i"     (*not used; useful??*)

rules

  (*least ordinal operator*)
  Least_def  "Least(P) == THE i. Ord(i) & P(i) & (ALL j. j<i --> ~P(j))"

  eqpoll_def "A eqpoll B == EX f. f: bij(A,B)"

  lepoll_def "A lepoll B == EX f. f: inj(A,B)"

  cardinal_def "|A| == LEAST i. i eqpoll A"

  Card_def     "Card(i) == ( i = |i| )"

  swap_def   "swap(A,x,y) == (lam z:A. if(z=x, y, if(z=y, x, z)))"

end