# HG changeset patch # User haftmann # Date 1310574971 -7200 # Node ID 534c5eb522a6cde7ad4ce987de13185d757c1d80 # Parent 9959c8732edf7333206efc0352363d4b4b443a79# Parent 097732301fc0aa921615dae2346491084b0a9e91 merged diff -r 9959c8732edf -r 534c5eb522a6 src/HOL/Complete_Lattice.thy --- a/src/HOL/Complete_Lattice.thy Wed Jul 13 15:50:45 2011 +0200 +++ b/src/HOL/Complete_Lattice.thy Wed Jul 13 18:36:11 2011 +0200 @@ -392,13 +392,41 @@ lemma Inter_Un_distrib: "\(A \ B) = \A \ \B" by (fact Inf_union_distrib) -(*lemma (in complete_lattice) Inf_top_conv [no_atp]: - "\A = \ \ (\x\A. x = \)"*) +lemma (in bounded_lattice_bot) bot_less: + -- {* FIXME: tighten classes bot, top to partial orders (uniqueness!), move lemmas there *} + "a \ bot \ bot < a" + by (auto simp add: less_le_not_le intro!: antisym) + +lemma (in bounded_lattice_top) less_top: + "a \ top \ a < top" + by (auto simp add: less_le_not_le intro!: antisym) + +lemma (in complete_lattice) Inf_top_conv [no_atp]: + "\A = \ \ (\x\A. x = \)" + "\ = \A \ (\x\A. x = \)" +proof - + show "\A = \ \ (\x\A. x = \)" + proof + assume "\x\A. x = \" + then have "A = {} \ A = {\}" by auto + then show "\A = \" by auto + next + assume "\A = \" + show "\x\A. x = \" + proof (rule ccontr) + assume "\ (\x\A. x = \)" + then obtain x where "x \ A" and "x \ \" by blast + then obtain B where "A = insert x B" by blast + with `\A = \` `x \ \` show False by (simp add: Inf_insert) + qed + qed + then show "\ = \A \ (\x\A. x = \)" by auto +qed lemma Inter_UNIV_conv [simp,no_atp]: "\A = UNIV \ (\x\A. x = UNIV)" "UNIV = \A \ (\x\A. x = UNIV)" - by blast+ + by (fact Inf_top_conv)+ lemma (in complete_lattice) Inf_anti_mono: "B \ A \ \A \ \B" by (auto intro: Inf_greatest Inf_lower) @@ -448,7 +476,7 @@ lemma Inter_image_eq [simp]: "\(B`A) = (\x\A. B x)" - by (rule sym) (fact INTER_eq_Inter_image) + by (rule sym) (fact INFI_def) lemma INT_iff [simp]: "(b: (INT x:A. B x)) = (ALL x:A. b: B x)" by (unfold INTER_def) blast