author | nipkow |
Fri, 04 Sep 1998 11:01:59 +0200 | |
changeset 5427 | 26c9a7c0b36b |
parent 2539 | ddd22ceee8cc |
child 9570 | e16e168984e1 |
permissions | -rw-r--r-- |
1478 | 1 |
(* Title: ZF/bool.thy |
0 | 2 |
ID: $Id$ |
1478 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
0 | 4 |
Copyright 1992 University of Cambridge |
5 |
||
6 |
Booleans in Zermelo-Fraenkel Set Theory |
|
837 | 7 |
|
8 |
2 is equal to bool, but serves a different purpose |
|
0 | 9 |
*) |
10 |
||
2469 | 11 |
Bool = upair + |
0 | 12 |
consts |
1401 | 13 |
bool :: i |
14 |
cond :: [i,i,i]=>i |
|
1478 | 15 |
not :: i=>i |
1401 | 16 |
"and" :: [i,i]=>i (infixl 70) |
1478 | 17 |
or :: [i,i]=>i (infixl 65) |
18 |
xor :: [i,i]=>i (infixl 65) |
|
0 | 19 |
|
2539 | 20 |
syntax |
21 |
"1" :: i ("1") |
|
22 |
"2" :: i ("2") |
|
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 | 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 | 28 |
defs |
1478 | 29 |
bool_def "bool == {0,1}" |
30 |
cond_def "cond(b,c,d) == if(b=1,c,d)" |
|
31 |
not_def "not(b) == cond(b,0,1)" |
|
32 |
and_def "a and b == cond(a,b,0)" |
|
33 |
or_def "a or b == cond(a,1,b)" |
|
34 |
xor_def "a xor b == cond(a,not(b),b)" |
|
0 | 35 |
end |