# HG changeset patch # User haftmann # Date 1555783342 0 # Node ID e626bffe28bca8364c1e97f22a07e5e08709cc93 # Parent 2082287357e694e8c7a098820a0fe5dca8fb23a1 more use of existing locales diff -r 2082287357e6 -r e626bffe28bc src/HOL/Library/Boolean_Algebra.thy --- a/src/HOL/Library/Boolean_Algebra.thy Sat Apr 20 18:02:21 2019 +0000 +++ b/src/HOL/Library/Boolean_Algebra.thy Sat Apr 20 18:02:22 2019 +0000 @@ -8,49 +8,59 @@ imports Main begin -locale boolean_algebra = - fixes conj :: "'a \ 'a \ 'a" (infixr "\" 70) - and disj :: "'a \ 'a \ 'a" (infixr "\" 65) - and compl :: "'a \ 'a" ("\ _" [81] 80) +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_assoc: "(x \ y) \ z = x \ (y \ z)" - and disj_assoc: "(x \ y) \ z = x \ (y \ z)" - and conj_commute: "x \ y = y \ x" - and disj_commute: "x \ y = y \ x" - and conj_disj_distrib: "x \ (y \ z) = (x \ y) \ (x \ z)" + 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 [simp]: "x \ \ = x" - and disj_zero_right [simp]: "x \ \ = x" + 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: abel_semigroup conj - by standard (fact conj_assoc conj_commute)+ - -sublocale disj: abel_semigroup disj - by standard (fact disj_assoc disj_commute)+ - -lemmas conj_left_commute = conj.left_commute -lemmas disj_left_commute = disj.left_commute - -lemmas conj_ac = conj.assoc conj.commute conj.left_commute -lemmas disj_ac = disj.assoc disj.commute disj.left_commute +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 -lemma dual: "boolean_algebra disj conj compl one zero" - apply (rule boolean_algebra.intro) - apply (rule disj_assoc) - apply (rule conj_assoc) - apply (rule disj_commute) - apply (rule conj_commute) - apply (rule disj_conj_distrib) - apply (rule conj_disj_distrib) - apply (rule disj_zero_right) - apply (rule conj_one_right) - apply (rule disj_cancel_right) - apply (rule conj_cancel_right) - done +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\ @@ -65,7 +75,7 @@ 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: conj_commute) + 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 \ \" @@ -80,9 +90,9 @@ lemma double_compl [simp]: "\ (\ x) = x" proof (rule compl_unique) show "\ x \ x = \" - by (simp only: conj_cancel_right conj_commute) + by (simp only: conj_cancel_right conj.commute) show "\ x \ x = \" - by (simp only: disj_cancel_right disj_commute) + by (simp only: disj_cancel_right disj.commute) qed lemma compl_eq_compl_iff [simp]: "\ x = \ y \ x = y" @@ -91,27 +101,12 @@ subsection \Conjunction\ -lemma conj_absorb [simp]: "x \ x = 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: conj_disj_distrib) - also have "\ = x \ \" - by simp - also have "\ = x" - by simp - finally show ?thesis . -qed - lemma conj_zero_right [simp]: "x \ \ = \" proof - from conj_cancel_right have "x \ \ = x \ (x \ \ x)" by simp also from conj_assoc have "\ = (x \ x) \ \ x" - by (simp only:) + by (simp only: ac_simps) also from conj_absorb have "\ = x \ \ x" by simp also have "\ = \" @@ -123,51 +118,80 @@ 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_one_left [simp]: "\ \ x = x" - by (subst conj_commute) (rule conj_one_right) + 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_left_absorb [simp]: "x \ (x \ y) = x \ y" - by (simp only: conj_assoc [symmetric] conj_absorb) + 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) + 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\ -lemma disj_absorb [simp]: "x \ x = x" - by (rule boolean_algebra.conj_absorb [OF dual]) +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 (rule boolean_algebra.conj_zero_right [OF dual]) - -lemma compl_zero [simp]: "\ \ = \" - by (rule boolean_algebra.compl_one [OF dual]) - -lemma disj_zero_left [simp]: "\ \ x = x" - by (rule boolean_algebra.conj_one_left [OF dual]) + by (fact dual.conj_zero_right) lemma disj_one_left [simp]: "\ \ x = \" - by (rule boolean_algebra.conj_zero_left [OF dual]) + by (fact dual.conj_zero_left) lemma disj_cancel_left [simp]: "\ x \ x = \" - by (rule boolean_algebra.conj_cancel_left [OF dual]) - -lemma disj_left_absorb [simp]: "x \ (x \ y) = x \ y" - by (rule boolean_algebra.conj_left_absorb [OF dual]) + by (rule dual.conj_cancel_left) lemma disj_conj_distrib2: "(y \ z) \ x = (y \ x) \ (z \ x)" - by (rule boolean_algebra.conj_disj_distrib2 [OF dual]) + 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\ @@ -189,7 +213,7 @@ qed lemma de_Morgan_disj [simp]: "\ (x \ y) = \ x \ \ y" - by (rule boolean_algebra.de_Morgan_conj [OF dual]) + using dual.boolean_algebra_axioms by (rule boolean_algebra.de_Morgan_conj) subsection \Symmetric Difference\