| author | wenzelm | 
| Thu, 04 Jul 2019 12:31:24 +0200 | |
| changeset 70361 | 34b271c4f400 | 
| parent 70189 | 6d2effbbf8d4 | 
| child 70737 | e4825ec20468 | 
| permissions | -rw-r--r-- | 
| 29629 | 1 | (* Title: HOL/Library/Boolean_Algebra.thy | 
| 2 | Author: Brian Huffman | |
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 3 | *) | 
| 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 4 | |
| 60500 | 5 | section \<open>Boolean Algebras\<close> | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 6 | |
| 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 7 | theory Boolean_Algebra | 
| 63462 | 8 | imports Main | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 9 | begin | 
| 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 10 | |
| 70189 | 11 | locale boolean_algebra = conj: abel_semigroup "(\<^bold>\<sqinter>)" + disj: abel_semigroup "(\<^bold>\<squnion>)" | 
| 12 | for conj :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<^bold>\<sqinter>" 70) | |
| 13 | and disj :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<^bold>\<squnion>" 65) + | |
| 70188 | 14 |   fixes compl :: "'a \<Rightarrow> 'a"  ("\<sim> _" [81] 80)
 | 
| 65343 | 15 |     and zero :: "'a"  ("\<zero>")
 | 
| 16 |     and one  :: "'a"  ("\<one>")
 | |
