src/ZF/Bool.thy
author nipkow
Mon, 11 Dec 1995 11:24:51 +0100
changeset 1402 b72ccd1cff02
parent 1401 0c439768f45c
child 1478 2b8c2a7547ab
permissions -rw-r--r--
layout
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     1
(*  Title: 	ZF/bool.thy
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     3
    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
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
124
858ab9a9b047 made pseudo theories for all ML files;
clasohm
parents: 14
diff changeset
    11
Bool = ZF + "simpdata" +
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
    "1"		:: i     	   ("1")
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 837
diff changeset
    14
    "2"         :: i             ("2")
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 837
diff changeset
    15
    bool        :: i
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 837
diff changeset
    16
    cond        :: [i,i,i]=>i
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 837
diff changeset
    17
    not		:: i=>i
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 837
diff changeset
    18
    "and"       :: [i,i]=>i      (infixl 70)
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 837
diff changeset
    19
    or		:: [i,i]=>i      (infixl 65)
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 837
diff changeset
    20
    xor		:: [i,i]=>i      (infixl 65)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    21
14
1c0926788772 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 0
diff changeset
    22
translations
1c0926788772 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 0
diff changeset
    23
   "1"  == "succ(0)"
837
778f01546669 Re-indented declarations; declared the number 2
lcp
parents: 799
diff changeset
    24
   "2"  == "succ(1)"
14
1c0926788772 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 0
diff changeset
    25
753
ec86863e87c8 replaced "rules" by "defs"
lcp
parents: 124
diff changeset
    26
defs
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    27
    bool_def	"bool == {0,1}"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    28
    cond_def	"cond(b,c,d) == if(b=1,c,d)"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    29
    not_def	"not(b) == cond(b,0,1)"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    30
    and_def	"a and b == cond(a,b,0)"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    31
    or_def	"a or b == cond(a,1,b)"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    32
    xor_def	"a xor b == cond(a,not(b),b)"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    33
end