# HG changeset patch # User wenzelm # Date 1491066979 -7200 # Node ID 0a8e30a7b10e1152c4b6403df16eb687b953d3d7 # Parent e32eb488c3a3077d36956606ef190262ec3429de tuned proofs; diff -r e32eb488c3a3 -r 0a8e30a7b10e src/HOL/Library/Boolean_Algebra.thy --- a/src/HOL/Library/Boolean_Algebra.thy Sat Apr 01 18:50:26 2017 +0200 +++ b/src/HOL/Library/Boolean_Algebra.thy Sat Apr 01 19:16:19 2017 +0200 @@ -9,21 +9,21 @@ begin locale boolean = - fixes conj :: "'a \ 'a \ 'a" (infixr "\" 70) - fixes disj :: "'a \ 'a \ 'a" (infixr "\" 65) - fixes compl :: "'a \ 'a" ("\ _" [81] 80) - fixes zero :: "'a" ("\") - fixes one :: "'a" ("\") + fixes conj :: "'a \ 'a \ 'a" (infixr "\" 70) + and disj :: "'a \ 'a \ 'a" (infixr "\" 65) + and compl :: "'a \ 'a" ("\ _" [81] 80) + and zero :: "'a" ("\") + and one :: "'a" ("\") assumes conj_assoc: "(x \ y) \ z = x \ (y \ z)" - assumes disj_assoc: "(x \ y) \ z = x \ (y \ z)" - assumes conj_commute: "x \ y = y \ x" - assumes disj_commute: "x \ y = y \ x" - assumes conj_disj_distrib: "x \ (y \ z) = (x \ y) \ (x \ z)" - assumes disj_conj_distrib: "x \ (y \ z) = (x \ y) \ (x \ z)" - assumes conj_one_right [simp]: "x \ \ = x" - assumes disj_zero_right [simp]: "x \ \ = x" - assumes conj_cancel_right [simp]: "x \ \ x = \" - assumes disj_cancel_right [simp]: "x \ \ x = \" + 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)" + 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_cancel_right [simp]: "x \ \ x = \" + and disj_cancel_right [simp]: "x \ \ x = \" begin sublocale conj: abel_semigroup conj @@ -33,7 +33,6 @@ 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 @@ -41,15 +40,15 @@ lemma dual: "boolean disj conj compl one zero" apply (rule boolean.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 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 @@ -63,16 +62,16 @@ assumes 4: "a \ y = \" shows "x = y" proof - - have "(a \ x) \ (x \ y) = (a \ y) \ (x \ y)" - using 1 3 by simp + from 1 3 have "(a \ x) \ (x \ y) = (a \ y) \ (x \ y)" + by simp then have "(x \ a) \ (x \ y) = (y \ a) \ (y \ x)" - using conj_commute by simp + by (simp add: conj_commute) then have "x \ (a \ y) = y \ (a \ x)" - using conj_disj_distrib by simp - then have "x \ \ = y \ \" - using 2 4 by simp + by (simp add: conj_disj_distrib) + with 2 4 have "x \ \ = y \ \" + by simp then show "x = y" - using conj_one_right by simp + by simp qed lemma compl_unique: "x \ y = \ \ x \ y = \ \ \ x = y" @@ -80,10 +79,10 @@ lemma double_compl [simp]: "\ (\ x) = x" proof (rule compl_unique) - from conj_cancel_right show "\ x \ x = \" - by (simp only: conj_commute) - from disj_cancel_right show "\ x \ x = \" - by (simp only: disj_commute) + 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" @@ -95,28 +94,28 @@ lemma conj_absorb [simp]: "x \ x = x" proof - have "x \ x = (x \ x) \ \" - using disj_zero_right by simp - also have "... = (x \ x) \ (x \ \ x)" - using conj_cancel_right by simp - also have "... = x \ (x \ \ x)" - using conj_disj_distrib by (simp only:) - also have "... = x \ \" - using disj_cancel_right by simp - also have "... = x" - using conj_one_right by simp + 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 - - have "x \ \ = x \ (x \ \ x)" - using conj_cancel_right by simp - also have "... = (x \ x) \ \ x" - using conj_assoc by (simp only:) - also have "... = x \ \ x" - using conj_absorb by simp - also have "... = \" - using conj_cancel_right by simp + from conj_cancel_right have "x \ \ = x \ (x \ \ x)" + by simp + also from conj_assoc have "\ = (x \ x) \ \ x" + by (simp only:) + also from conj_absorb have "\ = x \ \ x" + by simp + also have "\ = \" + by simp finally show ?thesis . qed @@ -176,14 +175,14 @@ 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))" + 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))" + 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) @@ -205,15 +204,12 @@ 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 \ 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)" - apply (simp only: xor_def de_Morgan_disj de_Morgan_conj double_compl) - apply (simp only: conj_disj_distribs conj_ac disj_ac) - done + 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) qed