| 70189 | 17 | assumes conj_disj_distrib: "x \<^bold>\<sqinter> (y \<^bold>\<squnion> z) = (x \<^bold>\<sqinter> y) \<^bold>\<squnion> (x \<^bold>\<sqinter> z)" | 
| 18 | and disj_conj_distrib: "x \<^bold>\<squnion> (y \<^bold>\<sqinter> z) = (x \<^bold>\<squnion> y) \<^bold>\<sqinter> (x \<^bold>\<squnion> z)" | |
| 19 | and conj_one_right: "x \<^bold>\<sqinter> \<one> = x" | |
| 20 | and disj_zero_right: "x \<^bold>\<squnion> \<zero> = x" | |
| 21 | and conj_cancel_right [simp]: "x \<^bold>\<sqinter> \<sim> x = \<zero>" | |
| 22 | and disj_cancel_right [simp]: "x \<^bold>\<squnion> \<sim> x = \<one>" | |
| 54868 | 23 | begin | 
| 34973 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
30663diff
changeset | 24 | |
| 70189 | 25 | sublocale conj: semilattice_neutr "(\<^bold>\<sqinter>)" "\<one>" | 
| 70188 | 26 | proof | 
| 70189 | 27 | show "x \<^bold>\<sqinter> \<one> = x" for x | 
| 70188 | 28 | by (fact conj_one_right) | 
| 70189 | 29 | show "x \<^bold>\<sqinter> x = x" for x | 
| 70188 | 30 | proof - | 
| 70189 | 31 | have "x \<^bold>\<sqinter> x = (x \<^bold>\<sqinter> x) \<^bold>\<squnion> \<zero>" | 
| 70188 | 32 | by (simp add: disj_zero_right) | 
| 70189 | 33 | also have "\<dots> = (x \<^bold>\<sqinter> x) \<^bold>\<squnion> (x \<^bold>\<sqinter> \<sim> x)" | 
| 70188 | 34 | by simp | 
| 70189 | 35 | also have "\<dots> = x \<^bold>\<sqinter> (x \<^bold>\<squnion> \<sim> x)" | 
| 70188 | 36 | by (simp only: conj_disj_distrib) | 
| 70189 | 37 | also have "\<dots> = x \<^bold>\<sqinter> \<one>" | 
| 70188 | 38 | by simp | 
| 39 | also have "\<dots> = x" | |
| 40 | by (simp add: conj_one_right) | |
| 41 | finally show ?thesis . | |
| 42 | qed | |
| 43 | qed | |
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 44 | |
| 70189 | 45 | sublocale disj: semilattice_neutr "(\<^bold>\<squnion>)" "\<zero>" | 
| 70188 | 46 | proof | 
| 70189 | 47 | show "x \<^bold>\<squnion> \<zero> = x" for x | 
| 70188 | 48 | by (fact disj_zero_right) | 
| 70189 | 49 | show "x \<^bold>\<squnion> x = x" for x | 
| 70188 | 50 | proof - | 
| 70189 | 51 | have "x \<^bold>\<squnion> x = (x \<^bold>\<squnion> x) \<^bold>\<sqinter> \<one>" | 
| 70188 | 52 | by simp | 
| 70189 | 53 | also have "\<dots> = (x \<^bold>\<squnion> x) \<^bold>\<sqinter> (x \<^bold>\<squnion> \<sim> x)" | 
| 70188 | 54 | by simp | 
| 70189 | 55 | also have "\<dots> = x \<^bold>\<squnion> (x \<^bold>\<sqinter> \<sim> x)" | 
| 70188 | 56 | by (simp only: disj_conj_distrib) | 
| 70189 | 57 | also have "\<dots> = x \<^bold>\<squnion> \<zero>" | 
| 70188 | 58 | by simp | 
| 59 | also have "\<dots> = x" | |
| 60 | by (simp add: disj_zero_right) | |
| 61 | finally show ?thesis . | |
| 62 | qed | |
| 63 | qed | |
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 64 | |
| 60855 | 65 | |
| 60500 | 66 | subsection \<open>Complement\<close> | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 67 | |
| 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 68 | lemma complement_unique: | 
| 70189 | 69 | assumes 1: "a \<^bold>\<sqinter> x = \<zero>" | 
| 70 | assumes 2: "a \<^bold>\<squnion> x = \<one>" | |
| 71 | assumes 3: "a \<^bold>\<sqinter> y = \<zero>" | |
| 72 | assumes 4: "a \<^bold>\<squnion> y = \<one>" | |
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 73 | shows "x = y" | 
| 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 74 | proof - | 
| 70189 | 75 | from 1 3 have "(a \<^bold>\<sqinter> x) \<^bold>\<squnion> (x \<^bold>\<sqinter> y) = (a \<^bold>\<sqinter> y) \<^bold>\<squnion> (x \<^bold>\<sqinter> y)" | 
| 65343 | 76 | by simp | 
| 70189 | 77 | then have "(x \<^bold>\<sqinter> a) \<^bold>\<squnion> (x \<^bold>\<sqinter> y) = (y \<^bold>\<sqinter> a) \<^bold>\<squnion> (y \<^bold>\<sqinter> x)" | 
| 70188 | 78 | by (simp add: ac_simps) | 
| 70189 | 79 | then have "x \<^bold>\<sqinter> (a \<^bold>\<squnion> y) = y \<^bold>\<sqinter> (a \<^bold>\<squnion> x)" | 
| 65343 | 80 | by (simp add: conj_disj_distrib) | 
| 70189 | 81 | with 2 4 have "x \<^bold>\<sqinter> \<one> = y \<^bold>\<sqinter> \<one>" | 
| 65343 | 82 | by simp | 
| 63462 | 83 | then show "x = y" | 
| 65343 | 84 | by simp | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 85 | qed | 
| 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 86 | |
| 70189 | 87 | lemma compl_unique: "x \<^bold>\<sqinter> y = \<zero> \<Longrightarrow> x \<^bold>\<squnion> y = \<one> \<Longrightarrow> \<sim> x = y" | 
| 63462 | 88 | by (rule complement_unique [OF conj_cancel_right disj_cancel_right]) | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 89 | |
| 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 90 | lemma double_compl [simp]: "\<sim> (\<sim> x) = x" | 
| 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 91 | proof (rule compl_unique) | 
| 70189 | 92 | show "\<sim> x \<^bold>\<sqinter> x = \<zero>" | 
| 70188 | 93 | by (simp only: conj_cancel_right conj.commute) | 
| 70189 | 94 | show "\<sim> x \<^bold>\<squnion> x = \<one>" | 
| 70188 | 95 | by (simp only: disj_cancel_right disj.commute) | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 96 | qed | 
| 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 97 | |
| 63462 | 98 | lemma compl_eq_compl_iff [simp]: "\<sim> x = \<sim> y \<longleftrightarrow> x = y" | 
| 99 | by (rule inj_eq [OF inj_on_inverseI]) (rule double_compl) | |
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 100 | |
| 60855 | 101 | |
| 60500 | 102 | subsection \<open>Conjunction\<close> | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 103 | |
| 70189 | 104 | lemma conj_zero_right [simp]: "x \<^bold>\<sqinter> \<zero> = \<zero>" | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 105 | proof - | 
| 70189 | 106 | from conj_cancel_right have "x \<^bold>\<sqinter> \<zero> = x \<^bold>\<sqinter> (x \<^bold>\<sqinter> \<sim> x)" | 
| 65343 | 107 | by simp | 
| 70189 | 108 | also from conj_assoc have "\<dots> = (x \<^bold>\<sqinter> x) \<^bold>\<sqinter> \<sim> x" | 
| 70188 | 109 | by (simp only: ac_simps) | 
| 70189 | 110 | also from conj_absorb have "\<dots> = x \<^bold>\<sqinter> \<sim> x" | 
| 65343 | 111 | by simp | 
| 112 | also have "\<dots> = \<zero>" | |
| 113 | by simp | |
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 114 | finally show ?thesis . | 
| 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 115 | qed | 
| 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 116 | |
| 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 117 | lemma compl_one [simp]: "\<sim> \<one> = \<zero>" | 
| 63462 | 118 | by (rule compl_unique [OF conj_zero_right disj_zero_right]) | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 119 | |
| 70189 | 120 | lemma conj_zero_left [simp]: "\<zero> \<^bold>\<sqinter> x = \<zero>" | 
| 70188 | 121 | by (subst conj.commute) (rule conj_zero_right) | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 122 | |
| 70189 | 123 | lemma conj_cancel_left [simp]: "\<sim> x \<^bold>\<sqinter> x = \<zero>" | 
| 70188 | 124 | by (subst conj.commute) (rule conj_cancel_right) | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 125 | |
| 70189 | 126 | lemma conj_disj_distrib2: "(y \<^bold>\<squnion> z) \<^bold>\<sqinter> x = (y \<^bold>\<sqinter> x) \<^bold>\<squnion> (z \<^bold>\<sqinter> x)" | 
| 70188 | 127 | by (simp only: conj.commute conj_disj_distrib) | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 128 | |
| 63462 | 129 | lemmas conj_disj_distribs = conj_disj_distrib conj_disj_distrib2 | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 130 | |
| 70189 | 131 | lemma conj_assoc: "(x \<^bold>\<sqinter> y) \<^bold>\<sqinter> z = x \<^bold>\<sqinter> (y \<^bold>\<sqinter> z)" | 
| 70188 | 132 | by (fact ac_simps) | 
| 133 | ||
| 70189 | 134 | lemma conj_commute: "x \<^bold>\<sqinter> y = y \<^bold>\<sqinter> x" | 
| 70188 | 135 | by (fact ac_simps) | 
| 136 | ||
| 137 | lemmas conj_left_commute = conj.left_commute | |
| 138 | lemmas conj_ac = conj.assoc conj.commute conj.left_commute | |
| 139 | ||
| 70189 | 140 | lemma conj_one_left: "\<one> \<^bold>\<sqinter> x = x" | 
| 70188 | 141 | by (fact conj.left_neutral) | 
| 142 | ||
| 70189 | 143 | lemma conj_left_absorb: "x \<^bold>\<sqinter> (x \<^bold>\<sqinter> y) = x \<^bold>\<sqinter> y" | 
| 70188 | 144 | by (fact conj.left_idem) | 
| 145 | ||
| 70189 | 146 | lemma conj_absorb: "x \<^bold>\<sqinter> x = x" | 
| 70188 | 147 | by (fact conj.idem) | 
| 148 | ||
| 60855 | 149 | |
| 60500 | 150 | subsection \<open>Disjunction\<close> | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 151 | |
| 70189 | 152 | interpretation dual: boolean_algebra "(\<^bold>\<squnion>)" "(\<^bold>\<sqinter>)" compl \<one> \<zero> | 
| 70188 | 153 | apply standard | 
| 154 | apply (rule disj_conj_distrib) | |
| 155 | apply (rule conj_disj_distrib) | |
| 156 | apply simp_all | |
| 157 | done | |
| 158 | ||
| 159 | lemma compl_zero [simp]: "\<sim> \<zero> = \<one>" | |
| 160 | by (fact dual.compl_one) | |
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 161 | |
| 70189 | 162 | lemma disj_one_right [simp]: "x \<^bold>\<squnion> \<one> = \<one>" | 
| 70188 | 163 | by (fact dual.conj_zero_right) | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 164 | |
| 70189 | 165 | lemma disj_one_left [simp]: "\<one> \<^bold>\<squnion> x = \<one>" | 
| 70188 | 166 | by (fact dual.conj_zero_left) | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 167 | |
| 70189 | 168 | lemma disj_cancel_left [simp]: "\<sim> x \<^bold>\<squnion> x = \<one>" | 
| 70188 | 169 | by (rule dual.conj_cancel_left) | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 170 | |
| 70189 | 171 | lemma disj_conj_distrib2: "(y \<^bold>\<sqinter> z) \<^bold>\<squnion> x = (y \<^bold>\<squnion> x) \<^bold>\<sqinter> (z \<^bold>\<squnion> x)" | 
| 70188 | 172 | by (rule dual.conj_disj_distrib2) | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 173 | |
| 63462 | 174 | lemmas disj_conj_distribs = disj_conj_distrib disj_conj_distrib2 | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 175 | |
| 70189 | 176 | lemma disj_assoc: "(x \<^bold>\<squnion> y) \<^bold>\<squnion> z = x \<^bold>\<squnion> (y \<^bold>\<squnion> z)" | 
| 70188 | 177 | by (fact ac_simps) | 
| 178 | ||
| 70189 | 179 | lemma disj_commute: "x \<^bold>\<squnion> y = y \<^bold>\<squnion> x" | 
| 70188 | 180 | by (fact ac_simps) | 
| 181 | ||
| 182 | lemmas disj_left_commute = disj.left_commute | |
| 183 | ||
| 184 | lemmas disj_ac = disj.assoc disj.commute disj.left_commute | |
| 185 | ||
| 70189 | 186 | lemma disj_zero_left: "\<zero> \<^bold>\<squnion> x = x" | 
| 70188 | 187 | by (fact disj.left_neutral) | 
| 188 | ||
| 70189 | 189 | lemma disj_left_absorb: "x \<^bold>\<squnion> (x \<^bold>\<squnion> y) = x \<^bold>\<squnion> y" | 
| 70188 | 190 | by (fact disj.left_idem) | 
| 191 | ||
| 70189 | 192 | lemma disj_absorb: "x \<^bold>\<squnion> x = x" | 
| 70188 | 193 | by (fact disj.idem) | 
| 194 | ||
| 60855 | 195 | |
| 60500 | 196 | subsection \<open>De Morgan's Laws\<close> | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 197 | |
| 70189 | 198 | lemma de_Morgan_conj [simp]: "\<sim> (x \<^bold>\<sqinter> y) = \<sim> x \<^bold>\<squnion> \<sim> y" | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 199 | proof (rule compl_unique) | 
| 70189 | 200 | have "(x \<^bold>\<sqinter> y) \<^bold>\<sqinter> (\<sim> x \<^bold>\<squnion> \<sim> y) = ((x \<^bold>\<sqinter> y) \<^bold>\<sqinter> \<sim> x) \<^bold>\<squnion> ((x \<^bold>\<sqinter> y) \<^bold>\<sqinter> \<sim> y)" | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 201 | by (rule conj_disj_distrib) | 
| 70189 | 202 | also have "\<dots> = (y \<^bold>\<sqinter> (x \<^bold>\<sqinter> \<sim> x)) \<^bold>\<squnion> (x \<^bold>\<sqinter> (y \<^bold>\<sqinter> \<sim> y))" | 
| 24357 | 203 | by (simp only: conj_ac) | 
| 70189 | 204 | finally show "(x \<^bold>\<sqinter> y) \<^bold>\<sqinter> (\<sim> x \<^bold>\<squnion> \<sim> y) = \<zero>" | 
| 24357 | 205 | by (simp only: conj_cancel_right conj_zero_right disj_zero_right) | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 206 | next | 
| 70189 | 207 | have "(x \<^bold>\<sqinter> y) \<^bold>\<squnion> (\<sim> x \<^bold>\<squnion> \<sim> y) = (x \<^bold>\<squnion> (\<sim> x \<^bold>\<squnion> \<sim> y)) \<^bold>\<sqinter> (y \<^bold>\<squnion> (\<sim> x \<^bold>\<squnion> \<sim> y))" | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 208 | by (rule disj_conj_distrib2) | 
| 70189 | 209 | also have "\<dots> = (\<sim> y \<^bold>\<squnion> (x \<^bold>\<squnion> \<sim> x)) \<^bold>\<sqinter> (\<sim> x \<^bold>\<squnion> (y \<^bold>\<squnion> \<sim> y))" | 
| 24357 | 210 | by (simp only: disj_ac) | 
| 70189 | 211 | finally show "(x \<^bold>\<sqinter> y) \<^bold>\<squnion> (\<sim> x \<^bold>\<squnion> \<sim> y) = \<one>" | 
| 24357 | 212 | by (simp only: disj_cancel_right disj_one_right conj_one_right) | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 213 | qed | 
| 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 214 | |
| 70189 | 215 | lemma de_Morgan_disj [simp]: "\<sim> (x \<^bold>\<squnion> y) = \<sim> x \<^bold>\<sqinter> \<sim> y" | 
| 70188 | 216 | using dual.boolean_algebra_axioms by (rule boolean_algebra.de_Morgan_conj) | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 217 | |
| 60855 | 218 | |
| 60500 | 219 | subsection \<open>Symmetric Difference\<close> | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 220 | |
| 70187 
2082287357e6
avoid separate type class for mere definitional extension
 haftmann parents: 
