# HG changeset patch # User haftmann # Date 1247340781 -7200 # Node ID 37390299214a075ca799e0260afc82733492415b # Parent 1d4d0b305f1697681c3d13d9c9be94c5b0754266 added boolean_algebra type class; tuned lattice duals diff -r 1d4d0b305f16 -r 37390299214a src/HOL/Finite_Set.thy --- a/src/HOL/Finite_Set.thy Fri Jul 10 09:24:50 2009 +0200 +++ b/src/HOL/Finite_Set.thy Sat Jul 11 21:33:01 2009 +0200 @@ -2719,7 +2719,7 @@ lemma (in upper_semilattice) ab_semigroup_idem_mult_sup: "ab_semigroup_idem_mult sup" by (rule lower_semilattice.ab_semigroup_idem_mult_inf) - (rule dual_lattice) + (rule dual_semilattice) context lattice begin @@ -2741,7 +2741,7 @@ apply(erule exE) apply(rule order_trans) apply(erule (1) fold1_belowI) -apply(erule (1) lower_semilattice.fold1_belowI [OF dual_lattice]) +apply(erule (1) lower_semilattice.fold1_belowI [OF dual_semilattice]) done lemma sup_Inf_absorb [simp]: @@ -2753,7 +2753,7 @@ lemma inf_Sup_absorb [simp]: "finite A \ a \ A \ inf a (\\<^bsub>fin\<^esub>A) = a" by (simp add: Sup_fin_def inf_absorb1 - lower_semilattice.fold1_belowI [OF dual_lattice]) + lower_semilattice.fold1_belowI [OF dual_semilattice]) end diff -r 1d4d0b305f16 -r 37390299214a src/HOL/Lattices.thy --- a/src/HOL/Lattices.thy Fri Jul 10 09:24:50 2009 +0200 +++ b/src/HOL/Lattices.thy Sat Jul 11 21:33:01 2009 +0200 @@ -29,7 +29,7 @@ text {* Dual lattice *} -lemma dual_lattice: +lemma dual_semilattice: "lower_semilattice (op \) (op >) sup" by (rule lower_semilattice.intro, rule dual_order) (unfold_locales, simp_all add: sup_least) @@ -180,6 +180,11 @@ context lattice begin +lemma dual_lattice: + "lattice (op \) (op >) sup inf" + by (rule lattice.intro, rule dual_semilattice, rule upper_semilattice.intro, rule dual_order) + (unfold_locales, auto) + lemma inf_sup_absorb: "x \ (x \ y) = x" by (blast intro: antisym inf_le1 inf_greatest sup_ge1) @@ -252,12 +257,148 @@ "(y \ z) \ x = (y \ x) \ (z \ x)" by(simp add:ACI inf_sup_distrib1) +lemma dual_distrib_lattice: + "distrib_lattice (op \) (op >) sup inf" + by (rule distrib_lattice.intro, rule dual_lattice) + (unfold_locales, fact inf_sup_distrib1) + lemmas distrib = sup_inf_distrib1 sup_inf_distrib2 inf_sup_distrib1 inf_sup_distrib2 end +subsection {* Boolean algebras *} + +class boolean_algebra = distrib_lattice + top + bot + minus + uminus + + assumes inf_compl_bot: "x \ - x = bot" + and sup_compl_top: "x \ - x = top" + assumes diff_eq: "x - y = x \ - y" +begin + +lemma dual_boolean_algebra: + "boolean_algebra (\x y. x \ - y) uminus (op \) (op >) (op \) (op \) top bot" + by (rule boolean_algebra.intro, rule dual_distrib_lattice) + (unfold_locales, + auto simp add: inf_compl_bot sup_compl_top diff_eq less_le_not_le) + +lemma compl_inf_bot: + "- x \ x = bot" + by (simp add: inf_commute inf_compl_bot) + +lemma compl_sup_top: + "- x \ x = top" + by (simp add: sup_commute sup_compl_top) + +lemma inf_bot_left [simp]: + "bot \ x = bot" + by (rule inf_absorb1) simp + +lemma inf_bot_right [simp]: + "x \ bot = bot" + by (rule inf_absorb2) simp + +lemma sup_top_left [simp]: + "top \ x = top" + by (rule sup_absorb1) simp + +lemma sup_top_right [simp]: + "x \ top = top" + by (rule sup_absorb2) simp + +lemma inf_top_left [simp]: + "top \ x = x" + by (rule inf_absorb2) simp + +lemma inf_top_right [simp]: + "x \ top = x" + by (rule inf_absorb1) simp + +lemma sup_bot_left [simp]: + "bot \ x = x" + by (rule sup_absorb2) simp + +lemma sup_bot_right [simp]: + "x \ bot = x" + by (rule sup_absorb1) simp + +lemma compl_unique: + assumes "x \ y = bot" + and "x \ y = top" + 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 \ top = y \ top" + using sup_compl_top assms(2) by simp + then show "- x = y" by (simp add: inf_top_right) +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 = bot" + and "- x \ y = top" + by (simp_all add: compl_inf_bot compl_sup_top) + then have "- (- x) = y" by (rule compl_unique) + then show "x = y" by simp +next + assume "x = y" + then show "- x = - y" by simp +qed + +lemma compl_bot_eq [simp]: + "- bot = top" +proof - + from sup_compl_top have "bot \ - bot = top" . + then show ?thesis by simp +qed + +lemma compl_top_eq [simp]: + "- top = bot" +proof - + from inf_compl_bot have "top \ - top = bot" . + then show ?thesis by simp +qed + +lemma compl_inf [simp]: + "- (x \ y) = - x \ - y" +proof (rule compl_unique) + have "(x \ y) \ (- x \ - y) = ((x \ y) \ - x) \ ((x \ y) \ - y)" + by (rule inf_sup_distrib1) + also have "... = (y \ (x \ - x)) \ (x \ (y \ - y))" + by (simp only: inf_commute inf_assoc inf_left_commute) + finally show "(x \ y) \ (- x \ - y) = bot" + by (simp add: inf_compl_bot) +next + have "(x \ y) \ (- x \ - y) = (x \ (- x \ - y)) \ (y \ (- x \ - y))" + by (rule sup_inf_distrib2) + also have "... = (- y \ (x \ - x)) \ (- x \ (y \ - y))" + by (simp only: sup_commute sup_assoc sup_left_commute) + finally show "(x \ y) \ (- x \ - y) = top" + by (simp add: sup_compl_top) +qed + +lemma compl_sup [simp]: + "- (x \ y) = - x \ - y" +proof - + interpret boolean_algebra "\x y. x \ - y" uminus "op \" "op >" "op \" "op \" top bot + by (rule dual_boolean_algebra) + then show ?thesis by simp +qed + +end + + subsection {* Uniqueness of inf and sup *} lemma (in lower_semilattice) inf_unique: @@ -330,17 +471,24 @@ subsection {* Bool as lattice *} -instantiation bool :: distrib_lattice +instantiation bool :: boolean_algebra begin definition + bool_Compl_def: "uminus = Not" + +definition + bool_diff_def: "A - B \ A \ \ B" + +definition inf_bool_eq: "P \ Q \ P \ Q" definition sup_bool_eq: "P \ Q \ P \ Q" -instance - by intro_classes (auto simp add: inf_bool_eq sup_bool_eq le_bool_def) +instance proof +qed (simp_all add: inf_bool_eq sup_bool_eq le_bool_def + bot_bool_eq top_bool_eq bool_Compl_def bool_diff_def, auto) end @@ -369,7 +517,33 @@ end instance "fun" :: (type, distrib_lattice) distrib_lattice - by default (auto simp add: inf_fun_eq sup_fun_eq sup_inf_distrib1) +proof +qed (auto simp add: inf_fun_eq sup_fun_eq sup_inf_distrib1) + +instantiation "fun" :: (type, uminus) uminus +begin + +definition + fun_Compl_def: "- A = (\x. - A x)" + +instance .. + +end + +instantiation "fun" :: (type, minus) minus +begin + +definition + fun_diff_def: "A - B = (\x. A x - B x)" + +instance .. + +end + +instance "fun" :: (type, boolean_algebra) boolean_algebra +proof +qed (simp_all add: inf_fun_eq sup_fun_eq bot_fun_eq top_fun_eq fun_Compl_def fun_diff_def + inf_compl_bot sup_compl_top diff_eq) text {* redundant bindings *} diff -r 1d4d0b305f16 -r 37390299214a src/HOL/Set.thy --- a/src/HOL/Set.thy Fri Jul 10 09:24:50 2009 +0200 +++ b/src/HOL/Set.thy Sat Jul 11 21:33:01 2009 +0200 @@ -10,7 +10,6 @@ text {* A set in HOL is simply a predicate. *} - subsection {* Basic syntax *} global @@ -363,46 +362,6 @@ Bex_def: "Bex A P == EX x. x:A & P(x)" Bex1_def: "Bex1 A P == EX! x. x:A & P(x)" -instantiation "fun" :: (type, minus) minus -begin - -definition - fun_diff_def: "A - B = (%x. A x - B x)" - -instance .. - -end - -instantiation bool :: minus -begin - -definition - bool_diff_def: "A - B = (A & ~ B)" - -instance .. - -end - -instantiation "fun" :: (type, uminus) uminus -begin - -definition - fun_Compl_def: "- A = (%x. - A x)" - -instance .. - -end - -instantiation bool :: uminus -begin - -definition - bool_Compl_def: "- A = (~ A)" - -instance .. - -end - definition Pow :: "'a set => 'a set set" where Pow_def: "Pow A = {B. B \ A}"