author | wenzelm |
Fri, 01 Dec 2000 19:39:54 +0100 | |
changeset 10562 | fcd29e58c40c |
parent 9907 | 473a6604da94 |
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 |
||
9907 | 9 |
bind_thms ("bool_defs", [bool_def,cond_def]); |
0 | 10 |
|
5067 | 11 |
Goalw [succ_def] "{0} = 1"; |
2493 | 12 |
by (rtac refl 1); |
2469 | 13 |
qed "singleton_0"; |
14 |
||
0 | 15 |
(* Introduction rules *) |
16 |
||
5067 | 17 |
Goalw bool_defs "1 : bool"; |
0 | 18 |
by (rtac (consI1 RS consI2) 1); |
760 | 19 |
qed "bool_1I"; |
0 | 20 |
|
5067 | 21 |
Goalw bool_defs "0 : bool"; |
0 | 22 |
by (rtac consI1 1); |
760 | 23 |
qed "bool_0I"; |
0 | 24 |
|
6053
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
paulson
parents:
5268
diff
changeset
|
25 |
Addsimps [bool_1I, bool_0I]; |
6153 | 26 |
AddTCs [bool_1I, bool_0I]; |
6053
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
paulson
parents:
5268
diff
changeset
|
27 |
|
5067 | 28 |
Goalw bool_defs "1~=0"; |
0 | 29 |
by (rtac succ_not_0 1); |
760 | 30 |
qed "one_not_0"; |
0 | 31 |
|
32 |
(** 1=0 ==> R **) |
|
9907 | 33 |
bind_thm ("one_neq_0", one_not_0 RS notE); |
0 | 34 |
|
5268 | 35 |
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
|
36 |
"[| c: bool; c=1 ==> P; c=0 ==> P |] ==> P"; |
129 | 37 |
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
|
38 |
by (REPEAT (eresolve_tac (singletonE::prems) 1)); |
760 | 39 |
qed "boolE"; |
0 | 40 |
|
41 |
(** cond **) |
|
42 |
||
43 |
(*1 means true*) |
|
5067 | 44 |
Goalw bool_defs "cond(1,c,d) = c"; |
0 | 45 |
by (rtac (refl RS if_P) 1); |
760 | 46 |
qed "cond_1"; |
0 | 47 |
|
48 |
(*0 means false*) |
|
5067 | 49 |
Goalw bool_defs "cond(0,c,d) = d"; |
0 | 50 |
by (rtac (succ_not_0 RS not_sym RS if_not_P) 1); |
760 | 51 |
qed "cond_0"; |
0 | 52 |
|
2469 | 53 |
Addsimps [cond_1, cond_0]; |
54 |
||
4091 | 55 |
fun bool_tac i = fast_tac (claset() addSEs [boolE] addss (simpset())) i; |
2469 | 56 |
|
57 |
||
6053
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
paulson
parents:
5268
diff
changeset
|
58 |
Goal "[| b: bool; c: A(1); d: A(0) |] ==> cond(b,c,d): A(b)"; |
2469 | 59 |
by (bool_tac 1); |
760 | 60 |
qed "cond_type"; |
6153 | 61 |
AddTCs [cond_type]; |
0 | 62 |
|
6053
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
paulson
parents:
5268
diff
changeset
|
63 |
(*For Simp_tac and Blast_tac*) |
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
paulson
parents:
5268
diff
changeset
|
64 |
Goal "[| b: bool; c: A; d: A |] ==> cond(b,c,d): A"; |
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
paulson
parents:
5268
diff
changeset
|
65 |
by (bool_tac 1); |
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
paulson
parents:
5268
diff
changeset
|
66 |
qed "cond_simple_type"; |
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
paulson
parents:
5268
diff
changeset
|
67 |
|
5268 | 68 |
val [rew] = Goal "[| !!b. j(b)==cond(b,c,d) |] ==> j(1) = c"; |
0 | 69 |
by (rewtac rew); |
70 |
by (rtac cond_1 1); |
|
760 | 71 |
qed "def_cond_1"; |
0 | 72 |
|
5268 | 73 |
val [rew] = Goal "[| !!b. j(b)==cond(b,c,d) |] ==> j(0) = d"; |
0 | 74 |
by (rewtac rew); |
75 |
by (rtac cond_0 1); |
|
760 | 76 |
qed "def_cond_0"; |
0 | 77 |
|
78 |
fun conds def = [standard (def RS def_cond_1), standard (def RS def_cond_0)]; |
|
79 |
||
9907 | 80 |
val [not_1, not_0] = conds not_def; |
81 |
val [and_1, and_0] = conds and_def; |
|
82 |
val [or_1, or_0] = conds or_def; |
|
83 |
val [xor_1, xor_0] = conds xor_def; |
|
84 |
||
85 |
bind_thm ("not_1", not_1); |
|
86 |
bind_thm ("not_0", not_0); |
|
87 |
bind_thm ("and_1", and_1); |
|
88 |
bind_thm ("and_0", and_0); |
|
89 |
bind_thm ("or_1", or_1); |
|
90 |
bind_thm ("or_0", or_0); |
|
91 |
bind_thm ("xor_1", xor_1); |
|
92 |
bind_thm ("xor_0", xor_0); |
|
0 | 93 |
|
2469 | 94 |
Addsimps [not_1,not_0,and_1,and_0,or_1,or_0,xor_1,xor_0]; |
0 | 95 |
|
6153 | 96 |
Goalw [not_def] "a:bool ==> not(a) : bool"; |
97 |
by (Asm_simp_tac 1); |
|
98 |
qed "not_type"; |
|
0 | 99 |
|
6153 | 100 |
Goalw [and_def] "[| a:bool; b:bool |] ==> a and b : bool"; |
101 |
by (Asm_simp_tac 1); |
|
102 |
qed "and_type"; |
|
0 | 103 |
|
6153 | 104 |
Goalw [or_def] "[| a:bool; b:bool |] ==> a or b : bool"; |
105 |
by (Asm_simp_tac 1); |
|
106 |
qed "or_type"; |
|
0 | 107 |
|
6153 | 108 |
AddTCs [not_type, and_type, or_type]; |
6053
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
paulson
parents:
5268
diff
changeset
|
109 |
|
6153 | 110 |
Goalw [xor_def] "[| a:bool; b:bool |] ==> a xor b : bool"; |
111 |
by (Asm_simp_tac 1); |
|
112 |
qed "xor_type"; |
|
0 | 113 |
|
6153 | 114 |
AddTCs [xor_type]; |
6053
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
paulson
parents:
5268
diff
changeset
|
115 |
|
9907 | 116 |
bind_thms ("bool_typechecks", |
117 |
[bool_1I, bool_0I, cond_type, not_type, and_type, 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
|
118 |
|
0e58da397b1d
boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents:
37
diff
changeset
|
119 |
(*** Laws for 'not' ***) |
0e58da397b1d
boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents:
37
diff
changeset
|
120 |
|
5137 | 121 |
Goal "a:bool ==> not(not(a)) = a"; |
2469 | 122 |
by (bool_tac 1); |
760 | 123 |
qed "not_not"; |
119
0e58da397b1d
boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents:
37
diff
changeset
|
124 |
|
5137 | 125 |
Goal "a:bool ==> not(a and b) = not(a) or not(b)"; |
2469 | 126 |
by (bool_tac 1); |
760 | 127 |
qed "not_and"; |
119
0e58da397b1d
boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents:
37
diff
changeset
|
128 |
|
5137 | 129 |
Goal "a:bool ==> not(a or b) = not(a) and not(b)"; |
2469 | 130 |
by (bool_tac 1); |
760 | 131 |
qed "not_or"; |
119
0e58da397b1d
boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents:
37
diff
changeset
|
132 |
|
2469 | 133 |
Addsimps [not_not, not_and, not_or]; |
134 |
||
119
0e58da397b1d
boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents:
37
diff
changeset
|
135 |
(*** Laws about 'and' ***) |
0e58da397b1d
boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents:
37
diff
changeset
|
136 |
|
5137 | 137 |
Goal "a: bool ==> a and a = a"; |
2469 | 138 |
by (bool_tac 1); |
760 | 139 |
qed "and_absorb"; |
119
0e58da397b1d
boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents:
37
diff
changeset
|
140 |
|
2469 | 141 |
Addsimps [and_absorb]; |
142 |
||
5137 | 143 |
Goal "[| a: bool; b:bool |] ==> a and b = b and a"; |
2469 | 144 |
by (bool_tac 1); |
760 | 145 |
qed "and_commute"; |
0 | 146 |
|
5137 | 147 |
Goal "a: bool ==> (a and b) and c = a and (b and c)"; |
2469 | 148 |
by (bool_tac 1); |
760 | 149 |
qed "and_assoc"; |
119
0e58da397b1d
boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents:
37
diff
changeset
|
150 |
|
5268 | 151 |
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
|
152 |
\ (a or b) and c = (a and c) or (b and c)"; |
2469 | 153 |
by (bool_tac 1); |
760 | 154 |
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
|
155 |
|
0e58da397b1d
boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents:
37
diff
changeset
|
156 |
(** binary orion **) |
0e58da397b1d
boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents:
37
diff
changeset
|
157 |
|
5137 | 158 |
Goal "a: bool ==> a or a = a"; |
2469 | 159 |
by (bool_tac 1); |
760 | 160 |
qed "or_absorb"; |
119
0e58da397b1d
boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents:
37
diff
changeset
|
161 |
|
2469 | 162 |
Addsimps [or_absorb]; |
163 |
||
5137 | 164 |
Goal "[| a: bool; b:bool |] ==> a or b = b or a"; |
2469 | 165 |
by (bool_tac 1); |
760 | 166 |
qed "or_commute"; |
119
0e58da397b1d
boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents:
37
diff
changeset
|
167 |
|
5137 | 168 |
Goal "a: bool ==> (a or b) or c = a or (b or c)"; |
2469 | 169 |
by (bool_tac 1); |
760 | 170 |
qed "or_assoc"; |
119
0e58da397b1d
boolE: changed to have equality assumptions instead of P(c); proved many boolean laws
lcp
parents:
37
diff
changeset
|
171 |
|
5268 | 172 |
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
|
173 |
\ (a and b) or c = (a or c) and (b or c)"; |
2469 | 174 |
by (bool_tac 1); |
760 | 175 |
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
|
176 |