70186diff
changeset | 221 | definition xor :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<oplus>" 65) | 
| 70189 | 222 | where "x \<oplus> y = (x \<^bold>\<sqinter> \<sim> y) \<^bold>\<squnion> (\<sim> x \<^bold>\<sqinter> y)" | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 223 | |
| 61605 | 224 | sublocale xor: abel_semigroup xor | 
| 60855 | 225 | proof | 
| 34973 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
30663diff
changeset | 226 | fix x y z :: 'a | 
| 70189 | 227 | let ?t = "(x \<^bold>\<sqinter> y \<^bold>\<sqinter> z) \<^bold>\<squnion> (x \<^bold>\<sqinter> \<sim> y \<^bold>\<sqinter> \<sim> z) \<^bold>\<squnion> (\<sim> x \<^bold>\<sqinter> y \<^bold>\<sqinter> \<sim> z) \<^bold>\<squnion> (\<sim> x \<^bold>\<sqinter> \<sim> y \<^bold>\<sqinter> z)" | 
| 228 | have "?t \<^bold>\<squnion> (z \<^bold>\<sqinter> x \<^bold>\<sqinter> \<sim> x) \<^bold>\<squnion> (z \<^bold>\<sqinter> y \<^bold>\<sqinter> \<sim> y) = ?t \<^bold>\<squnion> (x \<^bold>\<sqinter> y \<^bold>\<sqinter> \<sim> y) \<^bold>\<squnion> (x \<^bold>\<sqinter> z \<^bold>\<sqinter> \<sim> z)" | |
| 24357 | 229 | by (simp only: conj_cancel_right conj_zero_right) | 
| 63462 | 230 | then show "(x \<oplus> y) \<oplus> z = x \<oplus> (y \<oplus> z)" | 
| 65343 | 231 | by (simp only: xor_def de_Morgan_disj de_Morgan_conj double_compl) | 
| 232 | (simp only: conj_disj_distribs conj_ac disj_ac) | |
| 34973 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
30663diff
changeset | 233 | show "x \<oplus> y = y \<oplus> x" | 
| 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
30663diff
changeset | 234 | by (simp only: xor_def conj_commute disj_commute) | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 235 | qed | 
| 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 236 | |
| 34973 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
30663diff
changeset | 237 | lemmas xor_assoc = xor.assoc | 
| 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
30663diff
changeset | 238 | lemmas xor_commute = xor.commute | 
| 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
30663diff
changeset | 239 | lemmas xor_left_commute = xor.left_commute | 
| 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
30663diff
changeset | 240 | |
| 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
30663diff
changeset | 241 | lemmas xor_ac = xor.assoc xor.commute xor.left_commute | 
| 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
30663diff
changeset | 242 | |
| 70189 | 243 | lemma xor_def2: "x \<oplus> y = (x \<^bold>\<squnion> y) \<^bold>\<sqinter> (\<sim> x \<^bold>\<squnion> \<sim> y)" | 
| 63462 | 244 | by (simp only: xor_def conj_disj_distribs disj_ac conj_ac conj_cancel_right disj_zero_left) | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 245 | |
| 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 246 | lemma xor_zero_right [simp]: "x \<oplus> \<zero> = x" | 
| 63462 | 247 | by (simp only: xor_def compl_zero conj_one_right conj_zero_right disj_zero_right) | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 248 | |
| 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 249 | lemma xor_zero_left [simp]: "\<zero> \<oplus> x = x" | 
| 63462 | 250 | by (subst xor_commute) (rule xor_zero_right) | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 251 | |
| 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 252 | lemma xor_one_right [simp]: "x \<oplus> \<one> = \<sim> x" | 
| 63462 | 253 | by (simp only: xor_def compl_one conj_zero_right conj_one_right disj_zero_left) | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 254 | |
| 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 255 | lemma xor_one_left [simp]: "\<one> \<oplus> x = \<sim> x" | 
| 63462 | 256 | by (subst xor_commute) (rule xor_one_right) | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 257 | |
| 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 258 | lemma xor_self [simp]: "x \<oplus> x = \<zero>" | 
| 63462 | 259 | by (simp only: xor_def conj_cancel_right conj_cancel_left disj_zero_right) | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 260 | |
| 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 261 | lemma xor_left_self [simp]: "x \<oplus> (x \<oplus> y) = y" | 
| 63462 | 262 | by (simp only: xor_assoc [symmetric] xor_self xor_zero_left) | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 263 | |
| 29996 | 264 | lemma xor_compl_left [simp]: "\<sim> x \<oplus> y = \<sim> (x \<oplus> y)" | 
| 63462 | 265 | apply (simp only: xor_def de_Morgan_disj de_Morgan_conj double_compl) | 
| 266 | apply (simp only: conj_disj_distribs) | |
| 267 | apply (simp only: conj_cancel_right conj_cancel_left) | |
| 268 | apply (simp only: disj_zero_left disj_zero_right) | |
| 269 | apply (simp only: disj_ac conj_ac) | |
| 270 | done | |
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 271 | |
| 29996 | 272 | lemma xor_compl_right [simp]: "x \<oplus> \<sim> y = \<sim> (x \<oplus> y)" | 
| 63462 | 273 | apply (simp only: xor_def de_Morgan_disj de_Morgan_conj double_compl) | 
| 274 | apply (simp only: conj_disj_distribs) | |
| 275 | apply (simp only: conj_cancel_right conj_cancel_left) | |
| 276 | apply (simp only: disj_zero_left disj_zero_right) | |
| 277 | apply (simp only: disj_ac conj_ac) | |
| 278 | done | |
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 279 | |
| 29996 | 280 | lemma xor_cancel_right: "x \<oplus> \<sim> x = \<one>" | 
| 63462 | 281 | by (simp only: xor_compl_right xor_self compl_zero) | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 282 | |
| 29996 | 283 | lemma xor_cancel_left: "\<sim> x \<oplus> x = \<one>" | 
| 63462 | 284 | by (simp only: xor_compl_left xor_self compl_zero) | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 285 | |
| 70189 | 286 | lemma conj_xor_distrib: "x \<^bold>\<sqinter> (y \<oplus> z) = (x \<^bold>\<sqinter> y) \<oplus> (x \<^bold>\<sqinter> z)" | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 287 | proof - | 
| 70189 | 288 | have *: "(x \<^bold>\<sqinter> y \<^bold>\<sqinter> \<sim> z) \<^bold>\<squnion> (x \<^bold>\<sqinter> \<sim> y \<^bold>\<sqinter> z) = | 
| 289 | (y \<^bold>\<sqinter> x \<^bold>\<sqinter> \<sim> x) \<^bold>\<squnion> (z \<^bold>\<sqinter> x \<^bold>\<sqinter> \<sim> x) \<^bold>\<squnion> (x \<^bold>\<sqinter> y \<^bold>\<sqinter> \<sim> z) \<^bold>\<squnion> (x \<^bold>\<sqinter> \<sim> y \<^bold>\<sqinter> z)" | |
| 24357 | 290 | by (simp only: conj_cancel_right conj_zero_right disj_zero_left) | 
| 70189 | 291 | then show "x \<^bold>\<sqinter> (y \<oplus> z) = (x \<^bold>\<sqinter> y) \<oplus> (x \<^bold>\<sqinter> z)" | 
| 24357 | 292 | by (simp (no_asm_use) only: | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 293 | xor_def de_Morgan_disj de_Morgan_conj double_compl | 
| 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 294 | conj_disj_distribs conj_ac disj_ac) | 
| 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 295 | qed | 
| 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 296 | |
| 70189 | 297 | lemma conj_xor_distrib2: "(y \<oplus> z) \<^bold>\<sqinter> x = (y \<^bold>\<sqinter> x) \<oplus> (z \<^bold>\<sqinter> x)" | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 298 | proof - | 
| 70189 | 299 | have "x \<^bold>\<sqinter> (y \<oplus> z) = (x \<^bold>\<sqinter> y) \<oplus> (x \<^bold>\<sqinter> z)" | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 300 | by (rule conj_xor_distrib) | 
| 70189 | 301 | then show "(y \<oplus> z) \<^bold>\<sqinter> x = (y \<^bold>\<sqinter> x) \<oplus> (z \<^bold>\<sqinter> x)" | 
| 24357 | 302 | by (simp only: conj_commute) | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 303 | qed | 
| 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 304 | |
| 60855 | 305 | lemmas conj_xor_distribs = conj_xor_distrib conj_xor_distrib2 | 
| 24332 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 306 | |
| 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 307 | end | 
| 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 308 | |
| 
e3a2b75b1cf9
boolean algebras as locales and numbers as types by Brian Huffman
 kleing parents: diff
changeset | 309 | end |