author | paulson |
Fri, 11 Sep 1998 18:09:54 +0200 | |
changeset 5484 | e9430ed7e8d6 |
parent 5268 | 59ef39008514 |
child 6053 | 8a1059aa01f0 |
permissions | -rw-r--r-- |
1461 | 1 |
(* Title: ZF/bool |
0 | 2 |
ID: $Id$ |
1461 | 3 |
Author: Martin D Coen, Cambridge University Computer Laboratory |
0 | 4 |
Copyright 1992 University of Cambridge |
5 |
||
5268 | 6 |
Booleans in Zermelo-Fraenkel Set Theory |
0 | 7 |
*) |
8 |
||
9 |
open Bool; |
|
10 |
||
14
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
6
diff
changeset
|
11 |
val bool_defs = [bool_def,cond_def]; |
0 | 12 |
|
5067 | 13 |
Goalw [succ_def] "{0} = 1"; |
2493 | 14 |
by (rtac refl 1); |
2469 | 15 |
qed "singleton_0"; |
16 |
||
0 | 17 |
(* Introduction rules *) |
18 |
||
5067 | 19 |
Goalw bool_defs "1 : bool"; |
0 | 20 |
by (rtac (consI1 RS consI2) 1); |
760 | 21 |
qed "bool_1I"; |
0 | 22 |
|
5067 | 23 |
Goalw bool_defs "0 : bool"; |
0 | 24 |
by (rtac consI1 1); |
760 | 25 |
qed "bool_0I"; |
0 | 26 |
|
5067 | 27 |
Goalw bool_defs "1~=0"; |
0 | 28 |
by (rtac succ_not_0 1); |
760 | 29 |
qed "one_not_0"; |
0 | 30 |
|
31 |
(** 1=0 ==> R **) |
|
32 |
val one_neq_0 = one_not_0 RS notE; |
|
33 |
||
5268 | 34 |
val major::prems = Goalw bool_defs |
119
0e58da397b1d
boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents:
37
diff
changeset
|
35 |
"[| c: bool; c=1 ==> P; c=0 ==> P |] ==> P"; |
129 | 36 |
by (rtac (major RS consE) 1); |
119
0e58da397b1d
boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents:
37
diff
changeset
|
37 |
by (REPEAT (eresolve_tac (singletonE::prems) 1)); |
760 | 38 |
qed "boolE"; |
0 | 39 |
|
40 |
(** cond **) |
|
41 |
||
42 |
(*1 means true*) |
|
5067 | 43 |
Goalw bool_defs "cond(1,c,d) = c"; |
0 | 44 |
by (rtac (refl RS if_P) 1); |
760 | 45 |
qed "cond_1"; |
0 | 46 |
|
47 |
(*0 means false*) |
|
5067 | 48 |
Goalw bool_defs "cond(0,c,d) = d"; |
0 | 49 |
by (rtac (succ_not_0 RS not_sym RS if_not_P) 1); |
760 | 50 |
qed "cond_0"; |
0 | 51 |
|
2469 | 52 |
Addsimps [cond_1, cond_0]; |
53 |
||
4091 | 54 |
fun bool_tac i = fast_tac (claset() addSEs [boolE] addss (simpset())) i; |
2469 | 55 |
|
56 |
||
5268 | 57 |
Goal "[| b: bool; c: A(1); d: A(0) |] ==> cond(b,c,d): A(b)"; |
2469 | 58 |
by (bool_tac 1); |
760 | 59 |
qed "cond_type"; |
0 | 60 |
|
5268 | 61 |
val [rew] = Goal "[| !!b. j(b)==cond(b,c,d) |] ==> j(1) = c"; |
0 | 62 |
by (rewtac rew); |
63 |
by (rtac cond_1 1); |
|
760 | 64 |
qed "def_cond_1"; |
0 | 65 |
|
5268 | 66 |
val [rew] = Goal "[| !!b. j(b)==cond(b,c,d) |] ==> j(0) = d"; |
0 | 67 |
by (rewtac rew); |
68 |
by (rtac cond_0 1); |
|
760 | 69 |
qed "def_cond_0"; |
0 | 70 |
|
71 |
fun conds def = [standard (def RS def_cond_1), standard (def RS def_cond_0)]; |
|
72 |
||
73 |
val [not_1,not_0] = conds not_def; |
|
74 |
val [and_1,and_0] = conds and_def; |
|
2469 | 75 |
val [or_1,or_0] = conds or_def; |
76 |
val [xor_1,xor_0] = conds xor_def; |
|
0 | 77 |
|
2469 | 78 |
Addsimps [not_1,not_0,and_1,and_0,or_1,or_0,xor_1,xor_0]; |
0 | 79 |
|
760 | 80 |
qed_goalw "not_type" Bool.thy [not_def] |
0 | 81 |
"a:bool ==> not(a) : bool" |
82 |
(fn prems=> [ (typechk_tac (prems@[bool_1I, bool_0I, cond_type])) ]); |
|
83 |
||
760 | 84 |
qed_goalw "and_type" Bool.thy [and_def] |
0 | 85 |
"[| a:bool; b:bool |] ==> a and b : bool" |
86 |
(fn prems=> [ (typechk_tac (prems@[bool_1I, bool_0I, cond_type])) ]); |
|
87 |
||
760 | 88 |
qed_goalw "or_type" Bool.thy [or_def] |
0 | 89 |
"[| a:bool; b:bool |] ==> a or b : bool" |
90 |
(fn prems=> [ (typechk_tac (prems@[bool_1I, bool_0I, cond_type])) ]); |
|
91 |
||
760 | 92 |
qed_goalw "xor_type" Bool.thy [xor_def] |
0 | 93 |
"[| a:bool; b:bool |] ==> a xor b : bool" |
94 |
(fn prems=> [ (typechk_tac(prems@[bool_1I, bool_0I, cond_type, not_type])) ]); |
|
95 |
||
96 |
val bool_typechecks = [bool_1I, bool_0I, cond_type, not_type, and_type, |
|
2469 | 97 |
or_type, xor_type]; |
119
0e58da397b1d
boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents:
37
diff
changeset
|
98 |
|
0e58da397b1d
boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents:
37
diff
changeset
|
99 |
(*** Laws for 'not' ***) |
0e58da397b1d
boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents:
37
diff
changeset
|
100 |
|
5137 | 101 |
Goal "a:bool ==> not(not(a)) = a"; |
2469 | 102 |
by (bool_tac 1); |
760 | 103 |
qed "not_not"; |
119
0e58da397b1d
boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents:
37
diff
changeset
|
104 |
|
5137 | 105 |
Goal "a:bool ==> not(a and b) = not(a) or not(b)"; |
2469 | 106 |
by (bool_tac 1); |
760 | 107 |
qed "not_and"; |
119
0e58da397b1d
boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents:
37
diff
changeset
|
108 |
|
5137 | 109 |
Goal "a:bool ==> not(a or b) = not(a) and not(b)"; |
2469 | 110 |
by (bool_tac 1); |
760 | 111 |
qed "not_or"; |
119
0e58da397b1d
boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents:
37
diff
changeset
|
112 |
|
2469 | 113 |
Addsimps [not_not, not_and, not_or]; |
114 |
||
119
0e58da397b1d
boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents:
37
diff
changeset
|
115 |
(*** Laws about 'and' ***) |
0e58da397b1d
boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents:
37
diff
changeset
|
116 |
|
5137 | 117 |
Goal "a: bool ==> a and a = a"; |
2469 | 118 |
by (bool_tac 1); |
760 | 119 |
qed "and_absorb"; |
119
0e58da397b1d
boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents:
37
diff
changeset
|
120 |
|
2469 | 121 |
Addsimps [and_absorb]; |
122 |
||
5137 | 123 |
Goal "[| a: bool; b:bool |] ==> a and b = b and a"; |
2469 | 124 |
by (bool_tac 1); |
760 | 125 |
qed "and_commute"; |
0 | 126 |
|
5137 | 127 |
Goal "a: bool ==> (a and b) and c = a and (b and c)"; |
2469 | 128 |
by (bool_tac 1); |
760 | 129 |
qed "and_assoc"; |
119
0e58da397b1d
boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents:
37
diff
changeset
|
130 |
|
5268 | 131 |
Goal "[| a: bool; b:bool; c:bool |] ==> \ |
119
0e58da397b1d
boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents:
37
diff
changeset
|
132 |
\ (a or b) and c = (a and c) or (b and c)"; |
2469 | 133 |
by (bool_tac 1); |
760 | 134 |
qed "and_or_distrib"; |
119
0e58da397b1d
boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents:
37
diff
changeset
|
135 |
|
0e58da397b1d
boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents:
37
diff
changeset
|
136 |
(** binary orion **) |
0e58da397b1d
boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents:
37
diff
changeset
|
137 |
|
5137 | 138 |
Goal "a: bool ==> a or a = a"; |
2469 | 139 |
by (bool_tac 1); |
760 | 140 |
qed "or_absorb"; |
119
0e58da397b1d
boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents:
37
diff
changeset
|
141 |
|
2469 | 142 |
Addsimps [or_absorb]; |
143 |
||
5137 | 144 |
Goal "[| a: bool; b:bool |] ==> a or b = b or a"; |
2469 | 145 |
by (bool_tac 1); |
760 | 146 |
qed "or_commute"; |
119
0e58da397b1d
boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents:
37
diff
changeset
|
147 |
|
5137 | 148 |
Goal "a: bool ==> (a or b) or c = a or (b or c)"; |
2469 | 149 |
by (bool_tac 1); |
760 | 150 |
qed "or_assoc"; |
119
0e58da397b1d
boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents:
37
diff
changeset
|
151 |
|
5268 | 152 |
Goal "[| a: bool; b: bool; c: bool |] ==> \ |
119
0e58da397b1d
boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents:
37
diff
changeset
|
153 |
\ (a and b) or c = (a or c) and (b or c)"; |
2469 | 154 |
by (bool_tac 1); |
760 | 155 |
qed "or_and_distrib"; |
119
0e58da397b1d
boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents:
37
diff
changeset
|
156 |