Theory Boolean_Algebra
section ‹Boolean Algebras›
theory Boolean_Algebra
imports Main
begin
locale boolean_algebra = conj: abel_semigroup "(❙⊓)" + disj: abel_semigroup "(❙⊔)"
for conj :: "'a ⇒ 'a ⇒ 'a" (infixr "❙⊓" 70)
and disj :: "'a ⇒ 'a ⇒ 'a" (infixr "❙⊔" 65) +
fixes compl :: "'a ⇒ 'a" ("∼ _" [81] 80)
and zero :: "'a" ("𝟬")
and one :: "'a" ("𝟭")
assumes conj_disj_distrib: "x ❙⊓ (y ❙⊔ z) = (x ❙⊓ y) ❙⊔ (x ❙⊓ z)"
and disj_conj_distrib: "x ❙⊔ (y ❙⊓ z) = (x ❙⊔ y) ❙⊓ (x ❙⊔ z)"
and conj_one_right: "x ❙⊓ 𝟭 = x"
and disj_zero_right: "x ❙⊔ 𝟬 = x"
and conj_cancel_right [simp]: "x ❙⊓ ∼ x = 𝟬"
and disj_cancel_right [simp]: "x ❙⊔ ∼ x = 𝟭"
begin
sublocale conj: semilattice_neutr "(❙⊓)" "𝟭"
proof
show "x ❙⊓ 𝟭 = x" for x
by (fact conj_one_right)
show "x ❙⊓ x = x" for x
proof -
have "x ❙⊓ x = (x ❙⊓ x) ❙⊔ 𝟬"
by (simp add: disj_zero_right)
also have "… = (x ❙⊓ x) ❙⊔ (x ❙⊓ ∼ x)"
by simp
also have "… = x ❙⊓ (x ❙⊔ ∼ x)"
by (simp only: conj_disj_distrib)
also have "… = x ❙⊓ 𝟭"
by simp
also have "… = x"
by (simp add: conj_one_right)
finally show ?thesis .
qed
qed
sublocale disj: semilattice_neutr "(❙⊔)" "𝟬"
proof
show "x ❙⊔ 𝟬 = x" for x
by (fact disj_zero_right)
show "x ❙⊔ x = x" for x
proof -
have "x ❙⊔ x = (x ❙⊔ x) ❙⊓ 𝟭"
by simp
also have "… = (x ❙⊔ x) ❙⊓ (x ❙⊔ ∼ x)"
by simp
also have "… = x ❙⊔ (x ❙⊓ ∼ x)"
by (simp only: disj_conj_distrib)
also have "… = x ❙⊔ 𝟬"
by simp
also have "… = x"
by (simp add: disj_zero_right)
finally show ?thesis .
qed
qed
subsection ‹Complement›
lemma complement_unique:
assumes 1: "a ❙⊓ x = 𝟬"
assumes 2: "a ❙⊔ x = 𝟭"
assumes 3: "a ❙⊓ y = 𝟬"
assumes 4: "a ❙⊔ y = 𝟭"
shows "x = y"
proof -
from 1 3 have "(a ❙⊓ x) ❙⊔ (x ❙⊓ y) = (a ❙⊓ y) ❙⊔ (x ❙⊓ y)"
by simp
then have "(x ❙⊓ a) ❙⊔ (x ❙⊓ y) = (y ❙⊓ a) ❙⊔ (y ❙⊓ x)"
by (simp add: ac_simps)
then have "x ❙⊓ (a ❙⊔ y) = y ❙⊓ (a ❙⊔ x)"
by (simp add: conj_disj_distrib)
with 2 4 have "x ❙⊓ 𝟭 = y ❙⊓ 𝟭"
by simp
then show "x = y"
by simp
qed
lemma compl_unique: "x ❙⊓ y = 𝟬 ⟹ x ❙⊔ y = 𝟭 ⟹ ∼ x = y"
by (rule complement_unique [OF conj_cancel_right disj_cancel_right])
lemma double_compl [simp]: "∼ (∼ x) = x"
proof (rule compl_unique)
show "∼ x ❙⊓ x = 𝟬"
by (simp only: conj_cancel_right conj.commute)
show "∼ x ❙⊔ x = 𝟭"
by (simp only: disj_cancel_right disj.commute)
qed
lemma compl_eq_compl_iff [simp]: "∼ x = ∼ y ⟷ x = y"
by (rule inj_eq [OF inj_on_inverseI]) (rule double_compl)
subsection ‹Conjunction›
lemma conj_zero_right [simp]: "x ❙⊓ 𝟬 = 𝟬"
using conj.left_idem conj_cancel_right by fastforce
lemma compl_one [simp]: "∼ 𝟭 = 𝟬"
by (rule compl_unique [OF conj_zero_right disj_zero_right])
lemma conj_zero_left [simp]: "𝟬 ❙⊓ x = 𝟬"
by (subst conj.commute) (rule conj_zero_right)
lemma conj_cancel_left [simp]: "∼ x ❙⊓ x = 𝟬"
by (subst conj.commute) (rule conj_cancel_right)
lemma conj_disj_distrib2: "(y ❙⊔ z) ❙⊓ x = (y ❙⊓ x) ❙⊔ (z ❙⊓ x)"
by (simp only: conj.commute conj_disj_distrib)
lemmas conj_disj_distribs = conj_disj_distrib conj_disj_distrib2
lemma conj_assoc: "(x ❙⊓ y) ❙⊓ z = x ❙⊓ (y ❙⊓ z)"
by (fact ac_simps)
lemma conj_commute: "x ❙⊓ y = y ❙⊓ x"
by (fact ac_simps)
lemmas conj_left_commute = conj.left_commute
lemmas conj_ac = conj.assoc conj.commute conj.left_commute
lemma conj_one_left: "𝟭 ❙⊓ x = x"
by (fact conj.left_neutral)
lemma conj_left_absorb: "x ❙⊓ (x ❙⊓ y) = x ❙⊓ y"
by (fact conj.left_idem)
lemma conj_absorb: "x ❙⊓ x = x"
by (fact conj.idem)
subsection ‹Disjunction›
interpretation dual: boolean_algebra "(❙⊔)" "(❙⊓)" compl 𝟭 𝟬
apply standard
apply (rule disj_conj_distrib)
apply (rule conj_disj_distrib)
apply simp_all
done
lemma compl_zero [simp]: "∼ 𝟬 = 𝟭"
by (fact dual.compl_one)
lemma disj_one_right [simp]: "x ❙⊔ 𝟭 = 𝟭"
by (fact dual.conj_zero_right)
lemma disj_one_left [simp]: "𝟭 ❙⊔ x = 𝟭"
by (fact dual.conj_zero_left)
lemma disj_cancel_left [simp]: "∼ x ❙⊔ x = 𝟭"
by (rule dual.conj_cancel_left)
lemma disj_conj_distrib2: "(y ❙⊓ z) ❙⊔ x = (y ❙⊔ x) ❙⊓ (z ❙⊔ x)"
by (rule dual.conj_disj_distrib2)
lemmas disj_conj_distribs = disj_conj_distrib disj_conj_distrib2
lemma disj_assoc: "(x ❙⊔ y) ❙⊔ z = x ❙⊔ (y ❙⊔ z)"
by (fact ac_simps)
lemma disj_commute: "x ❙⊔ y = y ❙⊔ x"
by (fact ac_simps)
lemmas disj_left_commute = disj.left_commute
lemmas disj_ac = disj.assoc disj.commute disj.left_commute
lemma disj_zero_left: "𝟬 ❙⊔ x = x"
by (fact disj.left_neutral)
lemma disj_left_absorb: "x ❙⊔ (x ❙⊔ y) = x ❙⊔ y"
by (fact disj.left_idem)
lemma disj_absorb: "x ❙⊔ x = x"
by (fact disj.idem)
subsection ‹De Morgan's Laws›
lemma de_Morgan_conj [simp]: "∼ (x ❙⊓ y) = ∼ x ❙⊔ ∼ y"
proof (rule compl_unique)
have "(x ❙⊓ y) ❙⊓ (∼ x ❙⊔ ∼ y) = ((x ❙⊓ y) ❙⊓ ∼ x) ❙⊔ ((x ❙⊓ y) ❙⊓ ∼ y)"
by (rule conj_disj_distrib)
also have "… = (y ❙⊓ (x ❙⊓ ∼ x)) ❙⊔ (x ❙⊓ (y ❙⊓ ∼ y))"
by (simp only: conj_ac)
finally show "(x ❙⊓ y) ❙⊓ (∼ x ❙⊔ ∼ y) = 𝟬"
by (simp only: conj_cancel_right conj_zero_right disj_zero_right)
next
have "(x ❙⊓ y) ❙⊔ (∼ x ❙⊔ ∼ y) = (x ❙⊔ (∼ x ❙⊔ ∼ y)) ❙⊓ (y ❙⊔ (∼ x ❙⊔ ∼ y))"
by (rule disj_conj_distrib2)
also have "… = (∼ y ❙⊔ (x ❙⊔ ∼ x)) ❙⊓ (∼ x ❙⊔ (y ❙⊔ ∼ y))"
by (simp only: disj_ac)
finally show "(x ❙⊓ y) ❙⊔ (∼ x ❙⊔ ∼ y) = 𝟭"
by (simp only: disj_cancel_right disj_one_right conj_one_right)
qed
lemma de_Morgan_disj [simp]: "∼ (x ❙⊔ y) = ∼ x ❙⊓ ∼ y"
using dual.boolean_algebra_axioms by (rule boolean_algebra.de_Morgan_conj)
subsection ‹Symmetric Difference›
definition xor :: "'a ⇒ 'a ⇒ 'a" (infixr "⊕" 65)
where "x ⊕ y = (x ❙⊓ ∼ y) ❙⊔ (∼ x ❙⊓ y)"
sublocale xor: comm_monoid xor 𝟬
proof
fix x y z :: 'a
let ?t = "(x ❙⊓ y ❙⊓ z) ❙⊔ (x ❙⊓ ∼ y ❙⊓ ∼ z) ❙⊔ (∼ x ❙⊓ y ❙⊓ ∼ z) ❙⊔ (∼ x ❙⊓ ∼ y ❙⊓ z)"
have "?t ❙⊔ (z ❙⊓ x ❙⊓ ∼ x) ❙⊔ (z ❙⊓ y ❙⊓ ∼ y) = ?t ❙⊔ (x ❙⊓ y ❙⊓ ∼ y) ❙⊔ (x ❙⊓ z ❙⊓ ∼ z)"
by (simp only: conj_cancel_right conj_zero_right)
then show "(x ⊕ y) ⊕ z = x ⊕ (y ⊕ z)"
by (simp only: xor_def de_Morgan_disj de_Morgan_conj double_compl)
(simp only: conj_disj_distribs conj_ac disj_ac)
show "x ⊕ y = y ⊕ x"
by (simp only: xor_def conj_commute disj_commute)
show "x ⊕ 𝟬 = x"
by (simp add: xor_def)
qed
lemmas xor_assoc = xor.assoc
lemmas xor_commute = xor.commute
lemmas xor_left_commute = xor.left_commute
lemmas xor_ac = xor.assoc xor.commute xor.left_commute
lemma xor_def2: "x ⊕ y = (x ❙⊔ y) ❙⊓ (∼ x ❙⊔ ∼ y)"
using conj.commute conj_disj_distrib2 disj.commute xor_def by auto
lemma xor_zero_right [simp]: "x ⊕ 𝟬 = x"
by (simp only: xor_def compl_zero conj_one_right conj_zero_right disj_zero_right)
lemma xor_zero_left [simp]: "𝟬 ⊕ x = x"
by (subst xor_commute) (rule xor_zero_right)
lemma xor_one_right [simp]: "x ⊕ 𝟭 = ∼ x"
by (simp only: xor_def compl_one conj_zero_right conj_one_right disj_zero_left)
lemma xor_one_left [simp]: "𝟭 ⊕ x = ∼ x"
by (subst xor_commute) (rule xor_one_right)
lemma xor_self [simp]: "x ⊕ x = 𝟬"
by (simp only: xor_def conj_cancel_right conj_cancel_left disj_zero_right)
lemma xor_left_self [simp]: "x ⊕ (x ⊕ y) = y"
by (simp only: xor_assoc [symmetric] xor_self xor_zero_left)
lemma xor_compl_left [simp]: "∼ x ⊕ y = ∼ (x ⊕ y)"
by (metis xor_assoc xor_one_left)
lemma xor_compl_right [simp]: "x ⊕ ∼ y = ∼ (x ⊕ y)"
using xor_commute xor_compl_left by auto
lemma xor_cancel_right: "x ⊕ ∼ x = 𝟭"
by (simp only: xor_compl_right xor_self compl_zero)
lemma xor_cancel_left: "∼ x ⊕ x = 𝟭"
by (simp only: xor_compl_left xor_self compl_zero)
lemma conj_xor_distrib: "x ❙⊓ (y ⊕ z) = (x ❙⊓ y) ⊕ (x ❙⊓ z)"
proof -
have *: "(x ❙⊓ y ❙⊓ ∼ z) ❙⊔ (x ❙⊓ ∼ y ❙⊓ z) =
(y ❙⊓ x ❙⊓ ∼ x) ❙⊔ (z ❙⊓ x ❙⊓ ∼ x) ❙⊔ (x ❙⊓ y ❙⊓ ∼ z) ❙⊔ (x ❙⊓ ∼ y ❙⊓ z)"
by (simp only: conj_cancel_right conj_zero_right disj_zero_left)
then show "x ❙⊓ (y ⊕ z) = (x ❙⊓ y) ⊕ (x ❙⊓ z)"
by (simp (no_asm_use) only:
xor_def de_Morgan_disj de_Morgan_conj double_compl
conj_disj_distribs conj_ac disj_ac)
qed
lemma conj_xor_distrib2: "(y ⊕ z) ❙⊓ x = (y ❙⊓ x) ⊕ (z ❙⊓ x)"
by (simp add: conj.commute conj_xor_distrib)
lemmas conj_xor_distribs = conj_xor_distrib conj_xor_distrib2
end
end