# HG changeset patch # User haftmann # Date 1252910217 -7200 # Node ID 89518a3074e16fc72fca7b8a56dea82b06563d4c # Parent fdbfa0e35e786df6d3565d568af3b5fa5bcf895b some lemmas about strict order in lattices diff -r fdbfa0e35e78 -r 89518a3074e1 src/HOL/Lattices.thy --- a/src/HOL/Lattices.thy Fri Sep 11 09:05:26 2009 +0200 +++ b/src/HOL/Lattices.thy Mon Sep 14 08:36:57 2009 +0200 @@ -12,7 +12,9 @@ notation less_eq (infix "\" 50) and - less (infix "\" 50) + less (infix "\" 50) and + top ("\") and + bot ("\") class lower_semilattice = order + fixes inf :: "'a \ 'a \ 'a" (infixl "\" 70) @@ -220,6 +222,46 @@ end +subsubsection {* Strict order *} + +context lower_semilattice +begin + +lemma less_infI1: + "a \ x \ a \ b \ x" + by (auto simp add: less_le intro: le_infI1) + +lemma less_infI2: + "b \ x \ a \ b \ x" + by (auto simp add: less_le intro: le_infI2) + +end + +context upper_semilattice +begin + +lemma less_supI1: + "x < a \ x < a \ b" +proof - + interpret dual: lower_semilattice "op \" "op >" sup + by (fact dual_semilattice) + assume "x < a" + then show "x < a \ b" + by (fact dual.less_infI1) +qed + +lemma less_supI2: + "x < b \ x < a \ b" +proof - + interpret dual: lower_semilattice "op \" "op >" sup + by (fact dual_semilattice) + assume "x < b" + then show "x < a \ b" + by (fact dual.less_infI2) +qed + +end + subsection {* Distributive lattices *} @@ -306,6 +348,40 @@ "x \ bot = 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 sup_eq_bot_eq1: + assumes "A \ B = \" + shows "A = \" +proof - + interpret dual: boolean_algebra "\x y. x \ - y" uminus "op \" "op >" "op \" "op \" top bot + by (rule dual_boolean_algebra) + from dual.inf_eq_top_eq1 assms show ?thesis . +qed + +lemma sup_eq_bot_eq2: + assumes "A \ B = \" + shows "B = \" +proof - + interpret dual: boolean_algebra "\x y. x \ - y" uminus "op \" "op >" "op \" "op \" top bot + by (rule dual_boolean_algebra) + from dual.inf_eq_top_eq2 assms show ?thesis . +qed + lemma compl_unique: assumes "x \ y = bot" and "x \ y = top" @@ -517,6 +593,8 @@ less_eq (infix "\" 50) and less (infix "\" 50) and inf (infixl "\" 70) and - sup (infixl "\" 65) + sup (infixl "\" 65) and + top ("\") and + bot ("\") end