# HG changeset patch # User haftmann # Date 1628147569 0 # Node ID 7c5842b061141f9f81fa666a3181533a86da28f4 # Parent 7d3e818fe21f3d318697e0160a33cf19ec2f5ded clarified abstract and concrete boolean algebras diff -r 7d3e818fe21f -r 7c5842b06114 NEWS --- a/NEWS Thu Aug 05 07:12:49 2021 +0000 +++ b/NEWS Thu Aug 05 07:12:49 2021 +0000 @@ -191,11 +191,18 @@ * Theory Bit_Operations is now part of HOL-Main. Minor INCOMPATIBILITY. +* Simplified class hierarchy for bit operations: bit operations reside +in classes (semi)ring_bit_operations, class semiring_bit_shifts is +gone. + * Abbreviation "max_word" has been moved to session Word_Lib in the AFP, as also have constants "shiftl1", "shiftr1", "sshiftr1", "bshiftr1", "setBit", "clearBit". See there further the changelog in theory Guide. INCOMPATIBILITY. +* Reorganized classes and locales for boolean algebras. +INCOMPATIBILITY. + * New simp rules: less_exp, min.absorb1, min.absorb2, min.absorb3, min.absorb4, max.absorb1, max.absorb2, max.absorb3, max.absorb4. Minor INCOMPATIBILITY. diff -r 7d3e818fe21f -r 7c5842b06114 src/HOL/Analysis/Path_Connected.thy --- a/src/HOL/Analysis/Path_Connected.thy Thu Aug 05 07:12:49 2021 +0000 +++ b/src/HOL/Analysis/Path_Connected.thy Thu Aug 05 07:12:49 2021 +0000 @@ -3483,7 +3483,7 @@ fixes S :: "'a::real_normed_vector set" assumes "bounded S" shows "outside(frontier S) \ - closure S" - unfolding outside_inside Lattices.boolean_algebra_class.compl_le_compl_iff + unfolding outside_inside boolean_algebra_class.compl_le_compl_iff proof - { assume "interior S \ inside (frontier S)" hence "interior S \ inside (frontier S) = inside (frontier S)" diff -r 7d3e818fe21f -r 7c5842b06114 src/HOL/Bit_Operations.thy --- a/src/HOL/Bit_Operations.thy Thu Aug 05 07:12:49 2021 +0000 +++ b/src/HOL/Bit_Operations.thy Thu Aug 05 07:12:49 2021 +0000 @@ -1326,17 +1326,14 @@ sublocale "and": semilattice_neutr \(AND)\ \- 1\ by standard (rule bit_eqI, simp add: bit_and_iff) -sublocale bit: boolean_algebra \(AND)\ \(OR)\ NOT 0 \- 1\ - rewrites \bit.xor = (XOR)\ -proof - - interpret bit: boolean_algebra \(AND)\ \(OR)\ NOT 0 \- 1\ - by standard (auto simp add: bit_and_iff bit_or_iff bit_not_iff intro: bit_eqI) - show \boolean_algebra (AND) (OR) NOT 0 (- 1)\ - by standard - show \boolean_algebra.xor (AND) (OR) NOT = (XOR)\ - by (rule ext, rule ext, rule bit_eqI) - (auto simp add: bit.xor_def bit_and_iff bit_or_iff bit_xor_iff bit_not_iff) -qed +sublocale bit: abstract_boolean_algebra \(AND)\ \(OR)\ NOT 0 \- 1\ + by standard (auto simp add: bit_and_iff bit_or_iff bit_not_iff intro: bit_eqI) + +sublocale bit: abstract_boolean_algebra_sym_diff \(AND)\ \(OR)\ NOT 0 \- 1\ \(XOR)\ + apply standard + apply (rule bit_eqI) + apply (auto simp add: bit_simps) + done lemma and_eq_not_not_or: \a AND b = NOT (NOT a OR NOT b)\ @@ -3565,8 +3562,6 @@ \<^item> (Bounded) conversion from and to a list of bits: @{thm horner_sum_bit_eq_take_bit [where ?'a = int, no_vars]} \ -find_theorems \(_ AND _) * _ = _\ - no_notation "and" (infixr \AND\ 64) and or (infixr \OR\ 59) diff -r 7d3e818fe21f -r 7c5842b06114 src/HOL/Boolean_Algebra.thy --- a/src/HOL/Boolean_Algebra.thy Thu Aug 05 07:12:49 2021 +0000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,296 +0,0 @@ -(* Title: HOL/Boolean_Algebra.thy - Author: Brian Huffman -*) - -section \Abstract boolean Algebras\ - -theory Boolean_Algebra - imports Lattices -begin - -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" ("\<^bold>- _" [81] 80) - and zero :: "'a" ("\<^bold>0") - and one :: "'a" ("\<^bold>1") - 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>\ \<^bold>1 = x" - and disj_zero_right: "x \<^bold>\ \<^bold>0 = x" - and conj_cancel_right [simp]: "x \<^bold>\ \<^bold>- x = \<^bold>0" - and disj_cancel_right [simp]: "x \<^bold>\ \<^bold>- x = \<^bold>1" -begin - -sublocale conj: semilattice_neutr "(\<^bold>\)" "\<^bold>1" -proof - show "x \<^bold>\ \<^bold>1 = x" for x - by (fact conj_one_right) - show "x \<^bold>\ x = x" for x - proof - - have "x \<^bold>\ x = (x \<^bold>\ x) \<^bold>\ \<^bold>0" - by (simp add: disj_zero_right) - also have "\ = (x \<^bold>\ x) \<^bold>\ (x \<^bold>\ \<^bold>- x)" - by simp - also have "\ = x \<^bold>\ (x \<^bold>\ \<^bold>- x)" - by (simp only: conj_disj_distrib) - also have "\ = x \<^bold>\ \<^bold>1" - by simp - also have "\ = x" - by (simp add: conj_one_right) - finally show ?thesis . - qed -qed - -sublocale disj: semilattice_neutr "(\<^bold>\)" "\<^bold>0" -proof - show "x \<^bold>\ \<^bold>0 = x" for x - by (fact disj_zero_right) - show "x \<^bold>\ x = x" for x - proof - - have "x \<^bold>\ x = (x \<^bold>\ x) \<^bold>\ \<^bold>1" - by simp - also have "\ = (x \<^bold>\ x) \<^bold>\ (x \<^bold>\ \<^bold>- x)" - by simp - also have "\ = x \<^bold>\ (x \<^bold>\ \<^bold>- x)" - by (simp only: disj_conj_distrib) - also have "\ = x \<^bold>\ \<^bold>0" - by simp - also have "\ = x" - by (simp add: disj_zero_right) - finally show ?thesis . - qed -qed - - -subsection \Complement\ - -lemma complement_unique: - assumes 1: "a \<^bold>\ x = \<^bold>0" - assumes 2: "a \<^bold>\ x = \<^bold>1" - assumes 3: "a \<^bold>\ y = \<^bold>0" - assumes 4: "a \<^bold>\ y = \<^bold>1" - shows "x = y" -proof - - from 1 3 have "(a \<^bold>\ x) \<^bold>\ (x \<^bold>\ y) = (a \<^bold>\ y) \<^bold>\ (x \<^bold>\ y)" - by simp - then have "(x \<^bold>\ a) \<^bold>\ (x \<^bold>\ y) = (y \<^bold>\ a) \<^bold>\ (y \<^bold>\ x)" - by (simp add: ac_simps) - then have "x \<^bold>\ (a \<^bold>\ y) = y \<^bold>\ (a \<^bold>\ x)" - by (simp add: conj_disj_distrib) - with 2 4 have "x \<^bold>\ \<^bold>1 = y \<^bold>\ \<^bold>1" - by simp - then show "x = y" - by simp -qed - -lemma compl_unique: "x \<^bold>\ y = \<^bold>0 \ x \<^bold>\ y = \<^bold>1 \ \<^bold>- x = y" - by (rule complement_unique [OF conj_cancel_right disj_cancel_right]) - -lemma double_compl [simp]: "\<^bold>- (\<^bold>- x) = x" -proof (rule compl_unique) - show "\<^bold>- x \<^bold>\ x = \<^bold>0" - by (simp only: conj_cancel_right conj.commute) - show "\<^bold>- x \<^bold>\ x = \<^bold>1" - by (simp only: disj_cancel_right disj.commute) -qed - -lemma compl_eq_compl_iff [simp]: - \\<^bold>- x = \<^bold>- y \ x = y\ (is \?P \ ?Q\) -proof - assume \?Q\ - then show ?P by simp -next - assume \?P\ - then have \\<^bold>- (\<^bold>- x) = \<^bold>- (\<^bold>- y)\ - by simp - then show ?Q - by simp -qed - - -subsection \Conjunction\ - -lemma conj_zero_right [simp]: "x \<^bold>\ \<^bold>0 = \<^bold>0" - using conj.left_idem conj_cancel_right by fastforce - -lemma compl_one [simp]: "\<^bold>- \<^bold>1 = \<^bold>0" - by (rule compl_unique [OF conj_zero_right disj_zero_right]) - -lemma conj_zero_left [simp]: "\<^bold>0 \<^bold>\ x = \<^bold>0" - by (subst conj.commute) (rule conj_zero_right) - -lemma conj_cancel_left [simp]: "\<^bold>- x \<^bold>\ x = \<^bold>0" - by (subst conj.commute) (rule conj_cancel_right) - -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 \<^bold>\ y) \<^bold>\ z = x \<^bold>\ (y \<^bold>\ z)" - by (fact ac_simps) - -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: "\<^bold>1 \<^bold>\ x = x" - by (fact conj.left_neutral) - -lemma conj_left_absorb: "x \<^bold>\ (x \<^bold>\ y) = x \<^bold>\ y" - by (fact conj.left_idem) - -lemma conj_absorb: "x \<^bold>\ x = x" - by (fact conj.idem) - - -subsection \Disjunction\ - -interpretation dual: boolean_algebra "(\<^bold>\)" "(\<^bold>\)" compl \\<^bold>1\ \\<^bold>0\ - apply standard - apply (rule disj_conj_distrib) - apply (rule conj_disj_distrib) - apply simp_all - done - -lemma compl_zero [simp]: "\<^bold>- \<^bold>0 = \<^bold>1" - by (fact dual.compl_one) - -lemma disj_one_right [simp]: "x \<^bold>\ \<^bold>1 = \<^bold>1" - by (fact dual.conj_zero_right) - -lemma disj_one_left [simp]: "\<^bold>1 \<^bold>\ x = \<^bold>1" - by (fact dual.conj_zero_left) - -lemma disj_cancel_left [simp]: "\<^bold>- x \<^bold>\ x = \<^bold>1" - by (rule dual.conj_cancel_left) - -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 \<^bold>\ y) \<^bold>\ z = x \<^bold>\ (y \<^bold>\ z)" - by (fact ac_simps) - -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: "\<^bold>0 \<^bold>\ x = x" - by (fact disj.left_neutral) - -lemma disj_left_absorb: "x \<^bold>\ (x \<^bold>\ y) = x \<^bold>\ y" - by (fact disj.left_idem) - -lemma disj_absorb: "x \<^bold>\ x = x" - by (fact disj.idem) - - -subsection \De Morgan's Laws\ - -lemma de_Morgan_conj [simp]: "\<^bold>- (x \<^bold>\ y) = \<^bold>- x \<^bold>\ \<^bold>- y" -proof (rule compl_unique) - have "(x \<^bold>\ y) \<^bold>\ (\<^bold>- x \<^bold>\ \<^bold>- y) = ((x \<^bold>\ y) \<^bold>\ \<^bold>- x) \<^bold>\ ((x \<^bold>\ y) \<^bold>\ \<^bold>- y)" - by (rule conj_disj_distrib) - also have "\ = (y \<^bold>\ (x \<^bold>\ \<^bold>- x)) \<^bold>\ (x \<^bold>\ (y \<^bold>\ \<^bold>- y))" - by (simp only: conj_ac) - finally show "(x \<^bold>\ y) \<^bold>\ (\<^bold>- x \<^bold>\ \<^bold>- y) = \<^bold>0" - by (simp only: conj_cancel_right conj_zero_right disj_zero_right) -next - have "(x \<^bold>\ y) \<^bold>\ (\<^bold>- x \<^bold>\ \<^bold>- y) = (x \<^bold>\ (\<^bold>- x \<^bold>\ \<^bold>- y)) \<^bold>\ (y \<^bold>\ (\<^bold>- x \<^bold>\ \<^bold>- y))" - by (rule disj_conj_distrib2) - also have "\ = (\<^bold>- y \<^bold>\ (x \<^bold>\ \<^bold>- x)) \<^bold>\ (\<^bold>- x \<^bold>\ (y \<^bold>\ \<^bold>- y))" - by (simp only: disj_ac) - finally show "(x \<^bold>\ y) \<^bold>\ (\<^bold>- x \<^bold>\ \<^bold>- y) = \<^bold>1" - by (simp only: disj_cancel_right disj_one_right conj_one_right) -qed - -lemma de_Morgan_disj [simp]: "\<^bold>- (x \<^bold>\ y) = \<^bold>- x \<^bold>\ \<^bold>- y" - using dual.boolean_algebra_axioms by (rule boolean_algebra.de_Morgan_conj) - - -subsection \Symmetric Difference\ - -definition xor :: "'a \ 'a \ 'a" (infixr "\<^bold>\" 65) - where "x \<^bold>\ y = (x \<^bold>\ \<^bold>- y) \<^bold>\ (\<^bold>- x \<^bold>\ y)" - -sublocale xor: comm_monoid xor \\<^bold>0\ -proof - fix x y z :: 'a - let ?t = "(x \<^bold>\ y \<^bold>\ z) \<^bold>\ (x \<^bold>\ \<^bold>- y \<^bold>\ \<^bold>- z) \<^bold>\ (\<^bold>- x \<^bold>\ y \<^bold>\ \<^bold>- z) \<^bold>\ (\<^bold>- x \<^bold>\ \<^bold>- y \<^bold>\ z)" - have "?t \<^bold>\ (z \<^bold>\ x \<^bold>\ \<^bold>- x) \<^bold>\ (z \<^bold>\ y \<^bold>\ \<^bold>- y) = ?t \<^bold>\ (x \<^bold>\ y \<^bold>\ \<^bold>- y) \<^bold>\ (x \<^bold>\ z \<^bold>\ \<^bold>- z)" - by (simp only: conj_cancel_right conj_zero_right) - then show "(x \<^bold>\ y) \<^bold>\ z = x \<^bold>\ (y \<^bold>\ 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 \<^bold>\ y = y \<^bold>\ x" - by (simp only: xor_def conj_commute disj_commute) - show "x \<^bold>\ \<^bold>0 = 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 \<^bold>\ y = (x \<^bold>\ y) \<^bold>\ (\<^bold>- x \<^bold>\ \<^bold>- y)" - using conj.commute conj_disj_distrib2 disj.commute xor_def by auto - -lemma xor_zero_right: "x \<^bold>\ \<^bold>0 = x" - by (fact xor.comm_neutral) - -lemma xor_zero_left: "\<^bold>0 \<^bold>\ x = x" - by (fact xor.left_neutral) - -lemma xor_one_right [simp]: "x \<^bold>\ \<^bold>1 = \<^bold>- x" - by (simp only: xor_def compl_one conj_zero_right conj_one_right disj_zero_left) - -lemma xor_one_left [simp]: "\<^bold>1 \<^bold>\ x = \<^bold>- x" - by (subst xor_commute) (rule xor_one_right) - -lemma xor_self [simp]: "x \<^bold>\ x = \<^bold>0" - by (simp only: xor_def conj_cancel_right conj_cancel_left disj_zero_right) - -lemma xor_left_self [simp]: "x \<^bold>\ (x \<^bold>\ y) = y" - by (simp only: xor_assoc [symmetric] xor_self xor_zero_left) - -lemma xor_compl_left [simp]: "\<^bold>- x \<^bold>\ y = \<^bold>- (x \<^bold>\ y)" - by (simp add: ac_simps flip: xor_one_left) - -lemma xor_compl_right [simp]: "x \<^bold>\ \<^bold>- y = \<^bold>- (x \<^bold>\ y)" - using xor_commute xor_compl_left by auto - -lemma xor_cancel_right: "x \<^bold>\ \<^bold>- x = \<^bold>1" - by (simp only: xor_compl_right xor_self compl_zero) - -lemma xor_cancel_left: "\<^bold>- x \<^bold>\ x = \<^bold>1" - by (simp only: xor_compl_left xor_self compl_zero) - -lemma conj_xor_distrib: "x \<^bold>\ (y \<^bold>\ z) = (x \<^bold>\ y) \<^bold>\ (x \<^bold>\ z)" -proof - - have *: "(x \<^bold>\ y \<^bold>\ \<^bold>- z) \<^bold>\ (x \<^bold>\ \<^bold>- y \<^bold>\ z) = - (y \<^bold>\ x \<^bold>\ \<^bold>- x) \<^bold>\ (z \<^bold>\ x \<^bold>\ \<^bold>- x) \<^bold>\ (x \<^bold>\ y \<^bold>\ \<^bold>- z) \<^bold>\ (x \<^bold>\ \<^bold>- y \<^bold>\ z)" - by (simp only: conj_cancel_right conj_zero_right disj_zero_left) - then show "x \<^bold>\ (y \<^bold>\ z) = (x \<^bold>\ y) \<^bold>\ (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 \<^bold>\ z) \<^bold>\ x = (y \<^bold>\ x) \<^bold>\ (z \<^bold>\ x)" - by (simp add: conj.commute conj_xor_distrib) - -lemmas conj_xor_distribs = conj_xor_distrib conj_xor_distrib2 - -end - -end diff -r 7d3e818fe21f -r 7c5842b06114 src/HOL/Boolean_Algebras.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Boolean_Algebras.thy Thu Aug 05 07:12:49 2021 +0000 @@ -0,0 +1,573 @@ +(* Title: HOL/Boolean_Algebras.thy + Author: Brian Huffman + Author: Florian Haftmann +*) + +section \Boolean Algebras\ + +theory Boolean_Algebras + imports Lattices +begin + +subsection \Abstract boolean algebra\ + +locale abstract_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\ (\\<^bold>- _\ [81] 80) + and zero :: \'a\ (\\<^bold>0\) + and one :: \'a\ (\\<^bold>1\) + 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>\ \<^bold>1 = x\ + and disj_zero_right: \x \<^bold>\ \<^bold>0 = x\ + and conj_cancel_right [simp]: \x \<^bold>\ \<^bold>- x = \<^bold>0\ + and disj_cancel_right [simp]: \x \<^bold>\ \<^bold>- x = \<^bold>1\ +begin + +sublocale conj: semilattice_neutr \(\<^bold>\)\ \\<^bold>1\ +proof + show "x \<^bold>\ \<^bold>1 = x" for x + by (fact conj_one_right) + show "x \<^bold>\ x = x" for x + proof - + have "x \<^bold>\ x = (x \<^bold>\ x) \<^bold>\ \<^bold>0" + by (simp add: disj_zero_right) + also have "\ = (x \<^bold>\ x) \<^bold>\ (x \<^bold>\ \<^bold>- x)" + by simp + also have "\ = x \<^bold>\ (x \<^bold>\ \<^bold>- x)" + by (simp only: conj_disj_distrib) + also have "\ = x \<^bold>\ \<^bold>1" + by simp + also have "\ = x" + by (simp add: conj_one_right) + finally show ?thesis . + qed +qed + +sublocale disj: semilattice_neutr \(\<^bold>\)\ \\<^bold>0\ +proof + show "x \<^bold>\ \<^bold>0 = x" for x + by (fact disj_zero_right) + show "x \<^bold>\ x = x" for x + proof - + have "x \<^bold>\ x = (x \<^bold>\ x) \<^bold>\ \<^bold>1" + by simp + also have "\ = (x \<^bold>\ x) \<^bold>\ (x \<^bold>\ \<^bold>- x)" + by simp + also have "\ = x \<^bold>\ (x \<^bold>\ \<^bold>- x)" + by (simp only: disj_conj_distrib) + also have "\ = x \<^bold>\ \<^bold>0" + by simp + also have "\ = x" + by (simp add: disj_zero_right) + finally show ?thesis . + qed +qed + + +subsubsection \Complement\ + +lemma complement_unique: + assumes 1: "a \<^bold>\ x = \<^bold>0" + assumes 2: "a \<^bold>\ x = \<^bold>1" + assumes 3: "a \<^bold>\ y = \<^bold>0" + assumes 4: "a \<^bold>\ y = \<^bold>1" + shows "x = y" +proof - + from 1 3 have "(a \<^bold>\ x) \<^bold>\ (x \<^bold>\ y) = (a \<^bold>\ y) \<^bold>\ (x \<^bold>\ y)" + by simp + then have "(x \<^bold>\ a) \<^bold>\ (x \<^bold>\ y) = (y \<^bold>\ a) \<^bold>\ (y \<^bold>\ x)" + by (simp add: ac_simps) + then have "x \<^bold>\ (a \<^bold>\ y) = y \<^bold>\ (a \<^bold>\ x)" + by (simp add: conj_disj_distrib) + with 2 4 have "x \<^bold>\ \<^bold>1 = y \<^bold>\ \<^bold>1" + by simp + then show "x = y" + by simp +qed + +lemma compl_unique: "x \<^bold>\ y = \<^bold>0 \ x \<^bold>\ y = \<^bold>1 \ \<^bold>- x = y" + by (rule complement_unique [OF conj_cancel_right disj_cancel_right]) + +lemma double_compl [simp]: "\<^bold>- (\<^bold>- x) = x" +proof (rule compl_unique) + show "\<^bold>- x \<^bold>\ x = \<^bold>0" + by (simp only: conj_cancel_right conj.commute) + show "\<^bold>- x \<^bold>\ x = \<^bold>1" + by (simp only: disj_cancel_right disj.commute) +qed + +lemma compl_eq_compl_iff [simp]: + \\<^bold>- x = \<^bold>- y \ x = y\ (is \?P \ ?Q\) +proof + assume \?Q\ + then show ?P by simp +next + assume \?P\ + then have \\<^bold>- (\<^bold>- x) = \<^bold>- (\<^bold>- y)\ + by simp + then show ?Q + by simp +qed + + +subsubsection \Conjunction\ + +lemma conj_zero_right [simp]: "x \<^bold>\ \<^bold>0 = \<^bold>0" + using conj.left_idem conj_cancel_right by fastforce + +lemma compl_one [simp]: "\<^bold>- \<^bold>1 = \<^bold>0" + by (rule compl_unique [OF conj_zero_right disj_zero_right]) + +lemma conj_zero_left [simp]: "\<^bold>0 \<^bold>\ x = \<^bold>0" + by (subst conj.commute) (rule conj_zero_right) + +lemma conj_cancel_left [simp]: "\<^bold>- x \<^bold>\ x = \<^bold>0" + by (subst conj.commute) (rule conj_cancel_right) + +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 + + +subsubsection \Disjunction\ + +context +begin + +interpretation dual: abstract_boolean_algebra \(\<^bold>\)\ \(\<^bold>\)\ compl \\<^bold>1\ \\<^bold>0\ + apply standard + apply (rule disj_conj_distrib) + apply (rule conj_disj_distrib) + apply simp_all + done + +lemma disj_one_right [simp]: "x \<^bold>\ \<^bold>1 = \<^bold>1" + by (fact dual.conj_zero_right) + +lemma compl_zero [simp]: "\<^bold>- \<^bold>0 = \<^bold>1" + by (fact dual.compl_one) + +lemma disj_one_left [simp]: "\<^bold>1 \<^bold>\ x = \<^bold>1" + by (fact dual.conj_zero_left) + +lemma disj_cancel_left [simp]: "\<^bold>- x \<^bold>\ x = \<^bold>1" + by (fact dual.conj_cancel_left) + +lemma disj_conj_distrib2: "(y \<^bold>\ z) \<^bold>\ x = (y \<^bold>\ x) \<^bold>\ (z \<^bold>\ x)" + by (fact dual.conj_disj_distrib2) + +lemmas disj_conj_distribs = disj_conj_distrib disj_conj_distrib2 + +end + + +subsubsection \De Morgan's Laws\ + +lemma de_Morgan_conj [simp]: "\<^bold>- (x \<^bold>\ y) = \<^bold>- x \<^bold>\ \<^bold>- y" +proof (rule compl_unique) + have "(x \<^bold>\ y) \<^bold>\ (\<^bold>- x \<^bold>\ \<^bold>- y) = ((x \<^bold>\ y) \<^bold>\ \<^bold>- x) \<^bold>\ ((x \<^bold>\ y) \<^bold>\ \<^bold>- y)" + by (rule conj_disj_distrib) + also have "\ = (y \<^bold>\ (x \<^bold>\ \<^bold>- x)) \<^bold>\ (x \<^bold>\ (y \<^bold>\ \<^bold>- y))" + by (simp only: ac_simps) + finally show "(x \<^bold>\ y) \<^bold>\ (\<^bold>- x \<^bold>\ \<^bold>- y) = \<^bold>0" + by (simp only: conj_cancel_right conj_zero_right disj_zero_right) +next + have "(x \<^bold>\ y) \<^bold>\ (\<^bold>- x \<^bold>\ \<^bold>- y) = (x \<^bold>\ (\<^bold>- x \<^bold>\ \<^bold>- y)) \<^bold>\ (y \<^bold>\ (\<^bold>- x \<^bold>\ \<^bold>- y))" + by (rule disj_conj_distrib2) + also have "\ = (\<^bold>- y \<^bold>\ (x \<^bold>\ \<^bold>- x)) \<^bold>\ (\<^bold>- x \<^bold>\ (y \<^bold>\ \<^bold>- y))" + by (simp only: ac_simps) + finally show "(x \<^bold>\ y) \<^bold>\ (\<^bold>- x \<^bold>\ \<^bold>- y) = \<^bold>1" + by (simp only: disj_cancel_right disj_one_right conj_one_right) +qed + +context +begin + +interpretation dual: abstract_boolean_algebra \(\<^bold>\)\ \(\<^bold>\)\ compl \\<^bold>1\ \\<^bold>0\ + apply standard + apply (rule disj_conj_distrib) + apply (rule conj_disj_distrib) + apply simp_all + done + +lemma de_Morgan_disj [simp]: "\<^bold>- (x \<^bold>\ y) = \<^bold>- x \<^bold>\ \<^bold>- y" + by (fact dual.de_Morgan_conj) + +end + +end + + +subsection \Symmetric Difference\ + +locale abstract_boolean_algebra_sym_diff = abstract_boolean_algebra + + fixes xor :: \'a \ 'a \ 'a\ (infixr \\<^bold>\\ 65) + assumes xor_def : \x \<^bold>\ y = (x \<^bold>\ \<^bold>- y) \<^bold>\ (\<^bold>- x \<^bold>\ y)\ +begin + +sublocale xor: comm_monoid xor \\<^bold>0\ +proof + fix x y z :: 'a + let ?t = "(x \<^bold>\ y \<^bold>\ z) \<^bold>\ (x \<^bold>\ \<^bold>- y \<^bold>\ \<^bold>- z) \<^bold>\ (\<^bold>- x \<^bold>\ y \<^bold>\ \<^bold>- z) \<^bold>\ (\<^bold>- x \<^bold>\ \<^bold>- y \<^bold>\ z)" + have "?t \<^bold>\ (z \<^bold>\ x \<^bold>\ \<^bold>- x) \<^bold>\ (z \<^bold>\ y \<^bold>\ \<^bold>- y) = ?t \<^bold>\ (x \<^bold>\ y \<^bold>\ \<^bold>- y) \<^bold>\ (x \<^bold>\ z \<^bold>\ \<^bold>- z)" + by (simp only: conj_cancel_right conj_zero_right) + then show "(x \<^bold>\ y) \<^bold>\ z = x \<^bold>\ (y \<^bold>\ z)" + by (simp only: xor_def de_Morgan_disj de_Morgan_conj double_compl) + (simp only: conj_disj_distribs conj_ac ac_simps) + show "x \<^bold>\ y = y \<^bold>\ x" + by (simp only: xor_def ac_simps) + show "x \<^bold>\ \<^bold>0 = x" + by (simp add: xor_def) +qed + +lemma xor_def2: + \x \<^bold>\ y = (x \<^bold>\ y) \<^bold>\ (\<^bold>- x \<^bold>\ \<^bold>- y)\ +proof - + note xor_def [of x y] + also have \x \<^bold>\ \<^bold>- y \<^bold>\ \<^bold>- x \<^bold>\ y = ((x \<^bold>\ \<^bold>- x) \<^bold>\ (\<^bold>- y \<^bold>\ \<^bold>- x)) \<^bold>\ (x \<^bold>\ y) \<^bold>\ (\<^bold>- y \<^bold>\ y)\ + by (simp add: ac_simps disj_conj_distribs) + also have \\ = (x \<^bold>\ y) \<^bold>\ (\<^bold>- x \<^bold>\ \<^bold>- y)\ + by (simp add: ac_simps) + finally show ?thesis . +qed + +lemma xor_one_right [simp]: "x \<^bold>\ \<^bold>1 = \<^bold>- x" + by (simp only: xor_def compl_one conj_zero_right conj_one_right disj.left_neutral) + +lemma xor_one_left [simp]: "\<^bold>1 \<^bold>\ x = \<^bold>- x" + using xor_one_right [of x] by (simp add: ac_simps) + +lemma xor_self [simp]: "x \<^bold>\ x = \<^bold>0" + by (simp only: xor_def conj_cancel_right conj_cancel_left disj_zero_right) + +lemma xor_left_self [simp]: "x \<^bold>\ (x \<^bold>\ y) = y" + by (simp only: xor.assoc [symmetric] xor_self xor.left_neutral) + +lemma xor_compl_left [simp]: "\<^bold>- x \<^bold>\ y = \<^bold>- (x \<^bold>\ y)" + by (simp add: ac_simps flip: xor_one_left) + +lemma xor_compl_right [simp]: "x \<^bold>\ \<^bold>- y = \<^bold>- (x \<^bold>\ y)" + using xor.commute xor_compl_left by auto + +lemma xor_cancel_right [simp]: "x \<^bold>\ \<^bold>- x = \<^bold>1" + by (simp only: xor_compl_right xor_self compl_zero) + +lemma xor_cancel_left [simp]: "\<^bold>- x \<^bold>\ x = \<^bold>1" + by (simp only: xor_compl_left xor_self compl_zero) + +lemma conj_xor_distrib: "x \<^bold>\ (y \<^bold>\ z) = (x \<^bold>\ y) \<^bold>\ (x \<^bold>\ z)" +proof - + have *: "(x \<^bold>\ y \<^bold>\ \<^bold>- z) \<^bold>\ (x \<^bold>\ \<^bold>- y \<^bold>\ z) = + (y \<^bold>\ x \<^bold>\ \<^bold>- x) \<^bold>\ (z \<^bold>\ x \<^bold>\ \<^bold>- x) \<^bold>\ (x \<^bold>\ y \<^bold>\ \<^bold>- z) \<^bold>\ (x \<^bold>\ \<^bold>- y \<^bold>\ z)" + by (simp only: conj_cancel_right conj_zero_right disj.left_neutral) + then show "x \<^bold>\ (y \<^bold>\ z) = (x \<^bold>\ y) \<^bold>\ (x \<^bold>\ z)" + by (simp (no_asm_use) only: + xor_def de_Morgan_disj de_Morgan_conj double_compl + conj_disj_distribs ac_simps) +qed + +lemma conj_xor_distrib2: "(y \<^bold>\ z) \<^bold>\ x = (y \<^bold>\ x) \<^bold>\ (z \<^bold>\ x)" + by (simp add: conj.commute conj_xor_distrib) + +lemmas conj_xor_distribs = conj_xor_distrib conj_xor_distrib2 + +end + + +subsection \Type classes\ + +class boolean_algebra = distrib_lattice + bounded_lattice + minus + uminus + + assumes inf_compl_bot: \x \ - x = \\ + and sup_compl_top: \x \ - x = \\ + assumes diff_eq: \x - y = x \ - y\ +begin + +sublocale boolean_algebra: abstract_boolean_algebra \(\)\ \(\)\ uminus \ \ + apply standard + apply (rule inf_sup_distrib1) + apply (rule sup_inf_distrib1) + apply (simp_all add: ac_simps inf_compl_bot sup_compl_top) + done + +lemma compl_inf_bot: "- x \ x = \" + by (fact boolean_algebra.conj_cancel_left) + +lemma compl_sup_top: "- x \ x = \" + by (fact boolean_algebra.disj_cancel_left) + +lemma compl_unique: + assumes "x \ y = \" + and "x \ y = \" + shows "- x = y" + using assms by (rule boolean_algebra.compl_unique) + +lemma double_compl: "- (- x) = x" + by (fact boolean_algebra.double_compl) + +lemma compl_eq_compl_iff: "- x = - y \ x = y" + by (fact boolean_algebra.compl_eq_compl_iff) + +lemma compl_bot_eq: "- \ = \" + by (fact boolean_algebra.compl_zero) + +lemma compl_top_eq: "- \ = \" + by (fact boolean_algebra.compl_one) + +lemma compl_inf: "- (x \ y) = - x \ - y" + by (fact boolean_algebra.de_Morgan_conj) + +lemma compl_sup: "- (x \ y) = - x \ - y" + by (fact boolean_algebra.de_Morgan_disj) + +lemma compl_mono: + assumes "x \ y" + shows "- y \ - x" +proof - + from assms have "x \ y = y" by (simp only: le_iff_sup) + then have "- (x \ y) = - y" by simp + then have "- x \ - y = - y" by simp + then have "- y \ - x = - y" by (simp only: inf_commute) + then show ?thesis by (simp only: le_iff_inf) +qed + +lemma compl_le_compl_iff [simp]: "- x \ - y \ y \ x" + by (auto dest: compl_mono) + +lemma compl_le_swap1: + assumes "y \ - x" + shows "x \ -y" +proof - + from assms have "- (- x) \ - y" by (simp only: compl_le_compl_iff) + then show ?thesis by simp +qed + +lemma compl_le_swap2: + assumes "- y \ x" + shows "- x \ y" +proof - + from assms have "- x \ - (- y)" by (simp only: compl_le_compl_iff) + then show ?thesis by simp +qed + +lemma compl_less_compl_iff [simp]: "- x < - y \ y < x" + by (auto simp add: less_le) + +lemma compl_less_swap1: + assumes "y < - x" + shows "x < - y" +proof - + from assms have "- (- x) < - y" by (simp only: compl_less_compl_iff) + then show ?thesis by simp +qed + +lemma compl_less_swap2: + assumes "- y < x" + shows "- x < y" +proof - + from assms have "- x < - (- y)" + by (simp only: compl_less_compl_iff) + then show ?thesis by simp +qed + +lemma sup_cancel_left1: \x \ a \ (- x \ b) = \\ + by (simp add: ac_simps) + +lemma sup_cancel_left2: \- x \ a \ (x \ b) = \\ + by (simp add: ac_simps) + +lemma inf_cancel_left1: \x \ a \ (- x \ b) = \\ + by (simp add: ac_simps) + +lemma inf_cancel_left2: \- x \ a \ (x \ b) = \\ + by (simp add: ac_simps) + +lemma sup_compl_top_left1 [simp]: \- x \ (x \ y) = \\ + by (simp add: sup_assoc [symmetric]) + +lemma sup_compl_top_left2 [simp]: \x \ (- x \ y) = \\ + using sup_compl_top_left1 [of "- x" y] by simp + +lemma inf_compl_bot_left1 [simp]: \- x \ (x \ y) = \\ + by (simp add: inf_assoc [symmetric]) + +lemma inf_compl_bot_left2 [simp]: \x \ (- x \ y) = \\ + using inf_compl_bot_left1 [of "- x" y] by simp + +lemma inf_compl_bot_right [simp]: \x \ (y \ - x) = \\ + by (subst inf_left_commute) simp + +end + + +subsection \Lattice on \<^typ>\bool\\ + +instantiation bool :: boolean_algebra +begin + +definition bool_Compl_def [simp]: "uminus = Not" + +definition bool_diff_def [simp]: "A - B \ A \ \ B" + +definition [simp]: "P \ Q \ P \ Q" + +definition [simp]: "P \ Q \ P \ Q" + +instance by standard auto + +end + +lemma sup_boolI1: "P \ P \ Q" + by simp + +lemma sup_boolI2: "Q \ P \ Q" + by simp + +lemma sup_boolE: "P \ Q \ (P \ R) \ (Q \ R) \ R" + by auto + +instance "fun" :: (type, boolean_algebra) boolean_algebra + by standard (rule ext, simp_all add: inf_compl_bot sup_compl_top diff_eq)+ + + +subsection \Lattice on unary and binary predicates\ + +lemma inf1I: "A x \ B x \ (A \ B) x" + by (simp add: inf_fun_def) + +lemma inf2I: "A x y \ B x y \ (A \ B) x y" + by (simp add: inf_fun_def) + +lemma inf1E: "(A \ B) x \ (A x \ B x \ P) \ P" + by (simp add: inf_fun_def) + +lemma inf2E: "(A \ B) x y \ (A x y \ B x y \ P) \ P" + by (simp add: inf_fun_def) + +lemma inf1D1: "(A \ B) x \ A x" + by (rule inf1E) + +lemma inf2D1: "(A \ B) x y \ A x y" + by (rule inf2E) + +lemma inf1D2: "(A \ B) x \ B x" + by (rule inf1E) + +lemma inf2D2: "(A \ B) x y \ B x y" + by (rule inf2E) + +lemma sup1I1: "A x \ (A \ B) x" + by (simp add: sup_fun_def) + +lemma sup2I1: "A x y \ (A \ B) x y" + by (simp add: sup_fun_def) + +lemma sup1I2: "B x \ (A \ B) x" + by (simp add: sup_fun_def) + +lemma sup2I2: "B x y \ (A \ B) x y" + by (simp add: sup_fun_def) + +lemma sup1E: "(A \ B) x \ (A x \ P) \ (B x \ P) \ P" + by (simp add: sup_fun_def) iprover + +lemma sup2E: "(A \ B) x y \ (A x y \ P) \ (B x y \ P) \ P" + by (simp add: sup_fun_def) iprover + +text \ \<^medskip> Classical introduction rule: no commitment to \A\ vs \B\.\ + +lemma sup1CI: "(\ B x \ A x) \ (A \ B) x" + by (auto simp add: sup_fun_def) + +lemma sup2CI: "(\ B x y \ A x y) \ (A \ B) x y" + by (auto simp add: sup_fun_def) + + +subsection \Simproc setup\ + +locale boolean_algebra_cancel +begin + +lemma sup1: "(A::'a::semilattice_sup) \ sup k a \ sup A b \ sup k (sup a b)" + by (simp only: ac_simps) + +lemma sup2: "(B::'a::semilattice_sup) \ sup k b \ sup a B \ sup k (sup a b)" + by (simp only: ac_simps) + +lemma sup0: "(a::'a::bounded_semilattice_sup_bot) \ sup a bot" + by simp + +lemma inf1: "(A::'a::semilattice_inf) \ inf k a \ inf A b \ inf k (inf a b)" + by (simp only: ac_simps) + +lemma inf2: "(B::'a::semilattice_inf) \ inf k b \ inf a B \ inf k (inf a b)" + by (simp only: ac_simps) + +lemma inf0: "(a::'a::bounded_semilattice_inf_top) \ inf a top" + by simp + +end + +ML_file \Tools/boolean_algebra_cancel.ML\ + +simproc_setup boolean_algebra_cancel_sup ("sup a b::'a::boolean_algebra") = + \fn phi => fn ss => try Boolean_Algebra_Cancel.cancel_sup_conv\ + +simproc_setup boolean_algebra_cancel_inf ("inf a b::'a::boolean_algebra") = + \fn phi => fn ss => try Boolean_Algebra_Cancel.cancel_inf_conv\ + + +context boolean_algebra +begin + +lemma shunt1: "(x \ y \ z) \ (x \ -y \ z)" +proof + assume "x \ y \ z" + hence "-y \ (x \ y) \ -y \ z" + using sup.mono by blast + hence "-y \ x \ -y \ z" + by (simp add: sup_inf_distrib1) + thus "x \ -y \ z" + by simp +next + assume "x \ -y \ z" + hence "x \ y \ (-y \ z) \ y" + using inf_mono by auto + thus "x \ y \ z" + using inf.boundedE inf_sup_distrib2 by auto +qed + +lemma shunt2: "(x \ -y \ z) \ (x \ y \ z)" + by (simp add: shunt1) + +lemma inf_shunt: "(x \ y = \) \ (x \ - y)" + by (simp add: order.eq_iff shunt1) + +lemma sup_shunt: "(x \ y = \) \ (- x \ y)" + using inf_shunt [of \- x\ \- y\, symmetric] + by (simp flip: compl_sup compl_top_eq) + +lemma diff_shunt_var: "(x - y = \) \ (x \ y)" + by (simp add: diff_eq inf_shunt) + +lemma sup_neg_inf: + \p \ q \ r \ p \ -q \ r\ (is \?P \ ?Q\) +proof + assume ?P + then have \p \ - q \ (q \ r) \ - q\ + by (rule inf_mono) simp + then show ?Q + by (simp add: inf_sup_distrib2) +next + assume ?Q + then have \p \ - q \ q \ r \ q\ + by (rule sup_mono) simp + then show ?P + by (simp add: sup_inf_distrib ac_simps) +qed + +end + +end diff -r 7d3e818fe21f -r 7c5842b06114 src/HOL/Fun.thy --- a/src/HOL/Fun.thy Thu Aug 05 07:12:49 2021 +0000 +++ b/src/HOL/Fun.thy Thu Aug 05 07:12:49 2021 +0000 @@ -93,7 +93,7 @@ lemma (in group_add) minus_comp_minus [simp]: "uminus \ uminus = id" by (simp add: fun_eq_iff) -lemma (in Lattices.boolean_algebra) minus_comp_minus [simp]: "uminus \ uminus = id" +lemma (in boolean_algebra) minus_comp_minus [simp]: "uminus \ uminus = id" by (simp add: fun_eq_iff) code_printing diff -r 7d3e818fe21f -r 7c5842b06114 src/HOL/Hilbert_Choice.thy --- a/src/HOL/Hilbert_Choice.thy Thu Aug 05 07:12:49 2021 +0000 +++ b/src/HOL/Hilbert_Choice.thy Thu Aug 05 07:12:49 2021 +0000 @@ -1170,18 +1170,6 @@ end -context complete_boolean_algebra -begin - -lemma dual_complete_boolean_algebra: - "class.complete_boolean_algebra Sup Inf sup (\) (>) inf \ \ (\x y. x \ - y) uminus" - by (rule class.complete_boolean_algebra.intro, - rule dual_complete_distrib_lattice, - rule dual_boolean_algebra) -end - - - instantiation set :: (type) complete_distrib_lattice begin instance proof (standard, clarsimp) diff -r 7d3e818fe21f -r 7c5842b06114 src/HOL/Lattices.thy --- a/src/HOL/Lattices.thy Thu Aug 05 07:12:49 2021 +0000 +++ b/src/HOL/Lattices.thy Thu Aug 05 07:12:49 2021 +0000 @@ -462,7 +462,7 @@ end -subsection \Bounded lattices and boolean algebras\ +subsection \Bounded lattices\ class bounded_semilattice_inf_top = semilattice_inf + order_top begin @@ -549,196 +549,6 @@ end -class boolean_algebra = distrib_lattice + bounded_lattice + minus + uminus + - assumes inf_compl_bot: "x \ - x = \" - and sup_compl_top: "x \ - x = \" - assumes diff_eq: "x - y = x \ - y" -begin - -lemma dual_boolean_algebra: - "class.boolean_algebra (\x y. x \ - y) uminus sup greater_eq greater inf \ \" - by (rule class.boolean_algebra.intro, - rule dual_bounded_lattice, - rule dual_distrib_lattice) - (unfold_locales, auto simp add: inf_compl_bot sup_compl_top diff_eq) - -lemma compl_inf_bot [simp]: "- x \ x = \" - by (simp add: inf_commute inf_compl_bot) - -lemma compl_sup_top [simp]: "- x \ x = \" - by (simp add: sup_commute sup_compl_top) - -lemma compl_unique: - assumes "x \ y = \" - and "x \ y = \" - shows "- x = y" -proof - - have "(x \ - x) \ (- x \ y) = (x \ y) \ (- x \ y)" - using inf_compl_bot assms(1) by simp - then have "(- x \ x) \ (- x \ y) = (y \ x) \ (y \ - x)" - by (simp add: inf_commute) - then have "- x \ (x \ y) = y \ (x \ - x)" - by (simp add: inf_sup_distrib1) - then have "- x \ \ = y \ \" - using sup_compl_top assms(2) by simp - then show "- x = y" by simp -qed - -lemma double_compl [simp]: "- (- x) = x" - using compl_inf_bot compl_sup_top by (rule compl_unique) - -lemma compl_eq_compl_iff [simp]: "- x = - y \ x = y" -proof - assume "- x = - y" - then have "- (- x) = - (- y)" by (rule arg_cong) - then show "x = y" by simp -next - assume "x = y" - then show "- x = - y" by simp -qed - -lemma compl_bot_eq [simp]: "- \ = \" -proof - - from sup_compl_top have "\ \ - \ = \" . - then show ?thesis by simp -qed - -lemma compl_top_eq [simp]: "- \ = \" -proof - - from inf_compl_bot have "\ \ - \ = \" . - then show ?thesis by simp -qed - -lemma compl_inf [simp]: "- (x \ y) = - x \ - y" -proof (rule compl_unique) - have "(x \ y) \ (- x \ - y) = (y \ (x \ - x)) \ (x \ (y \ - y))" - by (simp only: inf_sup_distrib inf_aci) - then show "(x \ y) \ (- x \ - y) = \" - by (simp add: inf_compl_bot) -next - have "(x \ y) \ (- x \ - y) = (- y \ (x \ - x)) \ (- x \ (y \ - y))" - by (simp only: sup_inf_distrib sup_aci) - then show "(x \ y) \ (- x \ - y) = \" - by (simp add: sup_compl_top) -qed - -lemma compl_sup [simp]: "- (x \ y) = - x \ - y" - using dual_boolean_algebra - by (rule boolean_algebra.compl_inf) - -lemma compl_mono: - assumes "x \ y" - shows "- y \ - x" -proof - - from assms have "x \ y = y" by (simp only: le_iff_sup) - then have "- (x \ y) = - y" by simp - then have "- x \ - y = - y" by simp - then have "- y \ - x = - y" by (simp only: inf_commute) - then show ?thesis by (simp only: le_iff_inf) -qed - -lemma compl_le_compl_iff [simp]: "- x \ - y \ y \ x" - by (auto dest: compl_mono) - -lemma compl_le_swap1: - assumes "y \ - x" - shows "x \ -y" -proof - - from assms have "- (- x) \ - y" by (simp only: compl_le_compl_iff) - then show ?thesis by simp -qed - -lemma compl_le_swap2: - assumes "- y \ x" - shows "- x \ y" -proof - - from assms have "- x \ - (- y)" by (simp only: compl_le_compl_iff) - then show ?thesis by simp -qed - -lemma compl_less_compl_iff [simp]: "- x < - y \ y < x" - by (auto simp add: less_le) - -lemma compl_less_swap1: - assumes "y < - x" - shows "x < - y" -proof - - from assms have "- (- x) < - y" by (simp only: compl_less_compl_iff) - then show ?thesis by simp -qed - -lemma compl_less_swap2: - assumes "- y < x" - shows "- x < y" -proof - - from assms have "- x < - (- y)" - by (simp only: compl_less_compl_iff) - then show ?thesis by simp -qed - -lemma sup_cancel_left1: "sup (sup x a) (sup (- x) b) = top" - by (simp add: ac_simps sup_compl_top) - -lemma sup_cancel_left2: "sup (sup (- x) a) (sup x b) = top" - by (simp add: ac_simps sup_compl_top) - -lemma inf_cancel_left1: "inf (inf x a) (inf (- x) b) = bot" - by (simp add: ac_simps inf_compl_bot) - -lemma inf_cancel_left2: "inf (inf (- x) a) (inf x b) = bot" - by (simp add: ac_simps inf_compl_bot) - -declare inf_compl_bot [simp] - and sup_compl_top [simp] - -lemma sup_compl_top_left1 [simp]: "sup (- x) (sup x y) = top" - by (simp add: sup_assoc[symmetric]) - -lemma sup_compl_top_left2 [simp]: "sup x (sup (- x) y) = top" - using sup_compl_top_left1[of "- x" y] by simp - -lemma inf_compl_bot_left1 [simp]: "inf (- x) (inf x y) = bot" - by (simp add: inf_assoc[symmetric]) - -lemma inf_compl_bot_left2 [simp]: "inf x (inf (- x) y) = bot" - using inf_compl_bot_left1[of "- x" y] by simp - -lemma inf_compl_bot_right [simp]: "inf x (inf y (- x)) = bot" - by (subst inf_left_commute) simp - -end - -locale boolean_algebra_cancel -begin - -lemma sup1: "(A::'a::semilattice_sup) \ sup k a \ sup A b \ sup k (sup a b)" - by (simp only: ac_simps) - -lemma sup2: "(B::'a::semilattice_sup) \ sup k b \ sup a B \ sup k (sup a b)" - by (simp only: ac_simps) - -lemma sup0: "(a::'a::bounded_semilattice_sup_bot) \ sup a bot" - by simp - -lemma inf1: "(A::'a::semilattice_inf) \ inf k a \ inf A b \ inf k (inf a b)" - by (simp only: ac_simps) - -lemma inf2: "(B::'a::semilattice_inf) \ inf k b \ inf a B \ inf k (inf a b)" - by (simp only: ac_simps) - -lemma inf0: "(a::'a::bounded_semilattice_inf_top) \ inf a top" - by simp - -end - -ML_file \Tools/boolean_algebra_cancel.ML\ - -simproc_setup boolean_algebra_cancel_sup ("sup a b::'a::boolean_algebra") = - \fn phi => fn ss => try Boolean_Algebra_Cancel.cancel_sup_conv\ - -simproc_setup boolean_algebra_cancel_inf ("inf a b::'a::boolean_algebra") = - \fn phi => fn ss => try Boolean_Algebra_Cancel.cancel_inf_conv\ - subsection \\min/max\ as special case of lattice\ @@ -853,33 +663,6 @@ qed -subsection \Lattice on \<^typ>\bool\\ - -instantiation bool :: boolean_algebra -begin - -definition bool_Compl_def [simp]: "uminus = Not" - -definition bool_diff_def [simp]: "A - B \ A \ \ B" - -definition [simp]: "P \ Q \ P \ Q" - -definition [simp]: "P \ Q \ P \ Q" - -instance by standard auto - -end - -lemma sup_boolI1: "P \ P \ Q" - by simp - -lemma sup_boolI2: "Q \ P \ Q" - by simp - -lemma sup_boolE: "P \ Q \ (P \ R) \ (Q \ R) \ R" - by auto - - subsection \Lattice on \<^typ>\_ \ _\\ instantiation "fun" :: (type, semilattice_sup) semilattice_sup @@ -938,60 +721,4 @@ end -instance "fun" :: (type, boolean_algebra) boolean_algebra - by standard (rule ext, simp_all add: inf_compl_bot sup_compl_top diff_eq)+ - - -subsection \Lattice on unary and binary predicates\ - -lemma inf1I: "A x \ B x \ (A \ B) x" - by (simp add: inf_fun_def) - -lemma inf2I: "A x y \ B x y \ (A \ B) x y" - by (simp add: inf_fun_def) - -lemma inf1E: "(A \ B) x \ (A x \ B x \ P) \ P" - by (simp add: inf_fun_def) - -lemma inf2E: "(A \ B) x y \ (A x y \ B x y \ P) \ P" - by (simp add: inf_fun_def) - -lemma inf1D1: "(A \ B) x \ A x" - by (rule inf1E) - -lemma inf2D1: "(A \ B) x y \ A x y" - by (rule inf2E) - -lemma inf1D2: "(A \ B) x \ B x" - by (rule inf1E) - -lemma inf2D2: "(A \ B) x y \ B x y" - by (rule inf2E) - -lemma sup1I1: "A x \ (A \ B) x" - by (simp add: sup_fun_def) - -lemma sup2I1: "A x y \ (A \ B) x y" - by (simp add: sup_fun_def) - -lemma sup1I2: "B x \ (A \ B) x" - by (simp add: sup_fun_def) - -lemma sup2I2: "B x y \ (A \ B) x y" - by (simp add: sup_fun_def) - -lemma sup1E: "(A \ B) x \ (A x \ P) \ (B x \ P) \ P" - by (simp add: sup_fun_def) iprover - -lemma sup2E: "(A \ B) x y \ (A x y \ P) \ (B x y \ P) \ P" - by (simp add: sup_fun_def) iprover - -text \ \<^medskip> Classical introduction rule: no commitment to \A\ vs \B\.\ - -lemma sup1CI: "(\ B x \ A x) \ (A \ B) x" - by (auto simp add: sup_fun_def) - -lemma sup2CI: "(\ B x y \ A x y) \ (A \ B) x y" - by (auto simp add: sup_fun_def) - end diff -r 7d3e818fe21f -r 7c5842b06114 src/HOL/Set.thy --- a/src/HOL/Set.thy Thu Aug 05 07:12:49 2021 +0000 +++ b/src/HOL/Set.thy Thu Aug 05 07:12:49 2021 +0000 @@ -7,7 +7,7 @@ section \Set theory for higher-order logic\ theory Set - imports Lattices Boolean_Algebra + imports Lattices Boolean_Algebras begin subsection \Sets as predicates\