wenzelm@52265: (* Title: HOL/Conditionally_Complete_Lattices.thy hoelzl@51518: Author: Amine Chaieb and L C Paulson, University of Cambridge hoelzl@51643: Author: Johannes Hölzl, TU München hoelzl@54258: Author: Luke S. Serafin, Carnegie Mellon University hoelzl@51518: *) paulson@33269: hoelzl@51773: header {* Conditionally-complete Lattices *} paulson@33269: hoelzl@51773: theory Conditionally_Complete_Lattices hoelzl@54263: imports Main paulson@33269: begin paulson@33269: hoelzl@54259: lemma (in linorder) Sup_fin_eq_Max: "finite X \ X \ {} \ Sup_fin X = Max X" hoelzl@51475: by (induct X rule: finite_ne_induct) (simp_all add: sup_max) hoelzl@51475: hoelzl@54259: lemma (in linorder) Inf_fin_eq_Min: "finite X \ X \ {} \ Inf_fin X = Min X" hoelzl@51475: by (induct X rule: finite_ne_induct) (simp_all add: inf_min) hoelzl@51475: hoelzl@54258: context preorder hoelzl@54258: begin hoelzl@54258: hoelzl@54258: definition "bdd_above A \ (\M. \x \ A. x \ M)" hoelzl@54258: definition "bdd_below A \ (\m. \x \ A. m \ x)" hoelzl@54258: hoelzl@54258: lemma bdd_aboveI[intro]: "(\x. x \ A \ x \ M) \ bdd_above A" hoelzl@54258: by (auto simp: bdd_above_def) hoelzl@54258: hoelzl@54258: lemma bdd_belowI[intro]: "(\x. x \ A \ m \ x) \ bdd_below A" hoelzl@54258: by (auto simp: bdd_below_def) hoelzl@54258: hoelzl@54263: lemma bdd_aboveI2: "(\x. x \ A \ f x \ M) \ bdd_above (f`A)" hoelzl@54263: by force hoelzl@54263: hoelzl@54263: lemma bdd_belowI2: "(\x. x \ A \ m \ f x) \ bdd_below (f`A)" hoelzl@54263: by force hoelzl@54263: hoelzl@54258: lemma bdd_above_empty [simp, intro]: "bdd_above {}" hoelzl@54258: unfolding bdd_above_def by auto hoelzl@54258: hoelzl@54258: lemma bdd_below_empty [simp, intro]: "bdd_below {}" hoelzl@54258: unfolding bdd_below_def by auto hoelzl@54258: hoelzl@54258: lemma bdd_above_mono: "bdd_above B \ A \ B \ bdd_above A" hoelzl@54258: by (metis (full_types) bdd_above_def order_class.le_neq_trans psubsetD) hoelzl@54258: hoelzl@54258: lemma bdd_below_mono: "bdd_below B \ A \ B \ bdd_below A" hoelzl@54258: by (metis bdd_below_def order_class.le_neq_trans psubsetD) hoelzl@54258: hoelzl@54258: lemma bdd_above_Int1 [simp]: "bdd_above A \ bdd_above (A \ B)" hoelzl@54258: using bdd_above_mono by auto hoelzl@54258: hoelzl@54258: lemma bdd_above_Int2 [simp]: "bdd_above B \ bdd_above (A \ B)" hoelzl@54258: using bdd_above_mono by auto hoelzl@54258: hoelzl@54258: lemma bdd_below_Int1 [simp]: "bdd_below A \ bdd_below (A \ B)" hoelzl@54258: using bdd_below_mono by auto hoelzl@54258: hoelzl@54258: lemma bdd_below_Int2 [simp]: "bdd_below B \ bdd_below (A \ B)" hoelzl@54258: using bdd_below_mono by auto hoelzl@54258: hoelzl@54258: lemma bdd_above_Ioo [simp, intro]: "bdd_above {a <..< b}" hoelzl@54258: by (auto simp add: bdd_above_def intro!: exI[of _ b] less_imp_le) hoelzl@54258: hoelzl@54258: lemma bdd_above_Ico [simp, intro]: "bdd_above {a ..< b}" hoelzl@54258: by (auto simp add: bdd_above_def intro!: exI[of _ b] less_imp_le) hoelzl@54258: hoelzl@54258: lemma bdd_above_Iio [simp, intro]: "bdd_above {..< b}" hoelzl@54258: by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le) hoelzl@54258: hoelzl@54258: lemma bdd_above_Ioc [simp, intro]: "bdd_above {a <.. b}" hoelzl@54258: by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le) hoelzl@54258: hoelzl@54258: lemma bdd_above_Icc [simp, intro]: "bdd_above {a .. b}" hoelzl@54258: by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le) hoelzl@54258: hoelzl@54258: lemma bdd_above_Iic [simp, intro]: "bdd_above {.. b}" hoelzl@54258: by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le) hoelzl@54258: hoelzl@54258: lemma bdd_below_Ioo [simp, intro]: "bdd_below {a <..< b}" hoelzl@54258: by (auto simp add: bdd_below_def intro!: exI[of _ a] less_imp_le) hoelzl@54258: hoelzl@54258: lemma bdd_below_Ioc [simp, intro]: "bdd_below {a <.. b}" hoelzl@54258: by (auto simp add: bdd_below_def intro!: exI[of _ a] less_imp_le) hoelzl@54258: hoelzl@54258: lemma bdd_below_Ioi [simp, intro]: "bdd_below {a <..}" hoelzl@54258: by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le) hoelzl@54258: hoelzl@54258: lemma bdd_below_Ico [simp, intro]: "bdd_below {a ..< b}" hoelzl@54258: by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le) hoelzl@54258: hoelzl@54258: lemma bdd_below_Icc [simp, intro]: "bdd_below {a .. b}" hoelzl@54258: by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le) hoelzl@54258: hoelzl@54258: lemma bdd_below_Ici [simp, intro]: "bdd_below {a ..}" hoelzl@54258: by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le) hoelzl@54258: hoelzl@54258: end hoelzl@54258: hoelzl@54261: lemma (in order_top) bdd_above_top[simp, intro!]: "bdd_above A" hoelzl@54261: by (rule bdd_aboveI[of _ top]) simp hoelzl@54261: hoelzl@54261: lemma (in order_bot) bdd_above_bot[simp, intro!]: "bdd_below A" hoelzl@54261: by (rule bdd_belowI[of _ bot]) simp hoelzl@54261: hoelzl@54262: lemma bdd_above_uminus[simp]: hoelzl@54262: fixes X :: "'a::ordered_ab_group_add set" hoelzl@54262: shows "bdd_above (uminus ` X) \ bdd_below X" hoelzl@54262: by (auto simp: bdd_above_def bdd_below_def intro: le_imp_neg_le) (metis le_imp_neg_le minus_minus) hoelzl@54262: hoelzl@54262: lemma bdd_below_uminus[simp]: hoelzl@54262: fixes X :: "'a::ordered_ab_group_add set" hoelzl@54262: shows"bdd_below (uminus ` X) \ bdd_above X" hoelzl@54262: by (auto simp: bdd_above_def bdd_below_def intro: le_imp_neg_le) (metis le_imp_neg_le minus_minus) hoelzl@54262: hoelzl@54258: context lattice hoelzl@54258: begin hoelzl@54258: hoelzl@54258: lemma bdd_above_insert [simp]: "bdd_above (insert a A) = bdd_above A" hoelzl@54258: by (auto simp: bdd_above_def intro: le_supI2 sup_ge1) hoelzl@54258: hoelzl@54258: lemma bdd_below_insert [simp]: "bdd_below (insert a A) = bdd_below A" hoelzl@54258: by (auto simp: bdd_below_def intro: le_infI2 inf_le1) hoelzl@54258: hoelzl@54258: lemma bdd_finite [simp]: hoelzl@54258: assumes "finite A" shows bdd_above_finite: "bdd_above A" and bdd_below_finite: "bdd_below A" hoelzl@54258: using assms by (induct rule: finite_induct, auto) hoelzl@54258: hoelzl@54258: lemma bdd_above_Un [simp]: "bdd_above (A \ B) = (bdd_above A \ bdd_above B)" hoelzl@54258: proof hoelzl@54258: assume "bdd_above (A \ B)" hoelzl@54258: thus "bdd_above A \ bdd_above B" unfolding bdd_above_def by auto hoelzl@54258: next hoelzl@54258: assume "bdd_above A \ bdd_above B" hoelzl@54258: then obtain a b where "\x\A. x \ a" "\x\B. x \ b" unfolding bdd_above_def by auto hoelzl@54258: hence "\x \ A \ B. x \ sup a b" by (auto intro: Un_iff le_supI1 le_supI2) hoelzl@54258: thus "bdd_above (A \ B)" unfolding bdd_above_def .. hoelzl@54258: qed hoelzl@54258: hoelzl@54258: lemma bdd_below_Un [simp]: "bdd_below (A \ B) = (bdd_below A \ bdd_below B)" hoelzl@54258: proof hoelzl@54258: assume "bdd_below (A \ B)" hoelzl@54258: thus "bdd_below A \ bdd_below B" unfolding bdd_below_def by auto hoelzl@54258: next hoelzl@54258: assume "bdd_below A \ bdd_below B" hoelzl@54258: then obtain a b where "\x\A. a \ x" "\x\B. b \ x" unfolding bdd_below_def by auto hoelzl@54258: hence "\x \ A \ B. inf a b \ x" by (auto intro: Un_iff le_infI1 le_infI2) hoelzl@54258: thus "bdd_below (A \ B)" unfolding bdd_below_def .. hoelzl@54258: qed hoelzl@54258: hoelzl@54259: lemma bdd_above_sup[simp]: "bdd_above ((\x. sup (f x) (g x)) ` A) \ bdd_above (f`A) \ bdd_above (g`A)" hoelzl@54259: by (auto simp: bdd_above_def intro: le_supI1 le_supI2) hoelzl@54259: hoelzl@54259: lemma bdd_below_inf[simp]: "bdd_below ((\x. inf (f x) (g x)) ` A) \ bdd_below (f`A) \ bdd_below (g`A)" hoelzl@54259: by (auto simp: bdd_below_def intro: le_infI1 le_infI2) hoelzl@54259: hoelzl@54258: end hoelzl@54258: hoelzl@54258: hoelzl@51475: text {* hoelzl@51475: hoelzl@51475: To avoid name classes with the @{class complete_lattice}-class we prefix @{const Sup} and hoelzl@51475: @{const Inf} in theorem names with c. hoelzl@51475: hoelzl@51475: *} hoelzl@51475: hoelzl@51773: class conditionally_complete_lattice = lattice + Sup + Inf + hoelzl@54258: assumes cInf_lower: "x \ X \ bdd_below X \ Inf X \ x" hoelzl@51475: and cInf_greatest: "X \ {} \ (\x. x \ X \ z \ x) \ z \ Inf X" hoelzl@54258: assumes cSup_upper: "x \ X \ bdd_above X \ x \ Sup X" hoelzl@51475: and cSup_least: "X \ {} \ (\x. x \ X \ x \ z) \ Sup X \ z" paulson@33269: begin hoelzl@51475: hoelzl@54259: lemma cSup_upper2: "x \ X \ y \ x \ bdd_above X \ y \ Sup X" hoelzl@54259: by (metis cSup_upper order_trans) hoelzl@54259: hoelzl@54259: lemma cInf_lower2: "x \ X \ x \ y \ bdd_below X \ Inf X \ y" hoelzl@54259: by (metis cInf_lower order_trans) hoelzl@54259: hoelzl@54259: lemma cSup_mono: "B \ {} \ bdd_above A \ (\b. b \ B \ \a\A. b \ a) \ Sup B \ Sup A" hoelzl@54259: by (metis cSup_least cSup_upper2) hoelzl@54259: hoelzl@54259: lemma cInf_mono: "B \ {} \ bdd_below A \ (\b. b \ B \ \a\A. a \ b) \ Inf A \ Inf B" hoelzl@54259: by (metis cInf_greatest cInf_lower2) hoelzl@54259: hoelzl@54259: lemma cSup_subset_mono: "A \ {} \ bdd_above B \ A \ B \ Sup A \ Sup B" hoelzl@54259: by (metis cSup_least cSup_upper subsetD) hoelzl@54259: hoelzl@54259: lemma cInf_superset_mono: "A \ {} \ bdd_below B \ A \ B \ Inf B \ Inf A" hoelzl@54259: by (metis cInf_greatest cInf_lower subsetD) hoelzl@54259: hoelzl@54258: lemma cSup_eq_maximum: "z \ X \ (\x. x \ X \ x \ z) \ Sup X = z" hoelzl@54258: by (intro antisym cSup_upper[of z X] cSup_least[of X z]) auto hoelzl@51475: hoelzl@54258: lemma cInf_eq_minimum: "z \ X \ (\x. x \ X \ z \ x) \ Inf X = z" hoelzl@54258: by (intro antisym cInf_lower[of z X] cInf_greatest[of X z]) auto hoelzl@51475: hoelzl@54258: lemma cSup_le_iff: "S \ {} \ bdd_above S \ Sup S \ a \ (\x\S. x \ a)" hoelzl@51475: by (metis order_trans cSup_upper cSup_least) hoelzl@51475: hoelzl@54258: lemma le_cInf_iff: "S \ {} \ bdd_below S \ a \ Inf S \ (\x\S. a \ x)" hoelzl@51475: by (metis order_trans cInf_lower cInf_greatest) hoelzl@51475: hoelzl@51475: lemma cSup_eq_non_empty: hoelzl@51475: assumes 1: "X \ {}" hoelzl@51475: assumes 2: "\x. x \ X \ x \ a" hoelzl@51475: assumes 3: "\y. (\x. x \ X \ x \ y) \ a \ y" hoelzl@51475: shows "Sup X = a" hoelzl@51475: by (intro 3 1 antisym cSup_least) (auto intro: 2 1 cSup_upper) hoelzl@51475: hoelzl@51475: lemma cInf_eq_non_empty: hoelzl@51475: assumes 1: "X \ {}" hoelzl@51475: assumes 2: "\x. x \ X \ a \ x" hoelzl@51475: assumes 3: "\y. (\x. x \ X \ y \ x) \ y \ a" hoelzl@51475: shows "Inf X = a" hoelzl@51475: by (intro 3 1 antisym cInf_greatest) (auto intro: 2 1 cInf_lower) hoelzl@51475: hoelzl@54258: lemma cInf_cSup: "S \ {} \ bdd_below S \ Inf S = Sup {x. \s\S. x \ s}" hoelzl@54258: by (rule cInf_eq_non_empty) (auto intro!: cSup_upper cSup_least simp: bdd_below_def) hoelzl@51518: hoelzl@54258: lemma cSup_cInf: "S \ {} \ bdd_above S \ Sup S = Inf {x. \s\S. s \ x}" hoelzl@54258: by (rule cSup_eq_non_empty) (auto intro!: cInf_lower cInf_greatest simp: bdd_above_def) hoelzl@54258: hoelzl@54258: lemma cSup_insert: "X \ {} \ bdd_above X \ Sup (insert a X) = sup a (Sup X)" hoelzl@54258: by (intro cSup_eq_non_empty) (auto intro: le_supI2 cSup_upper cSup_least) paulson@33269: hoelzl@54258: lemma cInf_insert: "X \ {} \ bdd_below X \ Inf (insert a X) = inf a (Inf X)" hoelzl@54258: by (intro cInf_eq_non_empty) (auto intro: le_infI2 cInf_lower cInf_greatest) hoelzl@51475: hoelzl@54259: lemma cSup_singleton [simp]: "Sup {x} = x" hoelzl@54259: by (intro cSup_eq_maximum) auto hoelzl@54259: hoelzl@54259: lemma cInf_singleton [simp]: "Inf {x} = x" hoelzl@54259: by (intro cInf_eq_minimum) auto hoelzl@54259: hoelzl@54258: lemma cSup_insert_If: "bdd_above X \ Sup (insert a X) = (if X = {} then a else sup a (Sup X))" hoelzl@54258: using cSup_insert[of X] by simp hoelzl@51475: hoelzl@54259: lemma cInf_insert_If: "bdd_below X \ Inf (insert a X) = (if X = {} then a else inf a (Inf X))" hoelzl@54258: using cInf_insert[of X] by simp hoelzl@51475: hoelzl@51475: lemma le_cSup_finite: "finite X \ x \ X \ x \ Sup X" hoelzl@51475: proof (induct X arbitrary: x rule: finite_induct) hoelzl@51475: case (insert x X y) then show ?case hoelzl@54258: by (cases "X = {}") (auto simp: cSup_insert intro: le_supI2) hoelzl@51475: qed simp hoelzl@51475: hoelzl@51475: lemma cInf_le_finite: "finite X \ x \ X \ Inf X \ x" hoelzl@51475: proof (induct X arbitrary: x rule: finite_induct) hoelzl@51475: case (insert x X y) then show ?case hoelzl@54258: by (cases "X = {}") (auto simp: cInf_insert intro: le_infI2) hoelzl@51475: qed simp hoelzl@51475: hoelzl@51475: lemma cSup_eq_Sup_fin: "finite X \ X \ {} \ Sup X = Sup_fin X" hoelzl@54258: by (induct X rule: finite_ne_induct) (simp_all add: cSup_insert) hoelzl@51475: hoelzl@51475: lemma cInf_eq_Inf_fin: "finite X \ X \ {} \ Inf X = Inf_fin X" hoelzl@54258: by (induct X rule: finite_ne_induct) (simp_all add: cInf_insert) hoelzl@51475: hoelzl@51475: lemma cSup_atMost[simp]: "Sup {..x} = x" hoelzl@51475: by (auto intro!: cSup_eq_maximum) hoelzl@51475: hoelzl@51475: lemma cSup_greaterThanAtMost[simp]: "y < x \ Sup {y<..x} = x" hoelzl@51475: by (auto intro!: cSup_eq_maximum) hoelzl@51475: hoelzl@51475: lemma cSup_atLeastAtMost[simp]: "y \ x \ Sup {y..x} = x" hoelzl@51475: by (auto intro!: cSup_eq_maximum) hoelzl@51475: hoelzl@51475: lemma cInf_atLeast[simp]: "Inf {x..} = x" hoelzl@51475: by (auto intro!: cInf_eq_minimum) hoelzl@51475: hoelzl@51475: lemma cInf_atLeastLessThan[simp]: "y < x \ Inf {y.. x \ Inf {y..x} = y" hoelzl@51475: by (auto intro!: cInf_eq_minimum) hoelzl@51475: haftmann@56218: lemma cINF_lower: "bdd_below (f ` A) \ x \ A \ INFIMUM A f \ f x" haftmann@56166: using cInf_lower [of _ "f ` A"] by simp hoelzl@54259: haftmann@56218: lemma cINF_greatest: "A \ {} \ (\x. x \ A \ m \ f x) \ m \ INFIMUM A f" haftmann@56166: using cInf_greatest [of "f ` A"] by auto hoelzl@54259: haftmann@56218: lemma cSUP_upper: "x \ A \ bdd_above (f ` A) \ f x \ SUPREMUM A f" haftmann@56166: using cSup_upper [of _ "f ` A"] by simp hoelzl@54259: haftmann@56218: lemma cSUP_least: "A \ {} \ (\x. x \ A \ f x \ M) \ SUPREMUM A f \ M" haftmann@56166: using cSup_least [of "f ` A"] by auto hoelzl@54259: haftmann@56218: lemma cINF_lower2: "bdd_below (f ` A) \ x \ A \ f x \ u \ INFIMUM A f \ u" hoelzl@54259: by (auto intro: cINF_lower assms order_trans) hoelzl@54259: haftmann@56218: lemma cSUP_upper2: "bdd_above (f ` A) \ x \ A \ u \ f x \ u \ SUPREMUM A f" hoelzl@54259: by (auto intro: cSUP_upper assms order_trans) hoelzl@54259: hoelzl@54261: lemma cSUP_const: "A \ {} \ (SUP x:A. c) = c" hoelzl@54261: by (intro antisym cSUP_least) (auto intro: cSUP_upper) hoelzl@54261: hoelzl@54261: lemma cINF_const: "A \ {} \ (INF x:A. c) = c" hoelzl@54261: by (intro antisym cINF_greatest) (auto intro: cINF_lower) hoelzl@54261: haftmann@56218: lemma le_cINF_iff: "A \ {} \ bdd_below (f ` A) \ u \ INFIMUM A f \ (\x\A. u \ f x)" hoelzl@54259: by (metis cINF_greatest cINF_lower assms order_trans) hoelzl@54259: haftmann@56218: lemma cSUP_le_iff: "A \ {} \ bdd_above (f ` A) \ SUPREMUM A f \ u \ (\x\A. f x \ u)" hoelzl@54259: by (metis cSUP_least cSUP_upper assms order_trans) hoelzl@54259: hoelzl@54263: lemma less_cINF_D: "bdd_below (f`A) \ y < (INF i:A. f i) \ i \ A \ y < f i" hoelzl@54263: by (metis cINF_lower less_le_trans) hoelzl@54263: hoelzl@54263: lemma cSUP_lessD: "bdd_above (f`A) \ (SUP i:A. f i) < y \ i \ A \ f i < y" hoelzl@54263: by (metis cSUP_upper le_less_trans) hoelzl@54263: haftmann@56218: lemma cINF_insert: "A \ {} \ bdd_below (f ` A) \ INFIMUM (insert a A) f = inf (f a) (INFIMUM A f)" haftmann@56166: by (metis cInf_insert Inf_image_eq image_insert image_is_empty) hoelzl@54259: haftmann@56218: lemma cSUP_insert: "A \ {} \ bdd_above (f ` A) \ SUPREMUM (insert a A) f = sup (f a) (SUPREMUM A f)" haftmann@56166: by (metis cSup_insert Sup_image_eq image_insert image_is_empty) hoelzl@54259: haftmann@56218: lemma cINF_mono: "B \ {} \ bdd_below (f ` A) \ (\m. m \ B \ \n\A. f n \ g m) \ INFIMUM A f \ INFIMUM B g" haftmann@56166: using cInf_mono [of "g ` B" "f ` A"] by auto hoelzl@54259: haftmann@56218: lemma cSUP_mono: "A \ {} \ bdd_above (g ` B) \ (\n. n \ A \ \m\B. f n \ g m) \ SUPREMUM A f \ SUPREMUM B g" haftmann@56166: using cSup_mono [of "f ` A" "g ` B"] by auto hoelzl@54259: haftmann@56218: lemma cINF_superset_mono: "A \ {} \ bdd_below (g ` B) \ A \ B \ (\x. x \ B \ g x \ f x) \ INFIMUM B g \ INFIMUM A f" hoelzl@54259: by (rule cINF_mono) auto hoelzl@54259: haftmann@56218: lemma cSUP_subset_mono: "A \ {} \ bdd_above (g ` B) \ A \ B \ (\x. x \ B \ f x \ g x) \ SUPREMUM A f \ SUPREMUM B g" hoelzl@54259: by (rule cSUP_mono) auto hoelzl@54259: hoelzl@54259: lemma less_eq_cInf_inter: "bdd_below A \ bdd_below B \ A \ B \ {} \ inf (Inf A) (Inf B) \ Inf (A \ B)" hoelzl@54259: by (metis cInf_superset_mono lattice_class.inf_sup_ord(1) le_infI1) hoelzl@54259: hoelzl@54259: lemma cSup_inter_less_eq: "bdd_above A \ bdd_above B \ A \ B \ {} \ Sup (A \ B) \ sup (Sup A) (Sup B) " hoelzl@54259: by (metis cSup_subset_mono lattice_class.inf_sup_ord(1) le_supI1) hoelzl@54259: hoelzl@54259: lemma cInf_union_distrib: "A \ {} \ bdd_below A \ B \ {} \ bdd_below B \ Inf (A \ B) = inf (Inf A) (Inf B)" hoelzl@54259: by (intro antisym le_infI cInf_greatest cInf_lower) (auto intro: le_infI1 le_infI2 cInf_lower) hoelzl@54259: haftmann@56218: lemma cINF_union: "A \ {} \ bdd_below (f`A) \ B \ {} \ bdd_below (f`B) \ INFIMUM (A \ B) f = inf (INFIMUM A f) (INFIMUM B f)" haftmann@56166: using cInf_union_distrib [of "f ` A" "f ` B"] by (simp add: image_Un [symmetric]) hoelzl@54259: hoelzl@54259: lemma cSup_union_distrib: "A \ {} \ bdd_above A \ B \ {} \ bdd_above B \ Sup (A \ B) = sup (Sup A) (Sup B)" hoelzl@54259: by (intro antisym le_supI cSup_least cSup_upper) (auto intro: le_supI1 le_supI2 cSup_upper) hoelzl@54259: haftmann@56218: lemma cSUP_union: "A \ {} \ bdd_above (f`A) \ B \ {} \ bdd_above (f`B) \ SUPREMUM (A \ B) f = sup (SUPREMUM A f) (SUPREMUM B f)" haftmann@56166: using cSup_union_distrib [of "f ` A" "f ` B"] by (simp add: image_Un [symmetric]) hoelzl@54259: haftmann@56218: lemma cINF_inf_distrib: "A \ {} \ bdd_below (f`A) \ bdd_below (g`A) \ inf (INFIMUM A f) (INFIMUM A g) = (INF a:A. inf (f a) (g a))" hoelzl@54259: by (intro antisym le_infI cINF_greatest cINF_lower2) hoelzl@54259: (auto intro: le_infI1 le_infI2 cINF_greatest cINF_lower le_infI) hoelzl@54259: haftmann@56218: lemma SUP_sup_distrib: "A \ {} \ bdd_above (f`A) \ bdd_above (g`A) \ sup (SUPREMUM A f) (SUPREMUM A g) = (SUP a:A. sup (f a) (g a))" hoelzl@54259: by (intro antisym le_supI cSUP_least cSUP_upper2) hoelzl@54259: (auto intro: le_supI1 le_supI2 cSUP_least cSUP_upper le_supI) hoelzl@54259: hoelzl@57447: lemma cInf_le_cSup: hoelzl@57447: "A \ {} \ bdd_above A \ bdd_below A \ Inf A \ Sup A" hoelzl@57447: by (auto intro!: cSup_upper2[of "SOME a. a \ A"] intro: someI cInf_lower) hoelzl@57447: paulson@33269: end paulson@33269: hoelzl@51773: instance complete_lattice \ conditionally_complete_lattice hoelzl@51475: by default (auto intro: Sup_upper Sup_least Inf_lower Inf_greatest) hoelzl@51475: hoelzl@51475: lemma cSup_eq: hoelzl@51773: fixes a :: "'a :: {conditionally_complete_lattice, no_bot}" hoelzl@51475: assumes upper: "\x. x \ X \ x \ a" hoelzl@51475: assumes least: "\y. (\x. x \ X \ x \ y) \ a \ y" hoelzl@51475: shows "Sup X = a" hoelzl@51475: proof cases hoelzl@51475: assume "X = {}" with lt_ex[of a] least show ?thesis by (auto simp: less_le_not_le) hoelzl@51475: qed (intro cSup_eq_non_empty assms) hoelzl@51475: hoelzl@51475: lemma cInf_eq: hoelzl@51773: fixes a :: "'a :: {conditionally_complete_lattice, no_top}" hoelzl@51475: assumes upper: "\x. x \ X \ a \ x" hoelzl@51475: assumes least: "\y. (\x. x \ X \ y \ x) \ y \ a" hoelzl@51475: shows "Inf X = a" hoelzl@51475: proof cases hoelzl@51475: assume "X = {}" with gt_ex[of a] least show ?thesis by (auto simp: less_le_not_le) hoelzl@51475: qed (intro cInf_eq_non_empty assms) hoelzl@51475: hoelzl@51773: class conditionally_complete_linorder = conditionally_complete_lattice + linorder paulson@33269: begin hoelzl@51475: hoelzl@51475: lemma less_cSup_iff : (*REAL_SUP_LE in HOL4*) hoelzl@54258: "X \ {} \ bdd_above X \ y < Sup X \ (\x\X. y < x)" hoelzl@51475: by (rule iffI) (metis cSup_least not_less, metis cSup_upper less_le_trans) hoelzl@51475: hoelzl@54258: lemma cInf_less_iff: "X \ {} \ bdd_below X \ Inf X < y \ (\x\X. x < y)" hoelzl@51475: by (rule iffI) (metis cInf_greatest not_less, metis cInf_lower le_less_trans) hoelzl@51475: hoelzl@54263: lemma cINF_less_iff: "A \ {} \ bdd_below (f`A) \ (INF i:A. f i) < a \ (\x\A. f x < a)" haftmann@56166: using cInf_less_iff[of "f`A"] by auto hoelzl@54263: hoelzl@54263: lemma less_cSUP_iff: "A \ {} \ bdd_above (f`A) \ a < (SUP i:A. f i) \ (\x\A. a < f x)" haftmann@56166: using less_cSup_iff[of "f`A"] by auto hoelzl@54263: hoelzl@51475: lemma less_cSupE: hoelzl@51475: assumes "y < Sup X" "X \ {}" obtains x where "x \ X" "y < x" hoelzl@51475: by (metis cSup_least assms not_le that) hoelzl@51475: hoelzl@51518: lemma less_cSupD: hoelzl@51518: "X \ {} \ z < Sup X \ \x\X. z < x" hoelzl@54258: by (metis less_cSup_iff not_leE bdd_above_def) hoelzl@51518: hoelzl@51518: lemma cInf_lessD: hoelzl@51518: "X \ {} \ Inf X < z \ \x\X. x < z" hoelzl@54258: by (metis cInf_less_iff not_leE bdd_below_def) hoelzl@51518: hoelzl@51475: lemma complete_interval: hoelzl@51475: assumes "a < b" and "P a" and "\ P b" hoelzl@51475: shows "\c. a \ c \ c \ b \ (\x. a \ x \ x < c \ P x) \ hoelzl@51475: (\d. (\x. a \ x \ x < d \ P x) \ d \ c)" hoelzl@51475: proof (rule exI [where x = "Sup {d. \x. a \ x & x < d --> P x}"], auto) hoelzl@51475: show "a \ Sup {d. \c. a \ c \ c < d \ P c}" hoelzl@54258: by (rule cSup_upper, auto simp: bdd_above_def) hoelzl@51475: (metis `a < b` `\ P b` linear less_le) hoelzl@51475: next hoelzl@51475: show "Sup {d. \c. a \ c \ c < d \ P c} \ b" hoelzl@51475: apply (rule cSup_least) hoelzl@51475: apply auto hoelzl@51475: apply (metis less_le_not_le) hoelzl@51475: apply (metis `a x" and lt: "x < Sup {d. \c. a \ c \ c < d \ P c}" hoelzl@51475: show "P x" hoelzl@51475: apply (rule less_cSupE [OF lt], auto) hoelzl@51475: apply (metis less_le_not_le) hoelzl@51475: apply (metis x) hoelzl@51475: done hoelzl@51475: next hoelzl@51475: fix d hoelzl@51475: assume 0: "\x. a \ x \ x < d \ P x" hoelzl@51475: thus "d \ Sup {d. \c. a \ c \ c < d \ P c}" hoelzl@54258: by (rule_tac cSup_upper, auto simp: bdd_above_def) hoelzl@51475: (metis `a X \ {} \ Sup X = Max X" hoelzl@54259: using cSup_eq_Sup_fin[of X] Sup_fin_eq_Max[of X] by simp hoelzl@51775: hoelzl@54259: lemma cInf_eq_Min: "finite (X::'a::conditionally_complete_linorder set) \ X \ {} \ Inf X = Min X" hoelzl@54259: using cInf_eq_Inf_fin[of X] Inf_fin_eq_Min[of X] by simp hoelzl@51775: hoelzl@54257: lemma cSup_lessThan[simp]: "Sup {.. Sup {y<.. Sup {y.. Inf {y<..x::'a::{conditionally_complete_linorder, dense_linorder}} = y" hoelzl@57447: by (auto intro!: cInf_eq_non_empty intro: dense_ge_bounded) hoelzl@51475: hoelzl@57447: lemma cInf_greaterThanLessThan[simp]: "y < x \ Inf {y<..a b::'a. a \ b" hoelzl@54259: begin hoelzl@54259: hoelzl@54259: lemma ex_gt_or_lt: "\b. a < b \ b < a" hoelzl@54259: by (metis UNIV_not_singleton neq_iff) hoelzl@54259: paulson@33269: end hoelzl@54259: hoelzl@54281: instantiation nat :: conditionally_complete_linorder hoelzl@54281: begin hoelzl@54281: hoelzl@54281: definition "Sup (X::nat set) = Max X" hoelzl@54281: definition "Inf (X::nat set) = (LEAST n. n \ X)" hoelzl@54281: hoelzl@54281: lemma bdd_above_nat: "bdd_above X \ finite (X::nat set)" hoelzl@54281: proof hoelzl@54281: assume "bdd_above X" hoelzl@54281: then obtain z where "X \ {.. z}" hoelzl@54281: by (auto simp: bdd_above_def) hoelzl@54281: then show "finite X" hoelzl@54281: by (rule finite_subset) simp hoelzl@54281: qed simp hoelzl@54281: hoelzl@54281: instance hoelzl@54281: proof hoelzl@54281: fix x :: nat and X :: "nat set" hoelzl@54281: { assume "x \ X" "bdd_below X" then show "Inf X \ x" hoelzl@54281: by (simp add: Inf_nat_def Least_le) } hoelzl@54281: { assume "X \ {}" "\y. y \ X \ x \ y" then show "x \ Inf X" hoelzl@54281: unfolding Inf_nat_def ex_in_conv[symmetric] by (rule LeastI2_ex) } hoelzl@54281: { assume "x \ X" "bdd_above X" then show "x \ Sup X" hoelzl@54281: by (simp add: Sup_nat_def bdd_above_nat) } hoelzl@54281: { assume "X \ {}" "\y. y \ X \ y \ x" hoelzl@54281: moreover then have "bdd_above X" hoelzl@54281: by (auto simp: bdd_above_def) hoelzl@54281: ultimately show "Sup X \ x" hoelzl@54281: by (simp add: Sup_nat_def bdd_above_nat) } hoelzl@54281: qed hoelzl@54259: end hoelzl@54281: hoelzl@54281: instantiation int :: conditionally_complete_linorder hoelzl@54281: begin hoelzl@54281: hoelzl@54281: definition "Sup (X::int set) = (THE x. x \ X \ (\y\X. y \ x))" hoelzl@54281: definition "Inf (X::int set) = - (Sup (uminus ` X))" hoelzl@54281: hoelzl@54281: instance hoelzl@54281: proof hoelzl@54281: { fix x :: int and X :: "int set" assume "X \ {}" "bdd_above X" hoelzl@54281: then obtain x y where "X \ {..y}" "x \ X" hoelzl@54281: by (auto simp: bdd_above_def) hoelzl@54281: then have *: "finite (X \ {x..y})" "X \ {x..y} \ {}" and "x \ y" hoelzl@54281: by (auto simp: subset_eq) hoelzl@54281: have "\!x\X. (\y\X. y \ x)" hoelzl@54281: proof hoelzl@54281: { fix z assume "z \ X" hoelzl@54281: have "z \ Max (X \ {x..y})" hoelzl@54281: proof cases hoelzl@54281: assume "x \ z" with `z \ X` `X \ {..y}` *(1) show ?thesis hoelzl@54281: by (auto intro!: Max_ge) hoelzl@54281: next hoelzl@54281: assume "\ x \ z" hoelzl@54281: then have "z < x" by simp hoelzl@54281: also have "x \ Max (X \ {x..y})" hoelzl@54281: using `x \ X` *(1) `x \ y` by (intro Max_ge) auto hoelzl@54281: finally show ?thesis by simp hoelzl@54281: qed } hoelzl@54281: note le = this hoelzl@54281: with Max_in[OF *] show ex: "Max (X \ {x..y}) \ X \ (\z\X. z \ Max (X \ {x..y}))" by auto hoelzl@54281: hoelzl@54281: fix z assume *: "z \ X \ (\y\X. y \ z)" hoelzl@54281: with le have "z \ Max (X \ {x..y})" hoelzl@54281: by auto hoelzl@54281: moreover have "Max (X \ {x..y}) \ z" hoelzl@54281: using * ex by auto hoelzl@54281: ultimately show "z = Max (X \ {x..y})" hoelzl@54281: by auto hoelzl@54281: qed hoelzl@54281: then have "Sup X \ X \ (\y\X. y \ Sup X)" hoelzl@54281: unfolding Sup_int_def by (rule theI') } hoelzl@54281: note Sup_int = this hoelzl@54281: hoelzl@54281: { fix x :: int and X :: "int set" assume "x \ X" "bdd_above X" then show "x \ Sup X" hoelzl@54281: using Sup_int[of X] by auto } hoelzl@54281: note le_Sup = this hoelzl@54281: { fix x :: int and X :: "int set" assume "X \ {}" "\y. y \ X \ y \ x" then show "Sup X \ x" hoelzl@54281: using Sup_int[of X] by (auto simp: bdd_above_def) } hoelzl@54281: note Sup_le = this hoelzl@54281: hoelzl@54281: { fix x :: int and X :: "int set" assume "x \ X" "bdd_below X" then show "Inf X \ x" hoelzl@54281: using le_Sup[of "-x" "uminus ` X"] by (auto simp: Inf_int_def) } hoelzl@54281: { fix x :: int and X :: "int set" assume "X \ {}" "\y. y \ X \ x \ y" then show "x \ Inf X" hoelzl@54281: using Sup_le[of "uminus ` X" "-x"] by (force simp: Inf_int_def) } hoelzl@54281: qed hoelzl@54281: end hoelzl@54281: hoelzl@57275: lemma interval_cases: hoelzl@57275: fixes S :: "'a :: conditionally_complete_linorder set" hoelzl@57275: assumes ivl: "\a b x. a \ S \ b \ S \ a \ x \ x \ b \ x \ S" hoelzl@57275: shows "\a b. S = {} \ hoelzl@57275: S = UNIV \ hoelzl@57275: S = {.. hoelzl@57275: S = {..b} \ hoelzl@57275: S = {a<..} \ hoelzl@57275: S = {a..} \ hoelzl@57275: S = {a<.. hoelzl@57275: S = {a<..b} \ hoelzl@57275: S = {a.. hoelzl@57275: S = {a..b}" hoelzl@57275: proof - hoelzl@57275: def lower \ "{x. \s\S. s \ x}" and upper \ "{x. \s\S. x \ s}" hoelzl@57275: with ivl have "S = lower \ upper" hoelzl@57275: by auto hoelzl@57275: moreover hoelzl@57275: have "\a. upper = UNIV \ upper = {} \ upper = {.. a} \ upper = {..< a}" hoelzl@57275: proof cases hoelzl@57275: assume *: "bdd_above S \ S \ {}" hoelzl@57275: from * have "upper \ {.. Sup S}" hoelzl@57275: by (auto simp: upper_def intro: cSup_upper2) hoelzl@57275: moreover from * have "{..< Sup S} \ upper" hoelzl@57275: by (force simp add: less_cSup_iff upper_def subset_eq Ball_def) hoelzl@57275: ultimately have "upper = {.. Sup S} \ upper = {..< Sup S}" hoelzl@57275: unfolding ivl_disj_un(2)[symmetric] by auto hoelzl@57275: then show ?thesis by auto hoelzl@57275: next hoelzl@57275: assume "\ (bdd_above S \ S \ {})" hoelzl@57275: then have "upper = UNIV \ upper = {}" hoelzl@57275: by (auto simp: upper_def bdd_above_def not_le dest: less_imp_le) hoelzl@57275: then show ?thesis hoelzl@57275: by auto hoelzl@57275: qed hoelzl@57275: moreover hoelzl@57275: have "\b. lower = UNIV \ lower = {} \ lower = {b ..} \ lower = {b <..}" hoelzl@57275: proof cases hoelzl@57275: assume *: "bdd_below S \ S \ {}" hoelzl@57275: from * have "lower \ {Inf S ..}" hoelzl@57275: by (auto simp: lower_def intro: cInf_lower2) hoelzl@57275: moreover from * have "{Inf S <..} \ lower" hoelzl@57275: by (force simp add: cInf_less_iff lower_def subset_eq Ball_def) hoelzl@57275: ultimately have "lower = {Inf S ..} \ lower = {Inf S <..}" hoelzl@57275: unfolding ivl_disj_un(1)[symmetric] by auto hoelzl@57275: then show ?thesis by auto hoelzl@57275: next hoelzl@57275: assume "\ (bdd_below S \ S \ {})" hoelzl@57275: then have "lower = UNIV \ lower = {}" hoelzl@57275: by (auto simp: lower_def bdd_below_def not_le dest: less_imp_le) hoelzl@57275: then show ?thesis hoelzl@57275: by auto hoelzl@57275: qed hoelzl@57275: ultimately show ?thesis hoelzl@57275: unfolding greaterThanAtMost_def greaterThanLessThan_def atLeastAtMost_def atLeastLessThan_def hoelzl@57275: by (elim exE disjE) auto hoelzl@57275: qed hoelzl@57275: hoelzl@54281: end