| author | wenzelm | 
| Sat, 15 Dec 2007 14:52:55 +0100 | |
| changeset 25644 | d30391cdd9d9 | 
| parent 24893 | b8ef7afe3a6b | 
| child 35762 | af3ff2ba4c54 | 
| 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 | ||
| 13328 | 6 | *) | 
| 837 | 7 | |
| 13328 | 8 | header{*Booleans in Zermelo-Fraenkel Set Theory*}
 | 
| 0 | 9 | |
| 16417 | 10 | theory Bool imports pair begin | 
| 0 | 11 | |
| 24892 | 12 | abbreviation | 
| 13 |   one  ("1") where
 | |
| 14 | "1 == succ(0)" | |
| 2539 | 15 | |
| 24892 | 16 | abbreviation | 
| 17 |   two  ("2") where
 | |
| 18 | "2 == succ(1)" | |
| 14 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
0diff
changeset | 19 | |
| 13328 | 20 | text{*2 is equal to bool, but is used as a number rather than a type.*}
 | 
| 21 | ||
| 24893 | 22 | definition "bool == {0,1}"
 | 
| 23 | ||
| 24 | definition "cond(b,c,d) == if(b=1,c,d)" | |
| 13239 | 25 | |
| 24893 | 26 | definition "not(b) == cond(b,0,1)" | 
| 13239 | 27 | |
| 24893 | 28 | definition | 
| 29 | "and" :: "[i,i]=>i" (infixl "and" 70) where | |
| 13239 | 30 | "a and b == cond(a,b,0)" | 
| 31 | ||
| 24893 | 32 | definition | 
| 33 | or :: "[i,i]=>i" (infixl "or" 65) where | |
| 13239 | 34 | "a or b == cond(a,1,b)" | 
| 35 | ||
| 24893 | 36 | definition | 
| 37 | xor :: "[i,i]=>i" (infixl "xor" 65) where | |
| 13239 | 38 | "a xor b == cond(a,not(b),b)" | 
| 39 | ||
| 40 | ||
| 41 | lemmas bool_defs = bool_def cond_def | |
| 42 | ||
| 43 | lemma singleton_0: "{0} = 1"
 | |
