diff -r 62c033d7f3d8 -r d63ec23c9125 src/HOL/Complete_Lattices.thy --- a/src/HOL/Complete_Lattices.thy Mon Mar 04 09:57:54 2013 +0100 +++ b/src/HOL/Complete_Lattices.thy Wed Feb 20 12:04:42 2013 +0100 @@ -89,6 +89,22 @@ by (simp add: fun_eq_iff SUP_def complete_lattice.INF_def [OF dual_complete_lattice]) +lemma Sup_eqI: + "(\y. y \ A \ y \ x) \ (\y. (\z. z \ A \ z \ y) \ x \ y) \ \A = x" + by (blast intro: antisym Sup_least Sup_upper) + +lemma Inf_eqI: + "(\i. i \ A \ x \ i) \ (\y. (\i. i \ A \ y \ i) \ y \ x) \ \A = x" + by (blast intro: antisym Inf_greatest Inf_lower) + +lemma SUP_eqI: + "(\i. i \ A \ f i \ x) \ (\y. (\i. i \ A \ f i \ y) \ x \ y) \ (\i\A. f i) = x" + unfolding SUP_def by (rule Sup_eqI) auto + +lemma INF_eqI: + "(\i. i \ A \ x \ f i) \ (\y. (\i. i \ A \ f i \ y) \ x \ y) \ (\i\A. f i) = x" + unfolding INF_def by (rule Inf_eqI) auto + lemma INF_lower: "i \ A \ (\i\A. f i) \ f i" by (auto simp add: INF_def intro: Inf_lower) @@ -242,6 +258,18 @@ ultimately show ?thesis by (rule Sup_upper2) qed +lemma SUPR_eq: + assumes "\i. i \ A \ \j\B. f i \ g j" + assumes "\j. j \ B \ \i\A. g j \ f i" + shows "(SUP i:A. f i) = (SUP j:B. g j)" + by (intro antisym SUP_least) (blast intro: SUP_upper2 dest: assms)+ + +lemma INFI_eq: + assumes "\i. i \ A \ \j\B. f i \ g j" + assumes "\j. j \ B \ \i\A. g j \ f i" + shows "(INF i:A. f i) = (INF j:B. g j)" + by (intro antisym INF_greatest) (blast intro: INF_lower2 dest: assms)+ + lemma less_eq_Inf_inter: "\A \ \B \ \(A \ B)" by (auto intro: Inf_greatest Inf_lower) @@ -378,6 +406,12 @@ "(\b. A b) = A True \ A False" by (simp add: UNIV_bool SUP_insert sup_commute) +lemma Inf_le_Sup: "A \ {} \ Inf A \ Sup A" + by (blast intro: Sup_upper2 Inf_lower ex_in_conv) + +lemma INF_le_SUP: "A \ {} \ INFI A f \ SUPR A f" + unfolding INF_def SUP_def by (rule Inf_le_Sup) auto + end class complete_distrib_lattice = complete_lattice + @@ -530,9 +564,31 @@ "(\i\A. f i) = \ \ (\x>\. \i\A. f i < x)" unfolding INF_def by auto +lemma le_Sup_iff: "x \ \A \ (\ya\A. y < a)" +proof safe + fix y assume "x \ \A" "y < x" + then have "y < \A" by auto + then show "\a\A. y < a" + unfolding less_Sup_iff . +qed (auto elim!: allE[of _ "\A"] simp add: not_le[symmetric] Sup_upper) + +lemma le_SUP_iff: "x \ SUPR A f \ (\yi\A. y < f i)" + unfolding le_Sup_iff SUP_def by simp + +lemma Inf_le_iff: "\A \ x \ (\y>x. \a\A. y > a)" +proof safe + fix y assume "x \ \A" "y > x" + then have "y > \A" by auto + then show "\a\A. y > a" + unfolding Inf_less_iff . +qed (auto elim!: allE[of _ "\A"] simp add: not_le[symmetric] Inf_lower) + +lemma INF_le_iff: + "INFI A f \ x \ (\y>x. \i\A. y > f i)" + unfolding Inf_le_iff INF_def by simp + end - subsection {* Complete lattice on @{typ bool} *} instantiation bool :: complete_lattice