# HG changeset patch # User hoelzl # Date 1383641099 -3600 # Node ID 71c701dc5bf95e2cf0387f309272f591a06c31e9 # Parent adfc759263ab69120c69e90e59f842c49fc31c61 add SUP and INF for conditionally complete lattices diff -r adfc759263ab -r 71c701dc5bf9 src/HOL/Complete_Lattices.thy --- a/src/HOL/Complete_Lattices.thy Tue Nov 05 09:44:58 2013 +0100 +++ b/src/HOL/Complete_Lattices.thy Tue Nov 05 09:44:59 2013 +0100 @@ -20,6 +20,12 @@ definition INFI :: "'b set \ ('b \ 'a) \ 'a" where INF_def: "INFI A f = \(f ` A)" +lemma INF_image [simp]: "INFI (f`A) g = INFI A (\x. g (f x))" + by (simp add: INF_def image_image) + +lemma INF_cong: "A = B \ (\x. x \ B \ C x = D x) \ INFI A C = INFI B D" + by (simp add: INF_def image_def) + end class Sup = @@ -29,6 +35,12 @@ definition SUPR :: "'b set \ ('b \ 'a) \ 'a" where SUP_def: "SUPR A f = \(f ` A)" +lemma SUP_image [simp]: "SUPR (f`A) g = SUPR A (%x. g (f x))" + by (simp add: SUP_def image_image) + +lemma SUP_cong: "A = B \ (\x. x \ B \ C x = D x) \ SUPR A C = SUPR B D" + by (simp add: SUP_def image_def) + end text {* @@ -183,12 +195,6 @@ "\UNIV = \" by (auto intro!: antisym Sup_upper) -lemma INF_image [simp]: "(\x\f`A. g x) = (\x\A. g (f x))" - by (simp add: INF_def image_image) - -lemma SUP_image [simp]: "(\x\f`A. g x) = (\x\A. g (f x))" - by (simp add: SUP_def image_image) - lemma Inf_Sup: "\A = \{b. \a \ A. b \ a}" by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least) @@ -201,14 +207,6 @@ lemma Sup_subset_mono: "A \ B \ \A \ \B" by (auto intro: Sup_least Sup_upper) -lemma INF_cong: - "A = B \ (\x. x \ B \ C x = D x) \ (\x\A. C x) = (\x\B. D x)" - by (simp add: INF_def image_def) - -lemma SUP_cong: - "A = B \ (\x. x \ B \ C x = D x) \ (\x\A. C x) = (\x\B. D x)" - by (simp add: SUP_def image_def) - lemma Inf_mono: assumes "\b. b \ B \ \a\A. a \ b" shows "\A \ \B" diff -r adfc759263ab -r 71c701dc5bf9 src/HOL/Conditionally_Complete_Lattices.thy --- a/src/HOL/Conditionally_Complete_Lattices.thy Tue Nov 05 09:44:58 2013 +0100 +++ b/src/HOL/Conditionally_Complete_Lattices.thy Tue Nov 05 09:44:59 2013 +0100 @@ -10,10 +10,10 @@ imports Main Lubs begin -lemma Sup_fin_eq_Max: "finite X \ X \ {} \ Sup_fin X = Max X" +lemma (in linorder) Sup_fin_eq_Max: "finite X \ X \ {} \ Sup_fin X = Max X" by (induct X rule: finite_ne_induct) (simp_all add: sup_max) -lemma Inf_fin_eq_Min: "finite X \ X \ {} \ Inf_fin X = Min X" +lemma (in linorder) Inf_fin_eq_Min: "finite X \ X \ {} \ Inf_fin X = Min X" by (induct X rule: finite_ne_induct) (simp_all add: inf_min) context preorder @@ -125,6 +125,12 @@ thus "bdd_below (A \ B)" unfolding bdd_below_def .. qed +lemma bdd_above_sup[simp]: "bdd_above ((\x. sup (f x) (g x)) ` A) \ bdd_above (f`A) \ bdd_above (g`A)" + by (auto simp: bdd_above_def intro: le_supI1 le_supI2) + +lemma bdd_below_inf[simp]: "bdd_below ((\x. inf (f x) (g x)) ` A) \ bdd_below (f`A) \ bdd_below (g`A)" + by (auto simp: bdd_below_def intro: le_infI1 le_infI2) + end @@ -142,6 +148,24 @@ and cSup_least: "X \ {} \ (\x. x \ X \ x \ z) \ Sup X \ z" begin +lemma cSup_upper2: "x \ X \ y \ x \ bdd_above X \ y \ Sup X" + by (metis cSup_upper order_trans) + +lemma cInf_lower2: "x \ X \ x \ y \ bdd_below X \ Inf X \ y" + by (metis cInf_lower order_trans) + +lemma cSup_mono: "B \ {} \ bdd_above A \ (\b. b \ B \ \a\A. b \ a) \ Sup B \ Sup A" + by (metis cSup_least cSup_upper2) + +lemma cInf_mono: "B \ {} \ bdd_below A \ (\b. b \ B \ \a\A. a \ b) \ Inf A \ Inf B" + by (metis cInf_greatest cInf_lower2) + +lemma cSup_subset_mono: "A \ {} \ bdd_above B \ A \ B \ Sup A \ Sup B" + by (metis cSup_least cSup_upper subsetD) + +lemma cInf_superset_mono: "A \ {} \ bdd_below B \ A \ B \ Inf B \ Inf A" + by (metis cInf_greatest cInf_lower subsetD) + lemma cSup_eq_maximum: "z \ X \ (\x. x \ X \ x \ z) \ Sup X = z" by (intro antisym cSup_upper[of z X] cSup_least[of X z]) auto @@ -154,18 +178,6 @@ lemma le_cInf_iff: "S \ {} \ bdd_below S \ a \ Inf S \ (\x\S. a \ x)" by (metis order_trans cInf_lower cInf_greatest) -lemma cSup_singleton [simp]: "Sup {x} = x" - by (intro cSup_eq_maximum) auto - -lemma cInf_singleton [simp]: "Inf {x} = x" - by (intro cInf_eq_minimum) auto - -lemma cSup_upper2: "x \ X \ y \ x \ bdd_above X \ y \ Sup X" - by (metis cSup_upper order_trans) - -lemma cInf_lower2: "x \ X \ x \ y \ bdd_below X \ Inf X \ y" - by (metis cInf_lower order_trans) - lemma cSup_eq_non_empty: assumes 1: "X \ {}" assumes 2: "\x. x \ X \ x \ a" @@ -192,10 +204,16 @@ lemma cInf_insert: "X \ {} \ bdd_below X \ Inf (insert a X) = inf a (Inf X)" by (intro cInf_eq_non_empty) (auto intro: le_infI2 cInf_lower cInf_greatest) +lemma cSup_singleton [simp]: "Sup {x} = x" + by (intro cSup_eq_maximum) auto + +lemma cInf_singleton [simp]: "Inf {x} = x" + by (intro cInf_eq_minimum) auto + lemma cSup_insert_If: "bdd_above X \ Sup (insert a X) = (if X = {} then a else sup a (Sup X))" using cSup_insert[of X] by simp -lemma cInf_insert_if: "bdd_below X \ Inf (insert a X) = (if X = {} then a else inf a (Inf X))" +lemma cInf_insert_If: "bdd_below X \ Inf (insert a X) = (if X = {} then a else inf a (Inf X))" using cInf_insert[of X] by simp lemma le_cSup_finite: "finite X \ x \ X \ x \ Sup X" @@ -234,6 +252,74 @@ lemma cInf_atLeastAtMost[simp]: "y \ x \ Inf {y..x} = y" by (auto intro!: cInf_eq_minimum) +lemma cINF_lower: "bdd_below (f ` A) \ x \ A \ INFI A f \ f x" + unfolding INF_def by (rule cInf_lower) auto + +lemma cINF_greatest: "A \ {} \ (\x. x \ A \ m \ f x) \ m \ INFI A f" + unfolding INF_def by (rule cInf_greatest) auto + +lemma cSUP_upper: "x \ A \ bdd_above (f ` A) \ f x \ SUPR A f" + unfolding SUP_def by (rule cSup_upper) auto + +lemma cSUP_least: "A \ {} \ (\x. x \ A \ f x \ M) \ SUPR A f \ M" + unfolding SUP_def by (rule cSup_least) auto + +lemma cINF_lower2: "bdd_below (f ` A) \ x \ A \ f x \ u \ INFI A f \ u" + by (auto intro: cINF_lower assms order_trans) + +lemma cSUP_upper2: "bdd_above (f ` A) \ x \ A \ u \ f x \ u \ SUPR A f" + by (auto intro: cSUP_upper assms order_trans) + +lemma le_cINF_iff: "A \ {} \ bdd_below (f ` A) \ u \ INFI A f \ (\x\A. u \ f x)" + by (metis cINF_greatest cINF_lower assms order_trans) + +lemma cSUP_le_iff: "A \ {} \ bdd_above (f ` A) \ SUPR A f \ u \ (\x\A. f x \ u)" + by (metis cSUP_least cSUP_upper assms order_trans) + +lemma cINF_insert: "A \ {} \ bdd_below (f ` A) \ INFI (insert a A) f = inf (f a) (INFI A f)" + by (metis INF_def cInf_insert assms empty_is_image image_insert) + +lemma cSUP_insert: "A \ {} \ bdd_above (f ` A) \ SUPR (insert a A) f = sup (f a) (SUPR A f)" + by (metis SUP_def cSup_insert assms empty_is_image image_insert) + +lemma cINF_mono: "B \ {} \ bdd_below (f ` A) \ (\m. m \ B \ \n\A. f n \ g m) \ INFI A f \ INFI B g" + unfolding INF_def by (auto intro: cInf_mono) + +lemma cSUP_mono: "A \ {} \ bdd_above (g ` B) \ (\n. n \ A \ \m\B. f n \ g m) \ SUPR A f \ SUPR B g" + unfolding SUP_def by (auto intro: cSup_mono) + +lemma cINF_superset_mono: "A \ {} \ bdd_below (g ` B) \ A \ B \ (\x. x \ B \ g x \ f x) \ INFI B g \ INFI A f" + by (rule cINF_mono) auto + +lemma cSUP_subset_mono: "A \ {} \ bdd_above (g ` B) \ A \ B \ (\x. x \ B \ f x \ g x) \ SUPR A f \ SUPR B g" + by (rule cSUP_mono) auto + +lemma less_eq_cInf_inter: "bdd_below A \ bdd_below B \ A \ B \ {} \ inf (Inf A) (Inf B) \ Inf (A \ B)" + by (metis cInf_superset_mono lattice_class.inf_sup_ord(1) le_infI1) + +lemma cSup_inter_less_eq: "bdd_above A \ bdd_above B \ A \ B \ {} \ Sup (A \ B) \ sup (Sup A) (Sup B) " + by (metis cSup_subset_mono lattice_class.inf_sup_ord(1) le_supI1) + +lemma cInf_union_distrib: "A \ {} \ bdd_below A \ B \ {} \ bdd_below B \ Inf (A \ B) = inf (Inf A) (Inf B)" + by (intro antisym le_infI cInf_greatest cInf_lower) (auto intro: le_infI1 le_infI2 cInf_lower) + +lemma cINF_union: "A \ {} \ bdd_below (f`A) \ B \ {} \ bdd_below (f`B) \ INFI (A \ B) f = inf (INFI A f) (INFI B f)" + unfolding INF_def using assms by (auto simp add: image_Un intro: cInf_union_distrib) + +lemma cSup_union_distrib: "A \ {} \ bdd_above A \ B \ {} \ bdd_above B \ Sup (A \ B) = sup (Sup A) (Sup B)" + by (intro antisym le_supI cSup_least cSup_upper) (auto intro: le_supI1 le_supI2 cSup_upper) + +lemma cSUP_union: "A \ {} \ bdd_above (f`A) \ B \ {} \ bdd_above (f`B) \ SUPR (A \ B) f = sup (SUPR A f) (SUPR B f)" + unfolding SUP_def by (auto simp add: image_Un intro: cSup_union_distrib) + +lemma cINF_inf_distrib: "A \ {} \ bdd_below (f`A) \ bdd_below (g`A) \ inf (INFI A f) (INFI A g) = (INF a:A. inf (f a) (g a))" + by (intro antisym le_infI cINF_greatest cINF_lower2) + (auto intro: le_infI1 le_infI2 cINF_greatest cINF_lower le_infI) + +lemma SUP_sup_distrib: "A \ {} \ bdd_above (f`A) \ bdd_above (g`A) \ sup (SUPR A f) (SUPR A g) = (SUP a:A. sup (f a) (g a))" + by (intro antisym le_supI cSUP_least cSUP_upper2) + (auto intro: le_supI1 le_supI2 cSUP_least cSUP_upper le_supI) + end instance complete_lattice \ conditionally_complete_lattice @@ -323,14 +409,11 @@ end -class linear_continuum = conditionally_complete_linorder + dense_linorder + - assumes UNIV_not_singleton: "\a b::'a. a \ b" -begin +lemma cSup_eq_Max: "finite (X::'a::conditionally_complete_linorder set) \ X \ {} \ Sup X = Max X" + using cSup_eq_Sup_fin[of X] Sup_fin_eq_Max[of X] by simp -lemma ex_gt_or_lt: "\b. a < b \ b < a" - by (metis UNIV_not_singleton neq_iff) - -end +lemma cInf_eq_Min: "finite (X::'a::conditionally_complete_linorder set) \ X \ {} \ Inf X = Min X" + using cInf_eq_Inf_fin[of X] Inf_fin_eq_Min[of X] by simp lemma cSup_bounds: fixes S :: "'a :: conditionally_complete_lattice set" @@ -347,19 +430,12 @@ with b show ?thesis by blast qed - lemma cSup_unique: "(S::'a :: {conditionally_complete_linorder, no_bot} set) *<= b \ (\b'x\S. b' < x) \ Sup S = b" by (rule cSup_eq) (auto simp: not_le[symmetric] setle_def) lemma cInf_unique: "b <=* (S::'a :: {conditionally_complete_linorder, no_top} set) \ (\b'>b. \x\S. b' > x) \ Inf S = b" by (rule cInf_eq) (auto simp: not_le[symmetric] setge_def) -lemma cSup_eq_Max: "finite (X::'a::conditionally_complete_linorder set) \ X \ {} \ Sup X = Max X" - using cSup_eq_Sup_fin[of X] Sup_fin_eq_Max[of X] by simp - -lemma cInf_eq_Min: "finite (X::'a::conditionally_complete_linorder set) \ X \ {} \ Inf X = Min X" - using cInf_eq_Inf_fin[of X] Inf_fin_eq_Min[of X] by simp - lemma cSup_lessThan[simp]: "Sup {.. Inf {y<..a b::'a. a \ b" +begin + +lemma ex_gt_or_lt: "\b. a < b \ b < a" + by (metis UNIV_not_singleton neq_iff) + end + +end diff -r adfc759263ab -r 71c701dc5bf9 src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Tue Nov 05 09:44:58 2013 +0100 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Tue Nov 05 09:44:59 2013 +0100 @@ -2686,7 +2686,7 @@ lemma Inf_insert: fixes S :: "real set" shows "bounded S \ Inf (insert x S) = (if S = {} then x else min x (Inf S))" - by (auto simp: bounded_imp_bdd_below inf_min cInf_insert_if) + by (auto simp: bounded_imp_bdd_below inf_min cInf_insert_If) lemma Inf_insert_finite: fixes S :: "real set"