| 44 | by (simp add: succ_def) | |
| 45 | ||
| 46 | (* Introduction rules *) | |
| 47 | ||
| 48 | lemma bool_1I [simp,TC]: "1 : bool" | |
| 49 | by (simp add: bool_defs ) | |
| 50 | ||
| 51 | lemma bool_0I [simp,TC]: "0 : bool" | |
| 52 | by (simp add: bool_defs) | |
| 53 | ||
| 54 | lemma one_not_0: "1~=0" | |
| 55 | by (simp add: bool_defs ) | |
| 56 | ||
| 57 | (** 1=0 ==> R **) | |
| 58 | lemmas one_neq_0 = one_not_0 [THEN notE, standard] | |
| 59 | ||
| 60 | lemma boolE: | |
| 61 | "[| c: bool; c=1 ==> P; c=0 ==> P |] ==> P" | |
| 24892 | 62 | by (simp add: bool_defs, blast) | 
| 13239 | 63 | |
| 64 | (** cond **) | |
| 65 | ||
| 66 | (*1 means true*) | |
| 67 | lemma cond_1 [simp]: "cond(1,c,d) = c" | |
| 68 | by (simp add: bool_defs ) | |
| 69 | ||
| 70 | (*0 means false*) | |
| 71 | lemma cond_0 [simp]: "cond(0,c,d) = d" | |
| 72 | by (simp add: bool_defs ) | |
| 73 | ||
| 74 | lemma cond_type [TC]: "[| b: bool; c: A(1); d: A(0) |] ==> cond(b,c,d): A(b)" | |
| 13269 | 75 | by (simp add: bool_defs, blast) | 
| 13239 | 76 | |
| 77 | (*For Simp_tac and Blast_tac*) | |
| 78 | lemma cond_simple_type: "[| b: bool; c: A; d: A |] ==> cond(b,c,d): A" | |
| 79 | by (simp add: bool_defs ) | |
| 80 | ||
| 81 | lemma def_cond_1: "[| !!b. j(b)==cond(b,c,d) |] ==> j(1) = c" | |
| 82 | by simp | |
| 83 | ||
| 84 | lemma def_cond_0: "[| !!b. j(b)==cond(b,c,d) |] ==> j(0) = d" | |
| 85 | by simp | |
| 86 | ||
| 87 | lemmas not_1 = not_def [THEN def_cond_1, standard, simp] | |
| 88 | lemmas not_0 = not_def [THEN def_cond_0, standard, simp] | |
| 89 | ||
| 90 | lemmas and_1 = and_def [THEN def_cond_1, standard, simp] | |
| 91 | lemmas and_0 = and_def [THEN def_cond_0, standard, simp] | |
| 92 | ||
| 93 | lemmas or_1 = or_def [THEN def_cond_1, standard, simp] | |
| 94 | lemmas or_0 = or_def [THEN def_cond_0, standard, simp] | |
| 95 | ||
| 96 | lemmas xor_1 = xor_def [THEN def_cond_1, standard, simp] | |
| 97 | lemmas xor_0 = xor_def [THEN def_cond_0, standard, simp] | |
| 98 | ||
| 99 | lemma not_type [TC]: "a:bool ==> not(a) : bool" | |
| 100 | by (simp add: not_def) | |
| 101 | ||
| 102 | lemma and_type [TC]: "[| a:bool; b:bool |] ==> a and b : bool" | |
| 103 | by (simp add: and_def) | |
| 104 | ||
| 105 | lemma or_type [TC]: "[| a:bool; b:bool |] ==> a or b : bool" | |
| 106 | by (simp add: or_def) | |
| 107 | ||
| 108 | lemma xor_type [TC]: "[| a:bool; b:bool |] ==> a xor b : bool" | |
| 109 | by (simp add: xor_def) | |
| 110 | ||
| 111 | lemmas bool_typechecks = bool_1I bool_0I cond_type not_type and_type | |
| 112 | or_type xor_type | |
| 113 | ||
| 13356 | 114 | subsection{*Laws About 'not' *}
 | 
| 13239 | 115 | |
| 116 | lemma not_not [simp]: "a:bool ==> not(not(a)) = a" | |
| 117 | by (elim boolE, auto) | |
| 118 | ||
| 119 | lemma not_and [simp]: "a:bool ==> not(a and b) = not(a) or not(b)" | |
| 120 | by (elim boolE, auto) | |
| 121 | ||
| 122 | lemma not_or [simp]: "a:bool ==> not(a or b) = not(a) and not(b)" | |
| 123 | by (elim boolE, auto) | |
| 124 | ||
| 13356 | 125 | subsection{*Laws About 'and' *}
 | 
| 13239 | 126 | |
| 127 | lemma and_absorb [simp]: "a: bool ==> a and a = a" | |
| 128 | by (elim boolE, auto) | |
| 129 | ||
| 130 | lemma and_commute: "[| a: bool; b:bool |] ==> a and b = b and a" | |
| 131 | by (elim boolE, auto) | |
| 132 | ||
| 133 | lemma and_assoc: "a: bool ==> (a and b) and c = a and (b and c)" | |
| 134 | by (elim boolE, auto) | |
| 135 | ||
| 24892 | 136 | lemma and_or_distrib: "[| a: bool; b:bool; c:bool |] ==> | 
| 13239 | 137 | (a or b) and c = (a and c) or (b and c)" | 
| 138 | by (elim boolE, auto) | |
| 139 | ||
| 13356 | 140 | subsection{*Laws About 'or' *}
 | 
