# HG changeset patch # User haftmann # Date 1310908558 -7200 # Node ID db18f4d0cc7dbf2863426b457124f3b55c04714f # Parent 01b13e9a1a7eb90b9841ef117d538f4b4cbd5174 further generalization from sets to complete lattices diff -r 01b13e9a1a7e -r db18f4d0cc7d NEWS --- a/NEWS Sun Jul 17 08:45:06 2011 +0200 +++ b/NEWS Sun Jul 17 15:15:58 2011 +0200 @@ -63,6 +63,11 @@ * Classes bot and top require underlying partial order rather than preorder: uniqueness of bot and top is guaranteed. INCOMPATIBILITY. +* Class 'complete_lattice': generalized a couple of lemmas from sets; +generalized theorems INF_cong and SUP_cong. More consistent names: TBD. + +INCOMPATIBILITY. + * Archimedian_Field.thy: floor now is defined as parameter of a separate type class floor_ceiling. diff -r 01b13e9a1a7e -r db18f4d0cc7d src/HOL/Complete_Lattice.thy --- a/src/HOL/Complete_Lattice.thy Sun Jul 17 08:45:06 2011 +0200 +++ b/src/HOL/Complete_Lattice.thy Sun Jul 17 15:15:58 2011 +0200 @@ -152,66 +152,74 @@ context complete_lattice begin -lemma SUP_cong: "(\x. x \ A \ f x = g x) \ SUPR A f = SUPR A g" - by (simp add: SUPR_def cong: image_cong) - -lemma INF_cong: "(\x. x \ A \ f x = g x) \ INFI A f = INFI A g" - by (simp add: INFI_def cong: image_cong) +lemma INF_empty: "(\x\{}. f x) = \" + by (simp add: INFI_def) -lemma le_SUPI: "i \ A \ M i \ (\i\A. M i)" - by (auto simp add: SUPR_def intro: Sup_upper) +lemma INF_insert: "(\x\insert a A. f x) = f a \ INFI A f" + by (simp add: INFI_def Inf_insert) -lemma le_SUPI2: "i \ A \ u \ M i \ u \ (\i\A. M i)" - using le_SUPI[of i A M] by auto - -lemma SUP_leI: "(\i. i \ A \ M i \ u) \ (\i\A. M i) \ u" - by (auto simp add: SUPR_def intro: Sup_least) - -lemma INF_leI: "i \ A \ (\i\A. M i) \ M i" +lemma INF_leI: "i \ A \ (\i\A. f i) \ f i" by (auto simp add: INFI_def intro: Inf_lower) -lemma INF_leI2: "i \ A \ M i \ u \ (\i\A. M i) \ u" - using INF_leI[of i A M] by auto +lemma INF_leI2: "i \ A \ f i \ u \ (\i\A. f i) \ u" + using INF_leI [of i A f] by auto -lemma le_INFI: "(\i. i \ A \ u \ M i) \ u \ (\i\A. M i)" +lemma le_INFI: "(\i. i \ A \ u \ f i) \ u \ (\i\A. f i)" by (auto simp add: INFI_def intro: Inf_greatest) -lemma SUP_le_iff: "(\i\A. M i) \ u \ (\i \ A. M i \ u)" - unfolding SUPR_def by (auto simp add: Sup_le_iff) +lemma le_INF_iff: "u \ (\i\A. f i) \ (\i \ A. u \ f i)" + by (auto simp add: INFI_def le_Inf_iff) -lemma le_INF_iff: "u \ (\i\A. M i) \ (\i \ A. u \ M i)" - unfolding INFI_def by (auto simp add: le_Inf_iff) - -lemma INF_const[simp]: "A \ {} \ (\i\A. M) = M" +lemma INF_const [simp]: "A \ {} \ (\i\A. f) = f" by (auto intro: antisym INF_leI le_INFI) -lemma SUP_const[simp]: "A \ {} \ (\i\A. M) = M" - by (auto intro: antisym SUP_leI le_SUPI) +lemma INF_cong: + "A = B \ (\x. x \ B \ C x = D x) \ (\x\A. C x) = (\x\B. D x)" + by (simp add: INFI_def image_def) lemma INF_mono: "(\m. m \ B \ \n\A. f n \ g m) \ (\n\A. f n) \ (\n\B. g n)" by (force intro!: Inf_mono simp: INFI_def) +lemma INF_subset: "A \ B \ INFI B f \ INFI A f" + by (intro INF_mono) auto + +lemma INF_commute: "(\i\A. \j\B. f i j) = (\j\B. \i\A. f i j)" + by (iprover intro: INF_leI le_INFI order_trans antisym) + +lemma SUP_cong: + "A = B \ (\x. x \ B \ C x = D x) \ (\x\A. C x) = (\x\B. D x)" + by (simp add: SUPR_def image_def) + +lemma le_SUPI: "i \ A \ f i \ (\i\A. f i)" + by (auto simp add: SUPR_def intro: Sup_upper) + +lemma le_SUPI2: "i \ A \ u \ f i \ u \ (\i\A. f i)" + using le_SUPI [of i A f] by auto + +lemma SUP_leI: "(\i. i \ A \ f i \ u) \ (\i\A. f i) \ u" + by (auto simp add: SUPR_def intro: Sup_least) + +lemma SUP_le_iff: "(\i\A. f i) \ u \ (\i \ A. f i \ u)" + unfolding SUPR_def by (auto simp add: Sup_le_iff) + +lemma SUP_const [simp]: "A \ {} \ (\i\A. f) = f" + by (auto intro: antisym SUP_leI le_SUPI) + lemma SUP_mono: "(\n. n \ A \ \m\B. f n \ g m) \ (\n\A. f n) \ (\n\B. g n)" by (force intro!: Sup_mono simp: SUPR_def) -lemma INF_subset: "A \ B \ INFI B f \ INFI A f" - by (intro INF_mono) auto - lemma SUP_subset: "A \ B \ SUPR A f \ SUPR B f" by (intro SUP_mono) auto -lemma INF_commute: "(\i\A. \j\B. f i j) = (\j\B. \i\A. f i j)" - by (iprover intro: INF_leI le_INFI order_trans antisym) - lemma SUP_commute: "(\i\A. \j\B. f i j) = (\j\B. \i\A. f i j)" by (iprover intro: SUP_leI le_SUPI order_trans antisym) -lemma INFI_insert: "(\x\insert a A. B x) = B a \ INFI A B" - by (simp add: INFI_def Inf_insert) +lemma SUP_empty: "(\x\{}. f x) = \" + by (simp add: SUPR_def) -lemma SUPR_insert: "(\x\insert a A. B x) = B a \ SUPR A B" +lemma SUP_insert: "(\x\insert a A. f x) = f a \ SUPR A f" by (simp add: SUPR_def Sup_insert) end @@ -221,16 +229,16 @@ shows "\S \ a \ (\x\S. x \ a)" unfolding not_le [symmetric] le_Inf_iff by auto +lemma INF_less_iff: + fixes a :: "'a::{complete_lattice,linorder}" + shows "(\i\A. f i) \ a \ (\x\A. f x \ a)" + unfolding INFI_def Inf_less_iff by auto + lemma less_Sup_iff: fixes a :: "'a\{complete_lattice,linorder}" shows "a \ \S \ (\x\S. a \ x)" unfolding not_le [symmetric] Sup_le_iff by auto -lemma INF_less_iff: - fixes a :: "'a::{complete_lattice,linorder}" - shows "(\i\A. f i) \ a \ (\x\A. f x \ a)" - unfolding INFI_def Inf_less_iff by auto - lemma less_SUP_iff: fixes a :: "'a::{complete_lattice,linorder}" shows "a \ (\i\A. f i) \ (\x\A. a \ f x)" @@ -474,13 +482,9 @@ -- {* "Classical" elimination -- by the Excluded Middle on @{prop "a\A"}. *} by (unfold INTER_def) blast -lemma (in complete_lattice) INFI_cong: - "A = B \ (\x. x \ B \ C x = D x) \ (\x\A. C x) = (\x\B. D x)" - by (simp add: INFI_def image_def) - lemma INT_cong [cong]: "A = B \ (\x. x \ B \ C x = D x) \ (\x\A. C x) = (\x\B. D x)" - by (fact INFI_cong) + by (fact INF_cong) lemma Collect_ball_eq: "{x. \y\A. P x y} = (\y\A. {x. P x y})" by blast @@ -506,7 +510,7 @@ shows "A k \ (\i\I. A i) = (\i\I. A i)" proof - from assms obtain J where "I = insert k J" by blast - then show ?thesis by (simp add: INFI_insert) + then show ?thesis by (simp add: INF_insert) qed lemma INT_absorb: "k \ I \ A k \ (\i\I. A i) = (\i\I. A i)" @@ -516,41 +520,74 @@ by (fact le_INF_iff) lemma INT_insert [simp]: "(\x \ insert a A. B x) = B a \ INTER A B" - by (fact INFI_insert) + by (fact INF_insert) + +lemma (in complete_lattice) INF_union: + "(\i \ A \ B. M i) = (\i \ A. M i) \ (\i\B. M i)" + by (auto intro!: antisym INF_mono intro: le_infI1 le_infI2 le_INFI INF_leI) + +lemma INT_Un: "(\i \ A \ B. M i) = (\i \ A. M i) \ (\i\B. M i)" + by (fact INF_union) + +lemma INT_insert_distrib: + "u \ A \ (\x\A. insert a (B x)) = insert a (\x\A. B x)" + by blast -- {* continue generalization from here *} -lemma INT_Un: "(\i \ A \ B. M i) = (\i \ A. M i) \ (\i\B. M i)" - by blast - -lemma INT_insert_distrib: - "u \ A \ (\x\A. insert a (B x)) = insert a (\x\A. B x)" - by blast +lemma (in complete_lattice) INF_constant: + "(\y\A. c) = (if A = {} then \ else c)" + by (simp add: INF_empty) lemma INT_constant [simp]: "(\y\A. c) = (if A = {} then UNIV else c)" - by auto + by (fact INF_constant) + +lemma (in complete_lattice) INF_eq: + "(\x\A. B x) = \({Y. \x\A. Y = B x})" + by (simp add: INFI_def image_def) lemma INT_eq: "(\x\A. B x) = \({Y. \x\A. Y = B x})" -- {* Look: it has an \emph{existential} quantifier *} - by blast + by (fact INF_eq) + +lemma (in complete_lattice) INF_top_conv: + "\ = (\x\A. B x) \ (\x\A. B x = \)" + "(\x\A. B x) = \ \ (\x\A. B x = \)" + by (auto simp add: INFI_def Inf_top_conv) lemma INTER_UNIV_conv [simp]: "(UNIV = (\x\A. B x)) = (\x\A. B x = UNIV)" "((\x\A. B x) = UNIV) = (\x\A. B x = UNIV)" - by blast+ + by (fact INF_top_conv)+ + +lemma (in complete_lattice) INFI_UNIV_range: + "(\x\UNIV. f x) = \range f" + by (simp add: INFI_def) + +lemma UNIV_bool [no_atp]: -- "FIXME move" + "UNIV = {False, True}" + by auto -lemma INT_bool_eq: "(\b. A b) = (A True \ A False)" - by (auto intro: bool_induct) +lemma (in complete_lattice) INF_bool_eq: + "(\b. A b) = A True \ A False" + by (simp add: UNIV_bool INF_empty INF_insert inf_commute) + +lemma INT_bool_eq: "(\b. A b) = A True \ A False" + by (fact INF_bool_eq) + +lemma (in complete_lattice) INF_anti_mono: + "B \ A \ (\x. x \ A \ f x \ g x) \ (\x\A. f x) \ (\x\A. g x)" + -- {* The last inclusion is POSITIVE! *} + by (blast dest: subsetD) + +lemma INT_anti_mono: + "B \ A \ (\x. x \ A \ f x \ g x) \ (\x\A. f x) \ (\x\A. g x)" + -- {* The last inclusion is POSITIVE! *} + by (blast dest: subsetD) lemma Pow_INT_eq: "Pow (\x\A. B x) = (\x\A. Pow (B x))" by blast -lemma INT_anti_mono: - "B \ A \ (\x. x \ A \ f x \ g x) \ - (\x\A. f x) \ (\x\A. g x)" - -- {* The last inclusion is POSITIVE! *} - by (blast dest: subsetD) - lemma vimage_INT: "f -` (\x\A. B x) = (\x\A. f -` B x)" by blast