src/ZF/Bool.thy
author paulson
Tue, 18 Jun 2002 18:45:07 +0200
changeset 13220 62c899c77151
parent 9570 e16e168984e1
child 13239 f599984ec4c2
permissions -rw-r--r--
tidying
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     1
(*  Title:      ZF/bool.thy
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1992  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     6
Booleans in Zermelo-Fraenkel Set Theory 
837
778f01546669 Re-indented declarations; declared the number 2
lcp
parents: 799
diff changeset
     7
778f01546669 Re-indented declarations; declared the number 2
lcp
parents: 799
diff changeset
     8
2 is equal to bool, but serves a different purpose
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     9
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    10
9570
e16e168984e1 installation of cancellation simprocs for the integers
paulson
parents: 2539
diff changeset
    11
Bool = pair + 
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    12
consts
1401
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 837
diff changeset
    13
    bool        :: i
13220
62c899c77151 tidying
paulson
parents: 9570
diff changeset
    14
    cond        :: "[i,i,i]=>i"
62c899c77151 tidying
paulson
parents: 9570
diff changeset
    15
    not         :: "i=>i"
62c899c77151 tidying
paulson
parents: 9570
diff changeset
    16
    "and"       :: "[i,i]=>i"      (infixl 70)
62c899c77151 tidying
paulson
parents: 9570
diff changeset
    17
    or          :: "[i,i]=>i"      (infixl 65)
62c899c77151 tidying
paulson
parents: 9570
diff changeset
    18
    xor         :: "[i,i]=>i"      (infixl 65)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    19
2539
ddd22ceee8cc turned some consts into syntax;
wenzelm
parents: 2469
diff changeset
    20
syntax
ddd22ceee8cc turned some consts into syntax;
wenzelm
parents: 2469
diff changeset
    21
    "1"         :: i             ("1")
ddd22ceee8cc turned some consts into syntax;
wenzelm
parents: 2469
diff changeset
    22
    "2"         :: i             ("2")
ddd22ceee8cc turned some consts into syntax;
wenzelm
parents: 2469
diff changeset
    23
14
1c0926788772 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 0
diff changeset
    24
translations
1c0926788772 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 0
diff changeset
    25
   "1"  == "succ(0)"
837
778f01546669 Re-indented declarations; declared the number 2
lcp
parents: 799
diff changeset
    26
   "2"  == "succ(1)"
14
1c0926788772 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 0
diff changeset
    27
753
ec86863e87c8 replaced "rules" by "defs"
lcp
parents: 124
diff changeset
    28
defs
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    29
    bool_def    "bool == {0,1}"
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    30
    cond_def    "cond(b,c,d) == if(b=1,c,d)"
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    31
    not_def     "not(b) == cond(b,0,1)"
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    32
    and_def     "a and b == cond(a,b,0)"
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    33
    or_def      "a or b == cond(a,1,b)"
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    34
    xor_def     "a xor b == cond(a,not(b),b)"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    35
end