# HG changeset patch # User haftmann # Date 1311192879 -7200 # Node ID 26ca0bad226ae02060439e285242bb941e16b9ed # Parent aa04d1e1e2cc2999894b63e5442c4c283ea88849 class complete_linorder diff -r aa04d1e1e2cc -r 26ca0bad226a NEWS --- a/NEWS Wed Jul 20 16:15:33 2011 +0200 +++ b/NEWS Wed Jul 20 22:14:39 2011 +0200 @@ -63,15 +63,16 @@ * 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 and less -misunderstandable names: +* Class complete_lattice: generalized a couple of lemmas from sets; +generalized theorems INF_cong and SUP_cong. New type classes for complete +boolean algebras and complete linear orderes. Lemmas Inf_less_iff, +less_Sup_iff, INF_less_iff, less_SUP_iff now reside in class complete_linorder. +More consistent and less misunderstandable names: INFI_def ~> INF_def SUPR_def ~> SUP_def le_SUPI ~> le_SUP_I le_SUPI2 ~> le_SUP_I2 le_INFI ~> le_INF_I - INF_subset ~> INF_superset_mono INFI_bool_eq ~> INF_bool_eq SUPR_bool_eq ~> SUP_bool_eq INFI_apply ~> INF_apply diff -r aa04d1e1e2cc -r 26ca0bad226a src/HOL/Complete_Lattice.thy --- a/src/HOL/Complete_Lattice.thy Wed Jul 20 16:15:33 2011 +0200 +++ b/src/HOL/Complete_Lattice.thy Wed Jul 20 22:14:39 2011 +0200 @@ -292,12 +292,13 @@ by (force intro!: Sup_mono simp: SUP_def) lemma INF_superset_mono: - "B \ A \ INFI A f \ INFI B f" - by (rule INF_mono) auto + "B \ A \ (\x. x \ B \ f x \ g x) \ (\x\A. f x) \ (\x\B. g x)" + -- {* The last inclusion is POSITIVE! *} + by (blast intro: INF_mono dest: subsetD) lemma SUP_subset_mono: - "A \ B \ SUPR A f \ SUPR B f" - by (rule SUP_mono) auto + "A \ B \ (\x. x \ A \ f x \ g x) \ (\x\A. f x) \ (\x\B. g x)" + by (blast intro: SUP_mono dest: subsetD) lemma INF_commute: "(\i\A. \j\B. f i j) = (\j\B. \i\A. f i j)" by (iprover intro: INF_leI le_INF_I order_trans antisym) @@ -371,38 +372,8 @@ "(\b. A b) = A True \ A False" by (simp add: UNIV_bool SUP_empty SUP_insert sup_commute) -lemma INF_mono': - "B \ A \ (\x. x \ B \ f x \ g x) \ (\x\A. f x) \ (\x\B. g x)" - -- {* The last inclusion is POSITIVE! *} - by (rule INF_mono) auto - -lemma SUP_mono': - "A \ B \ (\x. x \ A \ f x \ g x) \ (\x\A. f x) \ (\x\B. g x)" - -- {* The last inclusion is POSITIVE! *} - by (blast intro: SUP_mono dest: subsetD) - end -lemma Inf_less_iff: - fixes a :: "'a\{complete_lattice,linorder}" - shows "\S \ a \ (\x\S. x \ a)" - unfolding not_le [symmetric] le_Inf_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 INF_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)" - unfolding SUP_def less_Sup_iff by auto - class complete_boolean_algebra = boolean_algebra + complete_lattice begin @@ -430,6 +401,27 @@ end +class complete_linorder = linorder + complete_lattice +begin + +lemma Inf_less_iff: + "\S \ a \ (\x\S. x \ a)" + unfolding not_le [symmetric] le_Inf_iff by auto + +lemma less_Sup_iff: + "a \ \S \ (\x\S. a \ x)" + unfolding not_le [symmetric] Sup_le_iff by auto + +lemma INF_less_iff: + "(\i\A. f i) \ a \ (\x\A. f x \ a)" + unfolding INF_def Inf_less_iff by auto + +lemma less_SUP_iff: + "a \ (\i\A. f i) \ (\x\A. a \ f x)" + unfolding SUP_def less_Sup_iff by auto + +end + subsection {* @{typ bool} and @{typ "_ \ _"} as complete lattice *} @@ -688,7 +680,7 @@ lemma INT_anti_mono: "A \ B \ (\x. x \ A \ f x \ g x) \ (\x\B. f x) \ (\x\A. g x)" -- {* The last inclusion is POSITIVE! *} - by (fact INF_mono') + by (fact INF_superset_mono) lemma Pow_INT_eq: "Pow (\x\A. B x) = (\x\A. Pow (B x))" by blast @@ -922,7 +914,7 @@ lemma UN_mono: "A \ B \ (\x. x \ A \ f x \ g x) \ (\x\A. f x) \ (\x\B. g x)" - by (fact SUP_mono') + by (fact SUP_subset_mono) lemma vimage_Union: "f -` (\A) = (\X\A. f -` X)" by blast @@ -1083,7 +1075,11 @@ lemmas (in complete_lattice) le_SUPI = le_SUP_I lemmas (in complete_lattice) le_SUPI2 = le_SUP_I2 lemmas (in complete_lattice) le_INFI = le_INF_I -lemmas (in complete_lattice) INF_subset = INF_superset_mono + +lemma (in complete_lattice) INF_subset: + "B \ A \ INFI A f \ INFI B f" + by (rule INF_superset_mono) auto + lemmas INFI_apply = INF_apply lemmas SUPR_apply = SUP_apply