diff -r 059c3568fdc8 -r abc6a2ea4b88 src/HOL/Lattices.thy --- a/src/HOL/Lattices.thy Fri Apr 02 13:33:48 2010 +0200 +++ b/src/HOL/Lattices.thy Wed Apr 07 19:17:10 2010 +0200 @@ -5,10 +5,45 @@ header {* Abstract lattices *} theory Lattices -imports Orderings +imports Orderings Groups +begin + +subsection {* Abstract semilattice *} + +text {* + This locales provide a basic structure for interpretation into + bigger structures; extensions require careful thinking, otherwise + undesired effects may occur due to interpretation. +*} + +locale semilattice = abel_semigroup + + assumes idem [simp]: "f a a = a" begin -subsection {* Lattices *} +lemma left_idem [simp]: + "f a (f a b) = f a b" + by (simp add: assoc [symmetric]) + +end + + +subsection {* Idempotent semigroup *} + +class ab_semigroup_idem_mult = ab_semigroup_mult + + assumes mult_idem: "x * x = x" + +sublocale ab_semigroup_idem_mult < times!: semilattice times proof +qed (fact mult_idem) + +context ab_semigroup_idem_mult +begin + +lemmas mult_left_idem = times.left_idem + +end + + +subsection {* Concrete lattices *} notation less_eq (infix "\" 50) and @@ -16,13 +51,13 @@ top ("\") and bot ("\") -class lower_semilattice = order + +class semilattice_inf = order + fixes inf :: "'a \ 'a \ 'a" (infixl "\" 70) assumes inf_le1 [simp]: "x \ y \ x" and inf_le2 [simp]: "x \ y \ y" and inf_greatest: "x \ y \ x \ z \ x \ y \ z" -class upper_semilattice = order + +class semilattice_sup = order + fixes sup :: "'a \ 'a \ 'a" (infixl "\" 65) assumes sup_ge1 [simp]: "x \ x \ y" and sup_ge2 [simp]: "y \ x \ y" @@ -32,18 +67,18 @@ text {* Dual lattice *} lemma dual_semilattice: - "lower_semilattice (op \) (op >) sup" -by (rule lower_semilattice.intro, rule dual_order) + "semilattice_inf (op \) (op >) sup" +by (rule semilattice_inf.intro, rule dual_order) (unfold_locales, simp_all add: sup_least) end -class lattice = lower_semilattice + upper_semilattice +class lattice = semilattice_inf + semilattice_sup subsubsection {* Intro and elim rules*} -context lower_semilattice +context semilattice_inf begin lemma le_infI1: @@ -55,10 +90,10 @@ by (rule order_trans) auto lemma le_infI: "x \ a \ x \ b \ x \ a \ b" - by (blast intro: inf_greatest) + by (rule inf_greatest) (* FIXME: duplicate lemma *) lemma le_infE: "x \ a \ b \ (x \ a \ x \ b \ P) \ P" - by (blast intro: order_trans le_infI1 le_infI2) + by (blast intro: order_trans inf_le1 inf_le2) lemma le_inf_iff [simp]: "x \ y \ z \ x \ y \ x \ z" @@ -68,14 +103,17 @@ "x \ y \ x \ y = x" by (auto intro: le_infI1 antisym dest: eq_iff [THEN iffD1]) +lemma inf_mono: "a \ c \ b \ d \ a \ b \ c \ d" + by (fast intro: inf_greatest le_infI1 le_infI2) + lemma mono_inf: - fixes f :: "'a \ 'b\lower_semilattice" + fixes f :: "'a \ 'b\semilattice_inf" shows "mono f \ f (A \ B) \ f A \ f B" by (auto simp add: mono_def intro: Lattices.inf_greatest) end -context upper_semilattice +context semilattice_sup begin lemma le_supI1: @@ -88,11 +126,11 @@ lemma le_supI: "a \ x \ b \ x \ a \ b \ x" - by (blast intro: sup_least) + by (rule sup_least) (* FIXME: duplicate lemma *) lemma le_supE: "a \ b \ x \ (a \ x \ b \ x \ P) \ P" - by (blast intro: le_supI1 le_supI2 order_trans) + by (blast intro: order_trans sup_ge1 sup_ge2) lemma le_sup_iff [simp]: "x \ y \ z \ x \ z \ y \ z" @@ -102,8 +140,11 @@ "x \ y \ x \ y = y" by (auto intro: le_supI2 antisym dest: eq_iff [THEN iffD1]) +lemma sup_mono: "a \ c \ b \ d \ a \ b \ c \ d" + by (fast intro: sup_least le_supI1 le_supI2) + lemma mono_sup: - fixes f :: "'a \ 'b\upper_semilattice" + fixes f :: "'a \ 'b\semilattice_sup" shows "mono f \ f A \ f B \ f (A \ B)" by (auto simp add: mono_def intro: Lattices.sup_least) @@ -112,48 +153,73 @@ subsubsection {* Equational laws *} -context lower_semilattice -begin +sublocale semilattice_inf < inf!: semilattice inf +proof + fix a b c + show "(a \ b) \ c = a \ (b \ c)" + by (rule antisym) (auto intro: le_infI1 le_infI2) + show "a \ b = b \ a" + by (rule antisym) auto + show "a \ a = a" + by (rule antisym) auto +qed -lemma inf_commute: "(x \ y) = (y \ x)" - by (rule antisym) auto +context semilattice_inf +begin lemma inf_assoc: "(x \ y) \ z = x \ (y \ z)" - by (rule antisym) (auto intro: le_infI1 le_infI2) + by (fact inf.assoc) + +lemma inf_commute: "(x \ y) = (y \ x)" + by (fact inf.commute) -lemma inf_idem[simp]: "x \ x = x" - by (rule antisym) auto +lemma inf_left_commute: "x \ (y \ z) = y \ (x \ z)" + by (fact inf.left_commute) -lemma inf_left_idem[simp]: "x \ (x \ y) = x \ y" - by (rule antisym) (auto intro: le_infI2) +lemma inf_idem: "x \ x = x" + by (fact inf.idem) + +lemma inf_left_idem: "x \ (x \ y) = x \ y" + by (fact inf.left_idem) lemma inf_absorb1: "x \ y \ x \ y = x" by (rule antisym) auto lemma inf_absorb2: "y \ x \ x \ y = y" by (rule antisym) auto - -lemma inf_left_commute: "x \ (y \ z) = y \ (x \ z)" - by (rule mk_left_commute [of inf]) (fact inf_assoc inf_commute)+ - + lemmas inf_aci = inf_commute inf_assoc inf_left_commute inf_left_idem end -context upper_semilattice -begin +sublocale semilattice_sup < sup!: semilattice sup +proof + fix a b c + show "(a \ b) \ c = a \ (b \ c)" + by (rule antisym) (auto intro: le_supI1 le_supI2) + show "a \ b = b \ a" + by (rule antisym) auto + show "a \ a = a" + by (rule antisym) auto +qed -lemma sup_commute: "(x \ y) = (y \ x)" - by (rule antisym) auto +context semilattice_sup +begin lemma sup_assoc: "(x \ y) \ z = x \ (y \ z)" - by (rule antisym) (auto intro: le_supI1 le_supI2) + by (fact sup.assoc) + +lemma sup_commute: "(x \ y) = (y \ x)" + by (fact sup.commute) -lemma sup_idem[simp]: "x \ x = x" - by (rule antisym) auto +lemma sup_left_commute: "x \ (y \ z) = y \ (x \ z)" + by (fact sup.left_commute) -lemma sup_left_idem[simp]: "x \ (x \ y) = x \ y" - by (rule antisym) (auto intro: le_supI2) +lemma sup_idem: "x \ x = x" + by (fact sup.idem) + +lemma sup_left_idem: "x \ (x \ y) = x \ y" + by (fact sup.left_idem) lemma sup_absorb1: "y \ x \ x \ y = x" by (rule antisym) auto @@ -161,9 +227,6 @@ lemma sup_absorb2: "x \ y \ x \ y = y" by (rule antisym) auto -lemma sup_left_commute: "x \ (y \ z) = y \ (x \ z)" - by (rule mk_left_commute [of sup]) (fact sup_assoc sup_commute)+ - lemmas sup_aci = sup_commute sup_assoc sup_left_commute sup_left_idem end @@ -173,7 +236,7 @@ lemma dual_lattice: "lattice (op \) (op >) sup inf" - by (rule lattice.intro, rule dual_semilattice, rule upper_semilattice.intro, rule dual_order) + by (rule lattice.intro, rule dual_semilattice, rule semilattice_sup.intro, rule dual_order) (unfold_locales, auto) lemma inf_sup_absorb: "x \ (x \ y) = x" @@ -224,7 +287,7 @@ subsubsection {* Strict order *} -context lower_semilattice +context semilattice_inf begin lemma less_infI1: @@ -237,13 +300,13 @@ end -context upper_semilattice +context semilattice_sup begin lemma less_supI1: "x \ a \ x \ a \ b" proof - - interpret dual: lower_semilattice "op \" "op >" sup + interpret dual: semilattice_inf "op \" "op >" sup by (fact dual_semilattice) assume "x \ a" then show "x \ a \ b" @@ -253,7 +316,7 @@ lemma less_supI2: "x \ b \ x \ a \ b" proof - - interpret dual: lower_semilattice "op \" "op >" sup + interpret dual: semilattice_inf "op \" "op >" sup by (fact dual_semilattice) assume "x \ b" then show "x \ a \ b" @@ -288,6 +351,12 @@ by (rule distrib_lattice.intro, rule dual_lattice) (unfold_locales, fact inf_sup_distrib1) +lemmas sup_inf_distrib = + sup_inf_distrib1 sup_inf_distrib2 + +lemmas inf_sup_distrib = + inf_sup_distrib1 inf_sup_distrib2 + lemmas distrib = sup_inf_distrib1 sup_inf_distrib2 inf_sup_distrib1 inf_sup_distrib2 @@ -335,39 +404,13 @@ "x \ \ = x" by (rule sup_absorb1) simp -lemma inf_eq_top_eq1: - assumes "A \ B = \" - shows "A = \" -proof (cases "B = \") - case True with assms show ?thesis by simp -next - case False with top_greatest have "B \ \" by (auto intro: neq_le_trans) - then have "A \ B \ \" by (rule less_infI2) - with assms show ?thesis by simp -qed - -lemma inf_eq_top_eq2: - assumes "A \ B = \" - shows "B = \" - by (rule inf_eq_top_eq1, unfold inf_commute [of B]) (fact assms) +lemma inf_eq_top_iff [simp]: + "x \ y = \ \ x = \ \ y = \" + by (simp add: eq_iff) -lemma sup_eq_bot_eq1: - assumes "A \ B = \" - shows "A = \" -proof - - interpret dual: bounded_lattice "op \" "op >" "op \" "op \" \ \ - by (rule dual_bounded_lattice) - from dual.inf_eq_top_eq1 assms show ?thesis . -qed - -lemma sup_eq_bot_eq2: - assumes "A \ B = \" - shows "B = \" -proof - - interpret dual: bounded_lattice "op \" "op >" "op \" "op \" \ \ - by (rule dual_bounded_lattice) - from dual.inf_eq_top_eq2 assms show ?thesis . -qed +lemma sup_eq_bot_iff [simp]: + "x \ y = \ \ x = \ \ y = \" + by (simp add: eq_iff) end @@ -414,10 +457,7 @@ "- x = - y \ x = y" proof assume "- x = - y" - then have "- x \ y = \" - and "- x \ y = \" - by (simp_all add: compl_inf_bot compl_sup_top) - then have "- (- x) = y" by (rule compl_unique) + then have "- (- x) = - (- y)" by (rule arg_cong) then show "x = y" by simp next assume "x = y" @@ -441,18 +481,14 @@ 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) = \" + 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) = (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) = \" + 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 @@ -464,12 +500,27 @@ then show ?thesis by simp qed +lemma compl_mono: + "x \ y \ - y \ - x" +proof - + assume "x \ y" + then 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 "- y \ - x" by (simp only: le_iff_inf) +qed + +lemma compl_le_compl_iff: (* TODO: declare [simp] ? *) + "- x \ - y \ y \ x" +by (auto dest: compl_mono) + end subsection {* Uniqueness of inf and sup *} -lemma (in lower_semilattice) inf_unique: +lemma (in semilattice_inf) inf_unique: fixes f (infixl "\" 70) assumes le1: "\x y. x \ y \ x" and le2: "\x y. x \ y \ y" and greatest: "\x y z. x \ y \ x \ z \ x \ y \ z" @@ -481,7 +532,7 @@ show "x \ y \ x \ y" by (rule leI) simp_all qed -lemma (in upper_semilattice) sup_unique: +lemma (in semilattice_sup) sup_unique: fixes f (infixl "\" 70) assumes ge1 [simp]: "\x y. x \ x \ y" and ge2: "\x y. y \ x \ y" and least: "\x y z. y \ x \ z \ x \ y \ z \ x" @@ -492,7 +543,7 @@ have leI: "\x y z. x \ z \ y \ z \ x \ y \ z" by (blast intro: least) show "x \ y \ x \ y" by (rule leI) simp_all qed - + subsection {* @{const min}/@{const max} on linear orders as special case of @{const inf}/@{const sup} *} @@ -504,20 +555,21 @@ by (auto simp add: min_def max_def) qed (auto simp add: min_def max_def not_le less_imp_le) -lemma inf_min: "inf = (min \ 'a\{lower_semilattice, linorder} \ 'a \ 'a)" +lemma inf_min: "inf = (min \ 'a\{semilattice_inf, linorder} \ 'a \ 'a)" by (rule ext)+ (auto intro: antisym) -lemma sup_max: "sup = (max \ 'a\{upper_semilattice, linorder} \ 'a \ 'a)" +lemma sup_max: "sup = (max \ 'a\{semilattice_sup, linorder} \ 'a \ 'a)" by (rule ext)+ (auto intro: antisym) lemmas le_maxI1 = min_max.sup_ge1 lemmas le_maxI2 = min_max.sup_ge2 -lemmas max_ac = min_max.sup_assoc min_max.sup_commute - mk_left_commute [of max, OF min_max.sup_assoc min_max.sup_commute] +lemmas min_ac = min_max.inf_assoc min_max.inf_commute + min_max.inf.left_commute -lemmas min_ac = min_max.inf_assoc min_max.inf_commute - mk_left_commute [of min, OF min_max.inf_assoc min_max.inf_commute] +lemmas max_ac = min_max.sup_assoc min_max.sup_commute + min_max.sup.left_commute + subsection {* Bool as lattice *}