| 13239 | 141 | |
| 142 | lemma or_absorb [simp]: "a: bool ==> a or a = a" | |
| 143 | by (elim boolE, auto) | |
| 144 | ||
| 145 | lemma or_commute: "[| a: bool; b:bool |] ==> a or b = b or a" | |
| 146 | by (elim boolE, auto) | |
| 147 | ||
| 148 | lemma or_assoc: "a: bool ==> (a or b) or c = a or (b or c)" | |
| 149 | by (elim boolE, auto) | |
| 150 | ||
| 24892 | 151 | lemma or_and_distrib: "[| a: bool; b: bool; c: bool |] ==> | 
| 13239 | 152 | (a and b) or c = (a or c) and (b or c)" | 
| 153 | by (elim boolE, auto) | |
| 154 | ||
| 13269 | 155 | |
| 24893 | 156 | definition | 
| 157 | bool_of_o :: "o=>i" where | |
| 13269 | 158 | "bool_of_o(P) == (if P then 1 else 0)" | 
| 159 | ||
| 160 | lemma [simp]: "bool_of_o(True) = 1" | |
| 24892 | 161 | by (simp add: bool_of_o_def) | 
| 13269 | 162 | |
| 163 | lemma [simp]: "bool_of_o(False) = 0" | |
| 24892 | 164 | by (simp add: bool_of_o_def) | 
| 13269 | 165 | |
| 166 | lemma [simp,TC]: "bool_of_o(P) \<in> bool" | |
| 24892 | 167 | by (simp add: bool_of_o_def) | 
| 13269 | 168 | |
| 169 | lemma [simp]: "(bool_of_o(P) = 1) <-> P" | |
| 24892 | 170 | by (simp add: bool_of_o_def) | 
| 13269 | 171 | |
| 172 | lemma [simp]: "(bool_of_o(P) = 0) <-> ~P" | |
| 24892 | 173 | by (simp add: bool_of_o_def) | 
| 13269 | 174 | |
| 13239 | 175 | ML | 
| 176 | {*
 | |
| 177 | val bool_def = thm "bool_def"; | |
| 178 | ||
| 179 | val bool_defs = thms "bool_defs"; | |
| 180 | val singleton_0 = thm "singleton_0"; | |
| 181 | val bool_1I = thm "bool_1I"; | |
| 182 | val bool_0I = thm "bool_0I"; | |
| 183 | val one_not_0 = thm "one_not_0"; | |
| 184 | val one_neq_0 = thm "one_neq_0"; | |
| 185 | val boolE = thm "boolE"; | |
| 186 | val cond_1 = thm "cond_1"; | |
| 187 | val cond_0 = thm "cond_0"; | |
| 188 | val cond_type = thm "cond_type"; | |
| 189 | val cond_simple_type = thm "cond_simple_type"; | |
| 190 | val def_cond_1 = thm "def_cond_1"; | |
| 191 | val def_cond_0 = thm "def_cond_0"; | |
| 192 | val not_1 = thm "not_1"; | |
| 193 | val not_0 = thm "not_0"; | |
| 194 | val and_1 = thm "and_1"; | |
| 195 | val and_0 = thm "and_0"; | |
| 196 | val or_1 = thm "or_1"; | |
| 197 | val or_0 = thm "or_0"; | |
| 198 | val xor_1 = thm "xor_1"; | |
| 199 | val xor_0 = thm "xor_0"; | |
| 200 | val not_type = thm "not_type"; | |
| 201 | val and_type = thm "and_type"; | |
| 202 | val or_type = thm "or_type"; | |
| 203 | val xor_type = thm "xor_type"; | |
| 204 | val bool_typechecks = thms "bool_typechecks"; | |
| 205 | val not_not = thm "not_not"; | |
| 206 | val not_and = thm "not_and"; | |
| 207 | val not_or = thm "not_or"; | |
| 208 | val and_absorb = thm "and_absorb"; | |
| 209 | val and_commute = thm "and_commute"; | |
| 210 | val and_assoc = thm "and_assoc"; | |
| 211 | val and_or_distrib = thm "and_or_distrib"; | |
| 212 | val or_absorb = thm "or_absorb"; | |
| 213 | val or_commute = thm "or_commute"; | |
| 214 | val or_assoc = thm "or_assoc"; | |
| 215 | val or_and_distrib = thm "or_and_distrib"; | |
| 216 | *} | |
| 217 | ||
| 0 | 218 | end |