haftmann@21249: (* Title: HOL/Lattices.thy haftmann@21249: Author: Tobias Nipkow haftmann@21249: *) haftmann@21249: haftmann@22454: header {* Abstract lattices *} haftmann@21249: haftmann@21249: theory Lattices blanchet@54555: imports Groups haftmann@21249: begin haftmann@21249: haftmann@35301: subsection {* Abstract semilattice *} haftmann@35301: haftmann@35301: text {* haftmann@51487: These locales provide a basic structure for interpretation into haftmann@35301: bigger structures; extensions require careful thinking, otherwise haftmann@35301: undesired effects may occur due to interpretation. haftmann@35301: *} haftmann@35301: haftmann@51487: no_notation times (infixl "*" 70) haftmann@51487: no_notation Groups.one ("1") haftmann@51487: haftmann@35301: locale semilattice = abel_semigroup + haftmann@51487: assumes idem [simp]: "a * a = a" haftmann@35301: begin haftmann@35301: haftmann@51487: lemma left_idem [simp]: "a * (a * b) = a * b" nipkow@50615: by (simp add: assoc [symmetric]) nipkow@50615: haftmann@51487: lemma right_idem [simp]: "(a * b) * b = a * b" nipkow@50615: by (simp add: assoc) haftmann@35301: haftmann@35301: end haftmann@35301: haftmann@51487: locale semilattice_neutr = semilattice + comm_monoid haftmann@35301: haftmann@51487: locale semilattice_order = semilattice + haftmann@51487: fixes less_eq :: "'a \ 'a \ bool" (infix "\" 50) haftmann@51487: and less :: "'a \ 'a \ bool" (infix "\" 50) haftmann@51487: assumes order_iff: "a \ b \ a = a * b" haftmann@51487: and semilattice_strict_iff_order: "a \ b \ a \ b \ a \ b" haftmann@51487: begin haftmann@51487: haftmann@51487: lemma orderI: haftmann@51487: "a = a * b \ a \ b" haftmann@51487: by (simp add: order_iff) haftmann@51487: haftmann@51487: lemma orderE: haftmann@51487: assumes "a \ b" haftmann@51487: obtains "a = a * b" haftmann@51487: using assms by (unfold order_iff) haftmann@51487: haftmann@52152: sublocale ordering less_eq less haftmann@51487: proof haftmann@51487: fix a b haftmann@51487: show "a \ b \ a \ b \ a \ b" haftmann@51487: by (fact semilattice_strict_iff_order) haftmann@51487: next haftmann@51487: fix a haftmann@51487: show "a \ a" haftmann@51487: by (simp add: order_iff) haftmann@51487: next haftmann@51487: fix a b haftmann@51487: assume "a \ b" "b \ a" haftmann@51487: then have "a = a * b" "a * b = b" haftmann@51487: by (simp_all add: order_iff commute) haftmann@51487: then show "a = b" by simp haftmann@51487: next haftmann@51487: fix a b c haftmann@51487: assume "a \ b" "b \ c" haftmann@51487: then have "a = a * b" "b = b * c" haftmann@51487: by (simp_all add: order_iff commute) haftmann@51487: then have "a = a * (b * c)" haftmann@51487: by simp haftmann@51487: then have "a = (a * b) * c" haftmann@51487: by (simp add: assoc) haftmann@51487: with `a = a * b` [symmetric] have "a = a * c" by simp haftmann@51487: then show "a \ c" by (rule orderI) haftmann@51487: qed haftmann@35301: haftmann@51487: lemma cobounded1 [simp]: haftmann@51487: "a * b \ a" haftmann@51487: by (simp add: order_iff commute) haftmann@51487: haftmann@51487: lemma cobounded2 [simp]: haftmann@51487: "a * b \ b" haftmann@51487: by (simp add: order_iff) haftmann@51487: haftmann@51487: lemma boundedI: haftmann@51487: assumes "a \ b" and "a \ c" haftmann@51487: shows "a \ b * c" haftmann@51487: proof (rule orderI) haftmann@51487: from assms obtain "a * b = a" and "a * c = a" by (auto elim!: orderE) haftmann@51487: then show "a = a * (b * c)" by (simp add: assoc [symmetric]) haftmann@51487: qed haftmann@51487: haftmann@51487: lemma boundedE: haftmann@51487: assumes "a \ b * c" haftmann@51487: obtains "a \ b" and "a \ c" haftmann@51487: using assms by (blast intro: trans cobounded1 cobounded2) haftmann@51487: haftmann@54859: lemma bounded_iff [simp]: haftmann@51487: "a \ b * c \ a \ b \ a \ c" haftmann@51487: by (blast intro: boundedI elim: boundedE) haftmann@51487: haftmann@51487: lemma strict_boundedE: haftmann@51487: assumes "a \ b * c" haftmann@51487: obtains "a \ b" and "a \ c" haftmann@54859: using assms by (auto simp add: commute strict_iff_order elim: orderE intro!: that)+ haftmann@51487: haftmann@51487: lemma coboundedI1: haftmann@51487: "a \ c \ a * b \ c" haftmann@51487: by (rule trans) auto haftmann@51487: haftmann@51487: lemma coboundedI2: haftmann@51487: "b \ c \ a * b \ c" haftmann@51487: by (rule trans) auto haftmann@51487: haftmann@54858: lemma strict_coboundedI1: haftmann@54858: "a \ c \ a * b \ c" haftmann@54858: using irrefl haftmann@54858: by (auto intro: not_eq_order_implies_strict coboundedI1 strict_implies_order elim: strict_boundedE) haftmann@54858: haftmann@54858: lemma strict_coboundedI2: haftmann@54858: "b \ c \ a * b \ c" haftmann@54858: using strict_coboundedI1 [of b c a] by (simp add: commute) haftmann@54858: haftmann@51487: lemma mono: "a \ c \ b \ d \ a * b \ c * d" haftmann@51487: by (blast intro: boundedI coboundedI1 coboundedI2) haftmann@51487: haftmann@51487: lemma absorb1: "a \ b \ a * b = a" haftmann@54859: by (rule antisym) (auto simp add: refl) haftmann@51487: haftmann@51487: lemma absorb2: "b \ a \ a * b = b" haftmann@54859: by (rule antisym) (auto simp add: refl) haftmann@54858: haftmann@54858: lemma absorb_iff1: "a \ b \ a * b = a" haftmann@54858: using order_iff by auto haftmann@54858: haftmann@54858: lemma absorb_iff2: "b \ a \ a * b = b" haftmann@54858: using order_iff by (auto simp add: commute) haftmann@35301: haftmann@35301: end haftmann@35301: haftmann@51487: locale semilattice_neutr_order = semilattice_neutr + semilattice_order haftmann@52152: begin haftmann@51487: haftmann@52152: sublocale ordering_top less_eq less 1 haftmann@51487: by default (simp add: order_iff) haftmann@51487: haftmann@52152: end haftmann@52152: haftmann@51487: notation times (infixl "*" 70) haftmann@51487: notation Groups.one ("1") haftmann@51487: haftmann@35301: haftmann@46691: subsection {* Syntactic infimum and supremum operations *} haftmann@41082: krauss@44845: class inf = krauss@44845: fixes inf :: "'a \ 'a \ 'a" (infixl "\" 70) haftmann@25206: krauss@44845: class sup = krauss@44845: fixes sup :: "'a \ 'a \ 'a" (infixl "\" 65) krauss@44845: haftmann@46691: haftmann@46691: subsection {* Concrete lattices *} haftmann@46691: haftmann@46691: notation haftmann@46691: less_eq (infix "\" 50) and haftmann@46691: less (infix "\" 50) haftmann@46691: krauss@44845: class semilattice_inf = order + inf + haftmann@22737: assumes inf_le1 [simp]: "x \ y \ x" haftmann@22737: and inf_le2 [simp]: "x \ y \ y" nipkow@21733: and inf_greatest: "x \ y \ x \ z \ x \ y \ z" haftmann@21249: krauss@44845: class semilattice_sup = order + sup + haftmann@22737: assumes sup_ge1 [simp]: "x \ x \ y" haftmann@22737: and sup_ge2 [simp]: "y \ x \ y" nipkow@21733: and sup_least: "y \ x \ z \ x \ y \ z \ x" haftmann@26014: begin haftmann@26014: haftmann@26014: text {* Dual lattice *} haftmann@26014: haftmann@31991: lemma dual_semilattice: krauss@44845: "class.semilattice_inf sup greater_eq greater" haftmann@36635: by (rule class.semilattice_inf.intro, rule dual_order) haftmann@27682: (unfold_locales, simp_all add: sup_least) haftmann@26014: haftmann@26014: end haftmann@21249: haftmann@35028: class lattice = semilattice_inf + semilattice_sup haftmann@21249: wenzelm@25382: haftmann@28562: subsubsection {* Intro and elim rules*} nipkow@21733: haftmann@35028: context semilattice_inf nipkow@21733: begin haftmann@21249: haftmann@32064: lemma le_infI1: haftmann@32064: "a \ x \ a \ b \ x" haftmann@32064: by (rule order_trans) auto haftmann@21249: haftmann@32064: lemma le_infI2: haftmann@32064: "b \ x \ a \ b \ x" haftmann@32064: by (rule order_trans) auto nipkow@21733: haftmann@32064: lemma le_infI: "x \ a \ x \ b \ x \ a \ b" haftmann@54857: by (fact inf_greatest) (* FIXME: duplicate lemma *) haftmann@21249: haftmann@32064: lemma le_infE: "x \ a \ b \ (x \ a \ x \ b \ P) \ P" huffman@36008: by (blast intro: order_trans inf_le1 inf_le2) haftmann@21249: haftmann@54859: lemma le_inf_iff: haftmann@32064: "x \ y \ z \ x \ y \ x \ z" haftmann@32064: by (blast intro: le_infI elim: le_infE) nipkow@21733: haftmann@32064: lemma le_iff_inf: haftmann@32064: "x \ y \ x \ y = x" haftmann@54859: by (auto intro: le_infI1 antisym dest: eq_iff [THEN iffD1] simp add: le_inf_iff) haftmann@21249: haftmann@43753: lemma inf_mono: "a \ c \ b \ d \ a \ b \ c \ d" huffman@36008: by (fast intro: inf_greatest le_infI1 le_infI2) huffman@36008: haftmann@25206: lemma mono_inf: haftmann@35028: fixes f :: "'a \ 'b\semilattice_inf" haftmann@34007: shows "mono f \ f (A \ B) \ f A \ f B" haftmann@25206: by (auto simp add: mono_def intro: Lattices.inf_greatest) nipkow@21733: haftmann@25206: end nipkow@21733: haftmann@35028: context semilattice_sup nipkow@21733: begin haftmann@21249: haftmann@32064: lemma le_supI1: haftmann@32064: "x \ a \ x \ a \ b" haftmann@25062: by (rule order_trans) auto haftmann@21249: haftmann@32064: lemma le_supI2: haftmann@32064: "x \ b \ x \ a \ b" haftmann@25062: by (rule order_trans) auto nipkow@21733: haftmann@32064: lemma le_supI: haftmann@32064: "a \ x \ b \ x \ a \ b \ x" haftmann@54857: by (fact sup_least) (* FIXME: duplicate lemma *) haftmann@21249: haftmann@32064: lemma le_supE: haftmann@32064: "a \ b \ x \ (a \ x \ b \ x \ P) \ P" huffman@36008: by (blast intro: order_trans sup_ge1 sup_ge2) haftmann@22422: haftmann@54859: lemma le_sup_iff: haftmann@32064: "x \ y \ z \ x \ z \ y \ z" haftmann@32064: by (blast intro: le_supI elim: le_supE) nipkow@21733: haftmann@32064: lemma le_iff_sup: haftmann@32064: "x \ y \ x \ y = y" haftmann@54859: by (auto intro: le_supI2 antisym dest: eq_iff [THEN iffD1] simp add: le_sup_iff) nipkow@21734: haftmann@43753: lemma sup_mono: "a \ c \ b \ d \ a \ b \ c \ d" huffman@36008: by (fast intro: sup_least le_supI1 le_supI2) huffman@36008: haftmann@25206: lemma mono_sup: haftmann@35028: fixes f :: "'a \ 'b\semilattice_sup" haftmann@34007: shows "mono f \ f A \ f B \ f (A \ B)" haftmann@25206: by (auto simp add: mono_def intro: Lattices.sup_least) nipkow@21733: haftmann@25206: end haftmann@23878: nipkow@21733: haftmann@32064: subsubsection {* Equational laws *} haftmann@21249: haftmann@52152: context semilattice_inf haftmann@52152: begin haftmann@52152: haftmann@52152: sublocale inf!: semilattice inf haftmann@34973: proof haftmann@34973: fix a b c haftmann@34973: show "(a \ b) \ c = a \ (b \ c)" haftmann@54859: by (rule antisym) (auto intro: le_infI1 le_infI2 simp add: le_inf_iff) haftmann@34973: show "a \ b = b \ a" haftmann@54859: by (rule antisym) (auto simp add: le_inf_iff) haftmann@34973: show "a \ a = a" haftmann@54859: by (rule antisym) (auto simp add: le_inf_iff) haftmann@34973: qed haftmann@34973: haftmann@52152: sublocale inf!: semilattice_order inf less_eq less haftmann@51487: by default (auto simp add: le_iff_inf less_le) haftmann@51487: haftmann@34973: lemma inf_assoc: "(x \ y) \ z = x \ (y \ z)" haftmann@34973: by (fact inf.assoc) nipkow@21733: haftmann@34973: lemma inf_commute: "(x \ y) = (y \ x)" haftmann@34973: by (fact inf.commute) nipkow@21733: haftmann@34973: lemma inf_left_commute: "x \ (y \ z) = y \ (x \ z)" haftmann@34973: by (fact inf.left_commute) nipkow@21733: huffman@44921: lemma inf_idem: "x \ x = x" huffman@44921: by (fact inf.idem) (* already simp *) haftmann@34973: nipkow@50615: lemma inf_left_idem: "x \ (x \ y) = x \ y" nipkow@50615: by (fact inf.left_idem) (* already simp *) nipkow@50615: nipkow@50615: lemma inf_right_idem: "(x \ y) \ y = x \ y" nipkow@50615: by (fact inf.right_idem) (* already simp *) nipkow@21733: haftmann@32642: lemma inf_absorb1: "x \ y \ x \ y = x" haftmann@32064: by (rule antisym) auto nipkow@21733: haftmann@32642: lemma inf_absorb2: "y \ x \ x \ y = y" haftmann@32064: by (rule antisym) auto haftmann@34973: haftmann@32064: lemmas inf_aci = inf_commute inf_assoc inf_left_commute inf_left_idem nipkow@21733: nipkow@21733: end nipkow@21733: haftmann@35028: context semilattice_sup nipkow@21733: begin haftmann@21249: haftmann@52152: sublocale sup!: semilattice sup haftmann@52152: proof haftmann@52152: fix a b c haftmann@52152: show "(a \ b) \ c = a \ (b \ c)" haftmann@54859: by (rule antisym) (auto intro: le_supI1 le_supI2 simp add: le_sup_iff) haftmann@52152: show "a \ b = b \ a" haftmann@54859: by (rule antisym) (auto simp add: le_sup_iff) haftmann@52152: show "a \ a = a" haftmann@54859: by (rule antisym) (auto simp add: le_sup_iff) haftmann@52152: qed haftmann@52152: haftmann@52152: sublocale sup!: semilattice_order sup greater_eq greater haftmann@52152: by default (auto simp add: le_iff_sup sup.commute less_le) haftmann@52152: haftmann@34973: lemma sup_assoc: "(x \ y) \ z = x \ (y \ z)" haftmann@34973: by (fact sup.assoc) nipkow@21733: haftmann@34973: lemma sup_commute: "(x \ y) = (y \ x)" haftmann@34973: by (fact sup.commute) nipkow@21733: haftmann@34973: lemma sup_left_commute: "x \ (y \ z) = y \ (x \ z)" haftmann@34973: by (fact sup.left_commute) nipkow@21733: huffman@44921: lemma sup_idem: "x \ x = x" huffman@44921: by (fact sup.idem) (* already simp *) haftmann@34973: noschinl@44918: lemma sup_left_idem [simp]: "x \ (x \ y) = x \ y" haftmann@34973: by (fact sup.left_idem) nipkow@21733: haftmann@32642: lemma sup_absorb1: "y \ x \ x \ y = x" haftmann@32064: by (rule antisym) auto nipkow@21733: haftmann@32642: lemma sup_absorb2: "x \ y \ x \ y = y" haftmann@32064: by (rule antisym) auto haftmann@21249: haftmann@32064: lemmas sup_aci = sup_commute sup_assoc sup_left_commute sup_left_idem nipkow@21733: nipkow@21733: end haftmann@21249: nipkow@21733: context lattice nipkow@21733: begin nipkow@21733: haftmann@31991: lemma dual_lattice: krauss@44845: "class.lattice sup (op \) (op >) inf" haftmann@36635: by (rule class.lattice.intro, rule dual_semilattice, rule class.semilattice_sup.intro, rule dual_order) haftmann@31991: (unfold_locales, auto) haftmann@31991: noschinl@44918: lemma inf_sup_absorb [simp]: "x \ (x \ y) = x" haftmann@25102: by (blast intro: antisym inf_le1 inf_greatest sup_ge1) nipkow@21733: noschinl@44918: lemma sup_inf_absorb [simp]: "x \ (x \ y) = x" haftmann@25102: by (blast intro: antisym sup_ge1 sup_least inf_le1) nipkow@21733: haftmann@32064: lemmas inf_sup_aci = inf_aci sup_aci nipkow@21734: haftmann@22454: lemmas inf_sup_ord = inf_le1 inf_le2 sup_ge1 sup_ge2 haftmann@22454: nipkow@21734: text{* Towards distributivity *} haftmann@21249: nipkow@21734: lemma distrib_sup_le: "x \ (y \ z) \ (x \ y) \ (x \ z)" haftmann@32064: by (auto intro: le_infI1 le_infI2 le_supI1 le_supI2) nipkow@21734: nipkow@21734: lemma distrib_inf_le: "(x \ y) \ (x \ z) \ x \ (y \ z)" haftmann@32064: by (auto intro: le_infI1 le_infI2 le_supI1 le_supI2) nipkow@21734: nipkow@21734: text{* If you have one of them, you have them all. *} haftmann@21249: nipkow@21733: lemma distrib_imp1: haftmann@21249: assumes D: "!!x y z. x \ (y \ z) = (x \ y) \ (x \ z)" haftmann@21249: shows "x \ (y \ z) = (x \ y) \ (x \ z)" haftmann@21249: proof- noschinl@44918: have "x \ (y \ z) = (x \ (x \ z)) \ (y \ z)" by simp noschinl@44918: also have "\ = x \ (z \ (x \ y))" noschinl@44918: by (simp add: D inf_commute sup_assoc del: sup_inf_absorb) haftmann@21249: also have "\ = ((x \ y) \ x) \ ((x \ y) \ z)" noschinl@44919: by(simp add: inf_commute) haftmann@21249: also have "\ = (x \ y) \ (x \ z)" by(simp add:D) haftmann@21249: finally show ?thesis . haftmann@21249: qed haftmann@21249: nipkow@21733: lemma distrib_imp2: haftmann@21249: assumes D: "!!x y z. x \ (y \ z) = (x \ y) \ (x \ z)" haftmann@21249: shows "x \ (y \ z) = (x \ y) \ (x \ z)" haftmann@21249: proof- noschinl@44918: have "x \ (y \ z) = (x \ (x \ z)) \ (y \ z)" by simp noschinl@44918: also have "\ = x \ (z \ (x \ y))" noschinl@44918: by (simp add: D sup_commute inf_assoc del: inf_sup_absorb) haftmann@21249: also have "\ = ((x \ y) \ x) \ ((x \ y) \ z)" noschinl@44919: by(simp add: sup_commute) haftmann@21249: also have "\ = (x \ y) \ (x \ z)" by(simp add:D) haftmann@21249: finally show ?thesis . haftmann@21249: qed haftmann@21249: nipkow@21733: end haftmann@21249: haftmann@32568: subsubsection {* Strict order *} haftmann@32568: haftmann@35028: context semilattice_inf haftmann@32568: begin haftmann@32568: haftmann@32568: lemma less_infI1: haftmann@32568: "a \ x \ a \ b \ x" haftmann@32642: by (auto simp add: less_le inf_absorb1 intro: le_infI1) haftmann@32568: haftmann@32568: lemma less_infI2: haftmann@32568: "b \ x \ a \ b \ x" haftmann@32642: by (auto simp add: less_le inf_absorb2 intro: le_infI2) haftmann@32568: haftmann@32568: end haftmann@32568: haftmann@35028: context semilattice_sup haftmann@32568: begin haftmann@32568: haftmann@32568: lemma less_supI1: haftmann@34007: "x \ a \ x \ a \ b" huffman@44921: using dual_semilattice huffman@44921: by (rule semilattice_inf.less_infI1) haftmann@32568: haftmann@32568: lemma less_supI2: haftmann@34007: "x \ b \ x \ a \ b" huffman@44921: using dual_semilattice huffman@44921: by (rule semilattice_inf.less_infI2) haftmann@32568: haftmann@32568: end haftmann@32568: haftmann@21249: haftmann@24164: subsection {* Distributive lattices *} haftmann@21249: haftmann@22454: class distrib_lattice = lattice + haftmann@21249: assumes sup_inf_distrib1: "x \ (y \ z) = (x \ y) \ (x \ z)" haftmann@21249: nipkow@21733: context distrib_lattice nipkow@21733: begin nipkow@21733: nipkow@21733: lemma sup_inf_distrib2: huffman@44921: "(y \ z) \ x = (y \ x) \ (z \ x)" huffman@44921: by (simp add: sup_commute sup_inf_distrib1) haftmann@21249: nipkow@21733: lemma inf_sup_distrib1: huffman@44921: "x \ (y \ z) = (x \ y) \ (x \ z)" huffman@44921: by (rule distrib_imp2 [OF sup_inf_distrib1]) haftmann@21249: nipkow@21733: lemma inf_sup_distrib2: huffman@44921: "(y \ z) \ x = (y \ x) \ (z \ x)" huffman@44921: by (simp add: inf_commute inf_sup_distrib1) haftmann@21249: haftmann@31991: lemma dual_distrib_lattice: krauss@44845: "class.distrib_lattice sup (op \) (op >) inf" haftmann@36635: by (rule class.distrib_lattice.intro, rule dual_lattice) haftmann@31991: (unfold_locales, fact inf_sup_distrib1) haftmann@31991: huffman@36008: lemmas sup_inf_distrib = huffman@36008: sup_inf_distrib1 sup_inf_distrib2 huffman@36008: huffman@36008: lemmas inf_sup_distrib = huffman@36008: inf_sup_distrib1 inf_sup_distrib2 huffman@36008: nipkow@21733: lemmas distrib = haftmann@21249: sup_inf_distrib1 sup_inf_distrib2 inf_sup_distrib1 inf_sup_distrib2 haftmann@21249: nipkow@21733: end nipkow@21733: haftmann@21249: haftmann@34007: subsection {* Bounded lattices and boolean algebras *} haftmann@31991: haftmann@52729: class bounded_semilattice_inf_top = semilattice_inf + order_top haftmann@52152: begin haftmann@51487: haftmann@52152: sublocale inf_top!: semilattice_neutr inf top haftmann@51546: + inf_top!: semilattice_neutr_order inf top less_eq less haftmann@51487: proof haftmann@51487: fix x haftmann@51487: show "x \ \ = x" haftmann@51487: by (rule inf_absorb1) simp haftmann@51487: qed haftmann@51487: haftmann@52152: end haftmann@51487: haftmann@52729: class bounded_semilattice_sup_bot = semilattice_sup + order_bot haftmann@52152: begin haftmann@52152: haftmann@52152: sublocale sup_bot!: semilattice_neutr sup bot haftmann@51546: + sup_bot!: semilattice_neutr_order sup bot greater_eq greater haftmann@51487: proof haftmann@51487: fix x haftmann@51487: show "x \ \ = x" haftmann@51487: by (rule sup_absorb1) simp haftmann@51487: qed haftmann@51487: haftmann@52152: end haftmann@52152: haftmann@52729: class bounded_lattice_bot = lattice + order_bot haftmann@31991: begin haftmann@31991: haftmann@51487: subclass bounded_semilattice_sup_bot .. haftmann@51487: haftmann@31991: lemma inf_bot_left [simp]: haftmann@34007: "\ \ x = \" haftmann@31991: by (rule inf_absorb1) simp haftmann@31991: haftmann@31991: lemma inf_bot_right [simp]: haftmann@34007: "x \ \ = \" haftmann@31991: by (rule inf_absorb2) simp haftmann@31991: haftmann@51487: lemma sup_bot_left: kaliszyk@36352: "\ \ x = x" haftmann@51487: by (fact sup_bot.left_neutral) kaliszyk@36352: haftmann@51487: lemma sup_bot_right: kaliszyk@36352: "x \ \ = x" haftmann@51487: by (fact sup_bot.right_neutral) kaliszyk@36352: kaliszyk@36352: lemma sup_eq_bot_iff [simp]: kaliszyk@36352: "x \ y = \ \ x = \ \ y = \" kaliszyk@36352: by (simp add: eq_iff) kaliszyk@36352: nipkow@51593: lemma bot_eq_sup_iff [simp]: nipkow@51593: "\ = x \ y \ x = \ \ y = \" nipkow@51593: by (simp add: eq_iff) nipkow@51593: kaliszyk@36352: end kaliszyk@36352: haftmann@52729: class bounded_lattice_top = lattice + order_top kaliszyk@36352: begin kaliszyk@36352: haftmann@51487: subclass bounded_semilattice_inf_top .. haftmann@51487: haftmann@31991: lemma sup_top_left [simp]: haftmann@34007: "\ \ x = \" haftmann@31991: by (rule sup_absorb1) simp haftmann@31991: haftmann@31991: lemma sup_top_right [simp]: haftmann@34007: "x \ \ = \" haftmann@31991: by (rule sup_absorb2) simp haftmann@31991: haftmann@51487: lemma inf_top_left: haftmann@34007: "\ \ x = x" haftmann@51487: by (fact inf_top.left_neutral) haftmann@31991: haftmann@51487: lemma inf_top_right: haftmann@34007: "x \ \ = x" haftmann@51487: by (fact inf_top.right_neutral) haftmann@31991: huffman@36008: lemma inf_eq_top_iff [simp]: huffman@36008: "x \ y = \ \ x = \ \ y = \" huffman@36008: by (simp add: eq_iff) haftmann@32568: kaliszyk@36352: end kaliszyk@36352: haftmann@52729: class bounded_lattice = lattice + order_bot + order_top kaliszyk@36352: begin kaliszyk@36352: haftmann@51487: subclass bounded_lattice_bot .. haftmann@51487: subclass bounded_lattice_top .. haftmann@51487: kaliszyk@36352: lemma dual_bounded_lattice: krauss@44845: "class.bounded_lattice sup greater_eq greater inf \ \" kaliszyk@36352: by unfold_locales (auto simp add: less_le_not_le) haftmann@32568: haftmann@34007: end haftmann@34007: haftmann@34007: class boolean_algebra = distrib_lattice + bounded_lattice + minus + uminus + haftmann@34007: assumes inf_compl_bot: "x \ - x = \" haftmann@34007: and sup_compl_top: "x \ - x = \" haftmann@34007: assumes diff_eq: "x - y = x \ - y" haftmann@34007: begin haftmann@34007: haftmann@34007: lemma dual_boolean_algebra: krauss@44845: "class.boolean_algebra (\x y. x \ - y) uminus sup greater_eq greater inf \ \" haftmann@36635: by (rule class.boolean_algebra.intro, rule dual_bounded_lattice, rule dual_distrib_lattice) haftmann@34007: (unfold_locales, auto simp add: inf_compl_bot sup_compl_top diff_eq) haftmann@34007: noschinl@44918: lemma compl_inf_bot [simp]: haftmann@34007: "- x \ x = \" haftmann@34007: by (simp add: inf_commute inf_compl_bot) haftmann@34007: noschinl@44918: lemma compl_sup_top [simp]: haftmann@34007: "- x \ x = \" haftmann@34007: by (simp add: sup_commute sup_compl_top) haftmann@34007: haftmann@31991: lemma compl_unique: haftmann@34007: assumes "x \ y = \" haftmann@34007: and "x \ y = \" haftmann@31991: shows "- x = y" haftmann@31991: proof - haftmann@31991: have "(x \ - x) \ (- x \ y) = (x \ y) \ (- x \ y)" haftmann@31991: using inf_compl_bot assms(1) by simp haftmann@31991: then have "(- x \ x) \ (- x \ y) = (y \ x) \ (y \ - x)" haftmann@31991: by (simp add: inf_commute) haftmann@31991: then have "- x \ (x \ y) = y \ (x \ - x)" haftmann@31991: by (simp add: inf_sup_distrib1) haftmann@34007: then have "- x \ \ = y \ \" haftmann@31991: using sup_compl_top assms(2) by simp krauss@34209: then show "- x = y" by simp haftmann@31991: qed haftmann@31991: haftmann@31991: lemma double_compl [simp]: haftmann@31991: "- (- x) = x" haftmann@31991: using compl_inf_bot compl_sup_top by (rule compl_unique) haftmann@31991: haftmann@31991: lemma compl_eq_compl_iff [simp]: haftmann@31991: "- x = - y \ x = y" haftmann@31991: proof haftmann@31991: assume "- x = - y" huffman@36008: then have "- (- x) = - (- y)" by (rule arg_cong) haftmann@31991: then show "x = y" by simp haftmann@31991: next haftmann@31991: assume "x = y" haftmann@31991: then show "- x = - y" by simp haftmann@31991: qed haftmann@31991: haftmann@31991: lemma compl_bot_eq [simp]: haftmann@34007: "- \ = \" haftmann@31991: proof - haftmann@34007: from sup_compl_top have "\ \ - \ = \" . haftmann@31991: then show ?thesis by simp haftmann@31991: qed haftmann@31991: haftmann@31991: lemma compl_top_eq [simp]: haftmann@34007: "- \ = \" haftmann@31991: proof - haftmann@34007: from inf_compl_bot have "\ \ - \ = \" . haftmann@31991: then show ?thesis by simp haftmann@31991: qed haftmann@31991: haftmann@31991: lemma compl_inf [simp]: haftmann@31991: "- (x \ y) = - x \ - y" haftmann@31991: proof (rule compl_unique) huffman@36008: have "(x \ y) \ (- x \ - y) = (y \ (x \ - x)) \ (x \ (y \ - y))" huffman@36008: by (simp only: inf_sup_distrib inf_aci) huffman@36008: then show "(x \ y) \ (- x \ - y) = \" haftmann@31991: by (simp add: inf_compl_bot) haftmann@31991: next huffman@36008: have "(x \ y) \ (- x \ - y) = (- y \ (x \ - x)) \ (- x \ (y \ - y))" huffman@36008: by (simp only: sup_inf_distrib sup_aci) huffman@36008: then show "(x \ y) \ (- x \ - y) = \" haftmann@31991: by (simp add: sup_compl_top) haftmann@31991: qed haftmann@31991: haftmann@31991: lemma compl_sup [simp]: haftmann@31991: "- (x \ y) = - x \ - y" huffman@44921: using dual_boolean_algebra huffman@44921: by (rule boolean_algebra.compl_inf) haftmann@31991: huffman@36008: lemma compl_mono: huffman@36008: "x \ y \ - y \ - x" huffman@36008: proof - huffman@36008: assume "x \ y" huffman@36008: then have "x \ y = y" by (simp only: le_iff_sup) huffman@36008: then have "- (x \ y) = - y" by simp huffman@36008: then have "- x \ - y = - y" by simp huffman@36008: then have "- y \ - x = - y" by (simp only: inf_commute) huffman@36008: then show "- y \ - x" by (simp only: le_iff_inf) huffman@36008: qed huffman@36008: noschinl@44918: lemma compl_le_compl_iff [simp]: haftmann@43753: "- x \ - y \ y \ x" haftmann@43873: by (auto dest: compl_mono) haftmann@43873: haftmann@43873: lemma compl_le_swap1: haftmann@43873: assumes "y \ - x" shows "x \ -y" haftmann@43873: proof - haftmann@43873: from assms have "- (- x) \ - y" by (simp only: compl_le_compl_iff) haftmann@43873: then show ?thesis by simp haftmann@43873: qed haftmann@43873: haftmann@43873: lemma compl_le_swap2: haftmann@43873: assumes "- y \ x" shows "- x \ y" haftmann@43873: proof - haftmann@43873: from assms have "- x \ - (- y)" by (simp only: compl_le_compl_iff) haftmann@43873: then show ?thesis by simp haftmann@43873: qed haftmann@43873: haftmann@43873: lemma compl_less_compl_iff: (* TODO: declare [simp] ? *) haftmann@43873: "- x \ - y \ y \ x" noschinl@44919: by (auto simp add: less_le) haftmann@43873: haftmann@43873: lemma compl_less_swap1: haftmann@43873: assumes "y \ - x" shows "x \ - y" haftmann@43873: proof - haftmann@43873: from assms have "- (- x) \ - y" by (simp only: compl_less_compl_iff) haftmann@43873: then show ?thesis by simp haftmann@43873: qed haftmann@43873: haftmann@43873: lemma compl_less_swap2: haftmann@43873: assumes "- y \ x" shows "- x \ y" haftmann@43873: proof - haftmann@43873: from assms have "- x \ - (- y)" by (simp only: compl_less_compl_iff) haftmann@43873: then show ?thesis by simp haftmann@43873: qed huffman@36008: haftmann@31991: end haftmann@31991: haftmann@31991: haftmann@51540: subsection {* @{text "min/max"} as special case of lattice *} haftmann@51540: haftmann@54861: context linorder haftmann@54861: begin haftmann@54861: haftmann@54861: sublocale min!: semilattice_order min less_eq less haftmann@51540: + max!: semilattice_order max greater_eq greater haftmann@51540: by default (auto simp add: min_def max_def) haftmann@51540: haftmann@54861: lemma min_le_iff_disj: haftmann@54861: "min x y \ z \ x \ z \ y \ z" haftmann@54861: unfolding min_def using linear by (auto intro: order_trans) haftmann@54861: haftmann@54861: lemma le_max_iff_disj: haftmann@54861: "z \ max x y \ z \ x \ z \ y" haftmann@54861: unfolding max_def using linear by (auto intro: order_trans) haftmann@54861: haftmann@54861: lemma min_less_iff_disj: haftmann@54861: "min x y < z \ x < z \ y < z" haftmann@54861: unfolding min_def le_less using less_linear by (auto intro: less_trans) haftmann@54861: haftmann@54861: lemma less_max_iff_disj: haftmann@54861: "z < max x y \ z < x \ z < y" haftmann@54861: unfolding max_def le_less using less_linear by (auto intro: less_trans) haftmann@54861: haftmann@54861: lemma min_less_iff_conj [simp]: haftmann@54861: "z < min x y \ z < x \ z < y" haftmann@54861: unfolding min_def le_less using less_linear by (auto intro: less_trans) haftmann@54861: haftmann@54861: lemma max_less_iff_conj [simp]: haftmann@54861: "max x y < z \ x < z \ y < z" haftmann@54861: unfolding max_def le_less using less_linear by (auto intro: less_trans) haftmann@54861: haftmann@54862: lemma min_max_distrib1: haftmann@54862: "min (max b c) a = max (min b a) (min c a)" haftmann@54862: by (auto simp add: min_def max_def not_le dest: le_less_trans less_trans intro: antisym) haftmann@54862: haftmann@54862: lemma min_max_distrib2: haftmann@54862: "min a (max b c) = max (min a b) (min a c)" haftmann@54862: by (auto simp add: min_def max_def not_le dest: le_less_trans less_trans intro: antisym) haftmann@54862: haftmann@54862: lemma max_min_distrib1: haftmann@54862: "max (min b c) a = min (max b a) (max c a)" haftmann@54862: by (auto simp add: min_def max_def not_le dest: le_less_trans less_trans intro: antisym) haftmann@54862: haftmann@54862: lemma max_min_distrib2: haftmann@54862: "max a (min b c) = min (max a b) (max a c)" haftmann@54862: by (auto simp add: min_def max_def not_le dest: le_less_trans less_trans intro: antisym) haftmann@54862: haftmann@54862: lemmas min_max_distribs = min_max_distrib1 min_max_distrib2 max_min_distrib1 max_min_distrib2 haftmann@54862: haftmann@54861: lemma split_min [no_atp]: haftmann@54861: "P (min i j) \ (i \ j \ P i) \ (\ i \ j \ P j)" haftmann@54861: by (simp add: min_def) haftmann@54861: haftmann@54861: lemma split_max [no_atp]: haftmann@54861: "P (max i j) \ (i \ j \ P j) \ (\ i \ j \ P i)" haftmann@54861: by (simp add: max_def) haftmann@54861: haftmann@54861: lemma min_of_mono: haftmann@54861: fixes f :: "'a \ 'b\linorder" haftmann@54861: shows "mono f \ min (f m) (f n) = f (min m n)" haftmann@54861: by (auto simp: mono_def Orderings.min_def min_def intro: Orderings.antisym) haftmann@54861: haftmann@54861: lemma max_of_mono: haftmann@54861: fixes f :: "'a \ 'b\linorder" haftmann@54861: shows "mono f \ max (f m) (f n) = f (max m n)" haftmann@54861: by (auto simp: mono_def Orderings.max_def max_def intro: Orderings.antisym) haftmann@54861: haftmann@54861: end haftmann@54861: haftmann@51540: lemma inf_min: "inf = (min \ 'a\{semilattice_inf, linorder} \ 'a \ 'a)" haftmann@51540: by (auto intro: antisym simp add: min_def fun_eq_iff) haftmann@51540: haftmann@51540: lemma sup_max: "sup = (max \ 'a\{semilattice_sup, linorder} \ 'a \ 'a)" haftmann@51540: by (auto intro: antisym simp add: max_def fun_eq_iff) haftmann@51540: haftmann@51540: haftmann@22454: subsection {* Uniqueness of inf and sup *} haftmann@22454: haftmann@35028: lemma (in semilattice_inf) inf_unique: haftmann@22454: fixes f (infixl "\" 70) haftmann@34007: assumes le1: "\x y. x \ y \ x" and le2: "\x y. x \ y \ y" haftmann@34007: and greatest: "\x y z. x \ y \ x \ z \ x \ y \ z" haftmann@22737: shows "x \ y = x \ y" haftmann@22454: proof (rule antisym) haftmann@34007: show "x \ y \ x \ y" by (rule le_infI) (rule le1, rule le2) haftmann@22454: next haftmann@34007: have leI: "\x y z. x \ y \ x \ z \ x \ y \ z" by (blast intro: greatest) haftmann@34007: show "x \ y \ x \ y" by (rule leI) simp_all haftmann@22454: qed haftmann@22454: haftmann@35028: lemma (in semilattice_sup) sup_unique: haftmann@22454: fixes f (infixl "\" 70) haftmann@34007: assumes ge1 [simp]: "\x y. x \ x \ y" and ge2: "\x y. y \ x \ y" haftmann@34007: and least: "\x y z. y \ x \ z \ x \ y \ z \ x" haftmann@22737: shows "x \ y = x \ y" haftmann@22454: proof (rule antisym) haftmann@34007: show "x \ y \ x \ y" by (rule le_supI) (rule ge1, rule ge2) haftmann@22454: next haftmann@34007: have leI: "\x y z. x \ z \ y \ z \ x \ y \ z" by (blast intro: least) haftmann@34007: show "x \ y \ x \ y" by (rule leI) simp_all haftmann@22454: qed huffman@36008: haftmann@22454: haftmann@46631: subsection {* Lattice on @{typ bool} *} haftmann@22454: haftmann@31991: instantiation bool :: boolean_algebra haftmann@25510: begin haftmann@25510: haftmann@25510: definition haftmann@41080: bool_Compl_def [simp]: "uminus = Not" haftmann@31991: haftmann@31991: definition haftmann@41080: bool_diff_def [simp]: "A - B \ A \ \ B" haftmann@31991: haftmann@31991: definition haftmann@41080: [simp]: "P \ Q \ P \ Q" haftmann@25510: haftmann@25510: definition haftmann@41080: [simp]: "P \ Q \ P \ Q" haftmann@25510: haftmann@31991: instance proof haftmann@41080: qed auto haftmann@22454: haftmann@25510: end haftmann@25510: haftmann@32781: lemma sup_boolI1: haftmann@32781: "P \ P \ Q" haftmann@41080: by simp haftmann@32781: haftmann@32781: lemma sup_boolI2: haftmann@32781: "Q \ P \ Q" haftmann@41080: by simp haftmann@32781: haftmann@32781: lemma sup_boolE: haftmann@32781: "P \ Q \ (P \ R) \ (Q \ R) \ R" haftmann@41080: by auto haftmann@32781: haftmann@23878: haftmann@46631: subsection {* Lattice on @{typ "_ \ _"} *} haftmann@23878: nipkow@51387: instantiation "fun" :: (type, semilattice_sup) semilattice_sup haftmann@25510: begin haftmann@25510: haftmann@25510: definition haftmann@41080: "f \ g = (\x. f x \ g x)" haftmann@41080: haftmann@49769: lemma sup_apply [simp, code]: haftmann@41080: "(f \ g) x = f x \ g x" haftmann@41080: by (simp add: sup_fun_def) haftmann@25510: haftmann@32780: instance proof noschinl@46884: qed (simp_all add: le_fun_def) haftmann@23878: haftmann@25510: end haftmann@23878: nipkow@51387: instantiation "fun" :: (type, semilattice_inf) semilattice_inf nipkow@51387: begin nipkow@51387: nipkow@51387: definition nipkow@51387: "f \ g = (\x. f x \ g x)" nipkow@51387: nipkow@51387: lemma inf_apply [simp, code]: nipkow@51387: "(f \ g) x = f x \ g x" nipkow@51387: by (simp add: inf_fun_def) nipkow@51387: nipkow@51387: instance proof nipkow@51387: qed (simp_all add: le_fun_def) nipkow@51387: nipkow@51387: end nipkow@51387: nipkow@51387: instance "fun" :: (type, lattice) lattice .. nipkow@51387: haftmann@41080: instance "fun" :: (type, distrib_lattice) distrib_lattice proof noschinl@46884: qed (rule ext, simp add: sup_inf_distrib1) haftmann@31991: haftmann@34007: instance "fun" :: (type, bounded_lattice) bounded_lattice .. haftmann@34007: haftmann@31991: instantiation "fun" :: (type, uminus) uminus haftmann@31991: begin haftmann@31991: haftmann@31991: definition haftmann@31991: fun_Compl_def: "- A = (\x. - A x)" haftmann@31991: haftmann@49769: lemma uminus_apply [simp, code]: haftmann@41080: "(- A) x = - (A x)" haftmann@41080: by (simp add: fun_Compl_def) haftmann@41080: haftmann@31991: instance .. haftmann@31991: haftmann@31991: end haftmann@31991: haftmann@31991: instantiation "fun" :: (type, minus) minus haftmann@31991: begin haftmann@31991: haftmann@31991: definition haftmann@31991: fun_diff_def: "A - B = (\x. A x - B x)" haftmann@31991: haftmann@49769: lemma minus_apply [simp, code]: haftmann@41080: "(A - B) x = A x - B x" haftmann@41080: by (simp add: fun_diff_def) haftmann@41080: haftmann@31991: instance .. haftmann@31991: haftmann@31991: end haftmann@31991: haftmann@41080: instance "fun" :: (type, boolean_algebra) boolean_algebra proof noschinl@46884: qed (rule ext, simp_all add: inf_compl_bot sup_compl_top diff_eq)+ berghofe@26794: haftmann@46631: haftmann@46631: subsection {* Lattice on unary and binary predicates *} haftmann@46631: haftmann@46631: lemma inf1I: "A x \ B x \ (A \ B) x" haftmann@46631: by (simp add: inf_fun_def) haftmann@46631: haftmann@46631: lemma inf2I: "A x y \ B x y \ (A \ B) x y" haftmann@46631: by (simp add: inf_fun_def) haftmann@46631: haftmann@46631: lemma inf1E: "(A \ B) x \ (A x \ B x \ P) \ P" haftmann@46631: by (simp add: inf_fun_def) haftmann@46631: haftmann@46631: lemma inf2E: "(A \ B) x y \ (A x y \ B x y \ P) \ P" haftmann@46631: by (simp add: inf_fun_def) haftmann@46631: haftmann@46631: lemma inf1D1: "(A \ B) x \ A x" haftmann@54857: by (rule inf1E) haftmann@46631: haftmann@46631: lemma inf2D1: "(A \ B) x y \ A x y" haftmann@54857: by (rule inf2E) haftmann@46631: haftmann@46631: lemma inf1D2: "(A \ B) x \ B x" haftmann@54857: by (rule inf1E) haftmann@46631: haftmann@46631: lemma inf2D2: "(A \ B) x y \ B x y" haftmann@54857: by (rule inf2E) haftmann@46631: haftmann@46631: lemma sup1I1: "A x \ (A \ B) x" haftmann@46631: by (simp add: sup_fun_def) haftmann@46631: haftmann@46631: lemma sup2I1: "A x y \ (A \ B) x y" haftmann@46631: by (simp add: sup_fun_def) haftmann@46631: haftmann@46631: lemma sup1I2: "B x \ (A \ B) x" haftmann@46631: by (simp add: sup_fun_def) haftmann@46631: haftmann@46631: lemma sup2I2: "B x y \ (A \ B) x y" haftmann@46631: by (simp add: sup_fun_def) haftmann@46631: haftmann@46631: lemma sup1E: "(A \ B) x \ (A x \ P) \ (B x \ P) \ P" haftmann@46631: by (simp add: sup_fun_def) iprover haftmann@46631: haftmann@46631: lemma sup2E: "(A \ B) x y \ (A x y \ P) \ (B x y \ P) \ P" haftmann@46631: by (simp add: sup_fun_def) iprover haftmann@46631: haftmann@46631: text {* haftmann@46631: \medskip Classical introduction rule: no commitment to @{text A} vs haftmann@46631: @{text B}. haftmann@46631: *} haftmann@46631: haftmann@46631: lemma sup1CI: "(\ B x \ A x) \ (A \ B) x" haftmann@46631: by (auto simp add: sup_fun_def) haftmann@46631: haftmann@46631: lemma sup2CI: "(\ B x y \ A x y) \ (A \ B) x y" haftmann@46631: by (auto simp add: sup_fun_def) haftmann@46631: haftmann@46631: haftmann@25062: no_notation haftmann@46691: less_eq (infix "\" 50) and haftmann@46691: less (infix "\" 50) haftmann@25062: haftmann@21249: end haftmann@46631: