# HG changeset patch # User huffman # Date 1269805754 25200 # Node ID 23dfa8678c7cad6330661f6f23231507b6e4a646 # Parent 095b1022e2ae8a54781b1125548322ad8e06abfa add/change some lemmas about lattices diff -r 095b1022e2ae -r 23dfa8678c7c src/HOL/Lattices.thy --- a/src/HOL/Lattices.thy Sun Mar 28 10:34:02 2010 -0700 +++ b/src/HOL/Lattices.thy Sun Mar 28 12:49:14 2010 -0700 @@ -90,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" @@ -103,6 +103,9 @@ "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\semilattice_inf" shows "mono f \ f (A \ B) \ f A \ f B" @@ -123,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" @@ -137,6 +140,9 @@ "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\semilattice_sup" shows "mono f \ f A \ f B \ f (A \ B)" @@ -345,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 @@ -393,39 +405,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 @@ -472,10 +458,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" @@ -499,18 +482,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 @@ -522,6 +501,21 @@ 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 @@ -550,7 +544,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} *} diff -r 095b1022e2ae -r 23dfa8678c7c src/HOL/Predicate.thy --- a/src/HOL/Predicate.thy Sun Mar 28 10:34:02 2010 -0700 +++ b/src/HOL/Predicate.thy Sun Mar 28 12:49:14 2010 -0700 @@ -516,7 +516,7 @@ lemma is_empty_sup: "is_empty (A \ B) \ is_empty A \ is_empty B" - by (auto simp add: is_empty_def intro: sup_eq_bot_eq1 sup_eq_bot_eq2) + by (auto simp add: is_empty_def) definition singleton :: "(unit => 'a) \ 'a pred \ 'a" where "singleton dfault A = (if \!x. eval A x then THE x. eval A x else dfault ())"