# HG changeset patch # User haftmann # Date 1555783342 0 # Node ID 6d2effbbf8d4fc27edaf40c74a5fab93ed655797 # Parent e626bffe28bca8364c1e97f22a07e5e08709cc93 follow convention of bold local syntax diff -r e626bffe28bc -r 6d2effbbf8d4 src/HOL/Library/Boolean_Algebra.thy --- a/src/HOL/Library/Boolean_Algebra.thy Sat Apr 20 18:02:22 2019 +0000 +++ b/src/HOL/Library/Boolean_Algebra.thy Sat Apr 20 18:02:22 2019 +0000 @@ -8,33 +8,33 @@ 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) + +locale boolean_algebra = conj: abel_semigroup "(\<^bold>\)" + disj: abel_semigroup "(\<^bold>\)" + for conj :: "'a \ 'a \ 'a" (infixr "\<^bold>\" 70) + and disj :: "'a \ 'a \ 'a" (infixr "\<^bold>\" 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 = \" + assumes conj_disj_distrib: "x \<^bold>\ (y \<^bold>\ z) = (x \<^bold>\ y) \<^bold>\ (x \<^bold>\ z)" + and disj_conj_distrib: "x \<^bold>\ (y \<^bold>\ z) = (x \<^bold>\ y) \<^bold>\ (x \<^bold>\ z)" + and conj_one_right: "x \<^bold>\ \ = x" + and disj_zero_right: "x \<^bold>\ \ = x" + and conj_cancel_right [simp]: "x \<^bold>\ \ x = \" + and disj_cancel_right [simp]: "x \<^bold>\ \ x = \" begin -sublocale conj: semilattice_neutr "(\)" "\" +sublocale conj: semilattice_neutr "(\<^bold>\)" "\" proof - show "x \ \ = x" for x + show "x \<^bold>\ \ = x" for x by (fact conj_one_right) - show "x \ x = x" for x + show "x \<^bold>\ x = x" for x proof - - have "x \ x = (x \ x) \ \" + have "x \<^bold>\ x = (x \<^bold>\ x) \<^bold>\ \" by (simp add: disj_zero_right) - also have "\ = (x \ x) \ (x \ \ x)" + also have "\ = (x \<^bold>\ x) \<^bold>\ (x \<^bold>\ \ x)" by simp - also have "\ = x \ (x \ \ x)" + also have "\ = x \<^bold>\ (x \<^bold>\ \ x)" by (simp only: conj_disj_distrib) - also have "\ = x \ \" + also have "\ = x \<^bold>\ \" by simp also have "\ = x" by (simp add: conj_one_right) @@ -42,19 +42,19 @@ qed qed -sublocale disj: semilattice_neutr "(\)" "\" +sublocale disj: semilattice_neutr "(\<^bold>\)" "\" proof - show "x \ \ = x" for x + show "x \<^bold>\ \ = x" for x by (fact disj_zero_right) - show "x \ x = x" for x + show "x \<^bold>\ x = x" for x proof - - have "x \ x = (x \ x) \ \" + have "x \<^bold>\ x = (x \<^bold>\ x) \<^bold>\ \" by simp - also have "\ = (x \ x) \ (x \ \ x)" + also have "\ = (x \<^bold>\ x) \<^bold>\ (x \<^bold>\ \ x)" by simp - also have "\ = x \ (x \ \ x)" + also have "\ = x \<^bold>\ (x \<^bold>\ \ x)" by (simp only: disj_conj_distrib) - also have "\ = x \ \" + also have "\ = x \<^bold>\ \" by simp also have "\ = x" by (simp add: disj_zero_right) @@ -66,32 +66,32 @@ subsection \Complement\ lemma complement_unique: - assumes 1: "a \ x = \" - assumes 2: "a \ x = \" - assumes 3: "a \ y = \" - assumes 4: "a \ y = \" + assumes 1: "a \<^bold>\ x = \" + assumes 2: "a \<^bold>\ x = \" + assumes 3: "a \<^bold>\ y = \" + assumes 4: "a \<^bold>\ y = \" shows "x = y" proof - - from 1 3 have "(a \ x) \ (x \ y) = (a \ y) \ (x \ y)" + from 1 3 have "(a \<^bold>\ x) \<^bold>\ (x \<^bold>\ y) = (a \<^bold>\ y) \<^bold>\ (x \<^bold>\ y)" by simp - then have "(x \ a) \ (x \ y) = (y \ a) \ (y \ x)" + then have "(x \<^bold>\ a) \<^bold>\ (x \<^bold>\ y) = (y \<^bold>\ a) \<^bold>\ (y \<^bold>\ x)" by (simp add: ac_simps) - then have "x \ (a \ y) = y \ (a \ x)" + then have "x \<^bold>\ (a \<^bold>\ y) = y \<^bold>\ (a \<^bold>\ x)" by (simp add: conj_disj_distrib) - with 2 4 have "x \ \ = y \ \" + with 2 4 have "x \<^bold>\ \ = y \<^bold>\ \" by simp then show "x = y" by simp qed -lemma compl_unique: "x \ y = \ \ x \ y = \ \ \ x = y" +lemma compl_unique: "x \<^bold>\ y = \ \ x \<^bold>\ 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 = \" + show "\ x \<^bold>\ x = \" by (simp only: conj_cancel_right conj.commute) - show "\ x \ x = \" + show "\ x \<^bold>\ x = \" by (simp only: disj_cancel_right disj.commute) qed @@ -101,13 +101,13 @@ subsection \Conjunction\ -lemma conj_zero_right [simp]: "x \ \ = \" +lemma conj_zero_right [simp]: "x \<^bold>\ \ = \" proof - - from conj_cancel_right have "x \ \ = x \ (x \ \ x)" + from conj_cancel_right have "x \<^bold>\ \ = x \<^bold>\ (x \<^bold>\ \ x)" by simp - also from conj_assoc have "\ = (x \ x) \ \ x" + also from conj_assoc have "\ = (x \<^bold>\ x) \<^bold>\ \ x" by (simp only: ac_simps) - also from conj_absorb have "\ = x \ \ x" + also from conj_absorb have "\ = x \<^bold>\ \ x" by simp also have "\ = \" by simp @@ -117,39 +117,39 @@ lemma compl_one [simp]: "\ \ = \" by (rule compl_unique [OF conj_zero_right disj_zero_right]) -lemma conj_zero_left [simp]: "\ \ x = \" +lemma conj_zero_left [simp]: "\ \<^bold>\ x = \" by (subst conj.commute) (rule conj_zero_right) -lemma conj_cancel_left [simp]: "\ x \ x = \" +lemma conj_cancel_left [simp]: "\ x \<^bold>\ x = \" by (subst conj.commute) (rule conj_cancel_right) -lemma conj_disj_distrib2: "(y \ z) \ x = (y \ x) \ (z \ x)" +lemma conj_disj_distrib2: "(y \<^bold>\ z) \<^bold>\ x = (y \<^bold>\ x) \<^bold>\ (z \<^bold>\ 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)" +lemma conj_assoc: "(x \<^bold>\ y) \<^bold>\ z = x \<^bold>\ (y \<^bold>\ z)" by (fact ac_simps) -lemma conj_commute: "x \ y = y \ x" +lemma conj_commute: "x \<^bold>\ y = y \<^bold>\ 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" +lemma conj_one_left: "\ \<^bold>\ x = x" by (fact conj.left_neutral) -lemma conj_left_absorb: "x \ (x \ y) = x \ y" +lemma conj_left_absorb: "x \<^bold>\ (x \<^bold>\ y) = x \<^bold>\ y" by (fact conj.left_idem) -lemma conj_absorb: "x \ x = x" +lemma conj_absorb: "x \<^bold>\ x = x" by (fact conj.idem) subsection \Disjunction\ -interpretation dual: boolean_algebra "(\)" "(\)" compl \ \ +interpretation dual: boolean_algebra "(\<^bold>\)" "(\<^bold>\)" compl \ \ apply standard apply (rule disj_conj_distrib) apply (rule conj_disj_distrib) @@ -159,73 +159,73 @@ lemma compl_zero [simp]: "\ \ = \" by (fact dual.compl_one) -lemma disj_one_right [simp]: "x \ \ = \" +lemma disj_one_right [simp]: "x \<^bold>\ \ = \" by (fact dual.conj_zero_right) -lemma disj_one_left [simp]: "\ \ x = \" +lemma disj_one_left [simp]: "\ \<^bold>\ x = \" by (fact dual.conj_zero_left) -lemma disj_cancel_left [simp]: "\ x \ x = \" +lemma disj_cancel_left [simp]: "\ x \<^bold>\ x = \" by (rule dual.conj_cancel_left) -lemma disj_conj_distrib2: "(y \ z) \ x = (y \ x) \ (z \ x)" +lemma disj_conj_distrib2: "(y \<^bold>\ z) \<^bold>\ x = (y \<^bold>\ x) \<^bold>\ (z \<^bold>\ 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)" +lemma disj_assoc: "(x \<^bold>\ y) \<^bold>\ z = x \<^bold>\ (y \<^bold>\ z)" by (fact ac_simps) -lemma disj_commute: "x \ y = y \ x" +lemma disj_commute: "x \<^bold>\ y = y \<^bold>\ 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" +lemma disj_zero_left: "\ \<^bold>\ x = x" by (fact disj.left_neutral) -lemma disj_left_absorb: "x \ (x \ y) = x \ y" +lemma disj_left_absorb: "x \<^bold>\ (x \<^bold>\ y) = x \<^bold>\ y" by (fact disj.left_idem) -lemma disj_absorb: "x \ x = x" +lemma disj_absorb: "x \<^bold>\ x = x" by (fact disj.idem) subsection \De Morgan's Laws\ -lemma de_Morgan_conj [simp]: "\ (x \ y) = \ x \ \ y" +lemma de_Morgan_conj [simp]: "\ (x \<^bold>\ y) = \ x \<^bold>\ \ y" proof (rule compl_unique) - have "(x \ y) \ (\ x \ \ y) = ((x \ y) \ \ x) \ ((x \ y) \ \ y)" + have "(x \<^bold>\ y) \<^bold>\ (\ x \<^bold>\ \ y) = ((x \<^bold>\ y) \<^bold>\ \ x) \<^bold>\ ((x \<^bold>\ y) \<^bold>\ \ y)" by (rule conj_disj_distrib) - also have "\ = (y \ (x \ \ x)) \ (x \ (y \ \ y))" + also have "\ = (y \<^bold>\ (x \<^bold>\ \ x)) \<^bold>\ (x \<^bold>\ (y \<^bold>\ \ y))" by (simp only: conj_ac) - finally show "(x \ y) \ (\ x \ \ y) = \" + finally show "(x \<^bold>\ y) \<^bold>\ (\ x \<^bold>\ \ 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))" + have "(x \<^bold>\ y) \<^bold>\ (\ x \<^bold>\ \ y) = (x \<^bold>\ (\ x \<^bold>\ \ y)) \<^bold>\ (y \<^bold>\ (\ x \<^bold>\ \ y))" by (rule disj_conj_distrib2) - also have "\ = (\ y \ (x \ \ x)) \ (\ x \ (y \ \ y))" + also have "\ = (\ y \<^bold>\ (x \<^bold>\ \ x)) \<^bold>\ (\ x \<^bold>\ (y \<^bold>\ \ y))" by (simp only: disj_ac) - finally show "(x \ y) \ (\ x \ \ y) = \" + finally show "(x \<^bold>\ y) \<^bold>\ (\ x \<^bold>\ \ y) = \" by (simp only: disj_cancel_right disj_one_right conj_one_right) qed -lemma de_Morgan_disj [simp]: "\ (x \ y) = \ x \ \ y" +lemma de_Morgan_disj [simp]: "\ (x \<^bold>\ y) = \ x \<^bold>\ \ 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)" + where "x \ y = (x \<^bold>\ \ y) \<^bold>\ (\ x \<^bold>\ y)" sublocale xor: abel_semigroup 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)" + let ?t = "(x \<^bold>\ y \<^bold>\ z) \<^bold>\ (x \<^bold>\ \ y \<^bold>\ \ z) \<^bold>\ (\ x \<^bold>\ y \<^bold>\ \ z) \<^bold>\ (\ x \<^bold>\ \ y \<^bold>\ z)" + have "?t \<^bold>\ (z \<^bold>\ x \<^bold>\ \ x) \<^bold>\ (z \<^bold>\ y \<^bold>\ \ y) = ?t \<^bold>\ (x \<^bold>\ y \<^bold>\ \ y) \<^bold>\ (x \<^bold>\ z \<^bold>\ \ 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) @@ -240,7 +240,7 @@ lemmas xor_ac = xor.assoc xor.commute xor.left_commute -lemma xor_def2: "x \ y = (x \ y) \ (\ x \ \ y)" +lemma xor_def2: "x \ y = (x \<^bold>\ y) \<^bold>\ (\ x \<^bold>\ \ y)" by (simp only: xor_def conj_disj_distribs disj_ac conj_ac conj_cancel_right disj_zero_left) lemma xor_zero_right [simp]: "x \ \ = x" @@ -283,22 +283,22 @@ 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)" +lemma conj_xor_distrib: "x \<^bold>\ (y \ z) = (x \<^bold>\ y) \ (x \<^bold>\ z)" proof - - have *: "(x \ y \ \ z) \ (x \ \ y \ z) = - (y \ x \ \ x) \ (z \ x \ \ x) \ (x \ y \ \ z) \ (x \ \ y \ z)" + have *: "(x \<^bold>\ y \<^bold>\ \ z) \<^bold>\ (x \<^bold>\ \ y \<^bold>\ z) = + (y \<^bold>\ x \<^bold>\ \ x) \<^bold>\ (z \<^bold>\ x \<^bold>\ \ x) \<^bold>\ (x \<^bold>\ y \<^bold>\ \ z) \<^bold>\ (x \<^bold>\ \ y \<^bold>\ z)" by (simp only: conj_cancel_right conj_zero_right disj_zero_left) - then show "x \ (y \ z) = (x \ y) \ (x \ z)" + then show "x \<^bold>\ (y \ z) = (x \<^bold>\ y) \ (x \<^bold>\ 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)" +lemma conj_xor_distrib2: "(y \ z) \<^bold>\ x = (y \<^bold>\ x) \ (z \<^bold>\ x)" proof - - have "x \ (y \ z) = (x \ y) \ (x \ z)" + have "x \<^bold>\ (y \ z) = (x \<^bold>\ y) \ (x \<^bold>\ z)" by (rule conj_xor_distrib) - then show "(y \ z) \ x = (y \ x) \ (z \ x)" + then show "(y \ z) \<^bold>\ x = (y \<^bold>\ x) \ (z \<^bold>\ x)" by (simp only: conj_commute) qed