src/HOL/Conditionally_Complete_Lattices.thy
author hoelzl
Tue, 05 Nov 2013 09:45:00 +0100
changeset 55713 89991ef58448
parent 55711 71c701dc5bf9
child 55714 326fd7103cb4
permissions -rw-r--r--
restrict Limsup and Liminf to complete lattices
     1 (*  Title:      HOL/Conditionally_Complete_Lattices.thy
     2     Author:     Amine Chaieb and L C Paulson, University of Cambridge
     3     Author:     Johannes Hölzl, TU München
     4     Author:     Luke S. Serafin, Carnegie Mellon University
     5 *)
     6 
     7 header {* Conditionally-complete Lattices *}
     8 
     9 theory Conditionally_Complete_Lattices
    10 imports Main Lubs
    11 begin
    12 
    13 lemma (in linorder) Sup_fin_eq_Max: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup_fin X = Max X"
    14   by (induct X rule: finite_ne_induct) (simp_all add: sup_max)
    15 
    16 lemma (in linorder) Inf_fin_eq_Min: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf_fin X = Min X"
    17   by (induct X rule: finite_ne_induct) (simp_all add: inf_min)
    18 
    19 context preorder
    20 begin
    21 
    22 definition "bdd_above A \<longleftrightarrow> (\<exists>M. \<forall>x \<in> A. x \<le> M)"
    23 definition "bdd_below A \<longleftrightarrow> (\<exists>m. \<forall>x \<in> A. m \<le> x)"
    24 
    25 lemma bdd_aboveI[intro]: "(\<And>x. x \<in> A \<Longrightarrow> x \<le> M) \<Longrightarrow> bdd_above A"
    26   by (auto simp: bdd_above_def)
    27 
    28 lemma bdd_belowI[intro]: "(\<And>x. x \<in> A \<Longrightarrow> m \<le> x) \<Longrightarrow> bdd_below A"
    29   by (auto simp: bdd_below_def)
    30 
    31 lemma bdd_above_empty [simp, intro]: "bdd_above {}"
    32   unfolding bdd_above_def by auto
    33 
    34 lemma bdd_below_empty [simp, intro]: "bdd_below {}"
    35   unfolding bdd_below_def by auto
    36 
    37 lemma bdd_above_mono: "bdd_above B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> bdd_above A"
    38   by (metis (full_types) bdd_above_def order_class.le_neq_trans psubsetD)
    39 
    40 lemma bdd_below_mono: "bdd_below B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> bdd_below A"
    41   by (metis bdd_below_def order_class.le_neq_trans psubsetD)
    42 
    43 lemma bdd_above_Int1 [simp]: "bdd_above A \<Longrightarrow> bdd_above (A \<inter> B)"
    44   using bdd_above_mono by auto
    45 
    46 lemma bdd_above_Int2 [simp]: "bdd_above B \<Longrightarrow> bdd_above (A \<inter> B)"
    47   using bdd_above_mono by auto
    48 
    49 lemma bdd_below_Int1 [simp]: "bdd_below A \<Longrightarrow> bdd_below (A \<inter> B)"
    50   using bdd_below_mono by auto
    51 
    52 lemma bdd_below_Int2 [simp]: "bdd_below B \<Longrightarrow> bdd_below (A \<inter> B)"
    53   using bdd_below_mono by auto
    54 
    55 lemma bdd_above_Ioo [simp, intro]: "bdd_above {a <..< b}"
    56   by (auto simp add: bdd_above_def intro!: exI[of _ b] less_imp_le)
    57 
    58 lemma bdd_above_Ico [simp, intro]: "bdd_above {a ..< b}"
    59   by (auto simp add: bdd_above_def intro!: exI[of _ b] less_imp_le)
    60 
    61 lemma bdd_above_Iio [simp, intro]: "bdd_above {..< b}"
    62   by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le)
    63 
    64 lemma bdd_above_Ioc [simp, intro]: "bdd_above {a <.. b}"
    65   by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le)
    66 
    67 lemma bdd_above_Icc [simp, intro]: "bdd_above {a .. b}"
    68   by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le)
    69 
    70 lemma bdd_above_Iic [simp, intro]: "bdd_above {.. b}"
    71   by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le)
    72 
    73 lemma bdd_below_Ioo [simp, intro]: "bdd_below {a <..< b}"
    74   by (auto simp add: bdd_below_def intro!: exI[of _ a] less_imp_le)
    75 
    76 lemma bdd_below_Ioc [simp, intro]: "bdd_below {a <.. b}"
    77   by (auto simp add: bdd_below_def intro!: exI[of _ a] less_imp_le)
    78 
    79 lemma bdd_below_Ioi [simp, intro]: "bdd_below {a <..}"
    80   by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le)
    81 
    82 lemma bdd_below_Ico [simp, intro]: "bdd_below {a ..< b}"
    83   by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le)
    84 
    85 lemma bdd_below_Icc [simp, intro]: "bdd_below {a .. b}"
    86   by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le)
    87 
    88 lemma bdd_below_Ici [simp, intro]: "bdd_below {a ..}"
    89   by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le)
    90 
    91 end
    92 
    93 lemma (in order_top) bdd_above_top[simp, intro!]: "bdd_above A"
    94   by (rule bdd_aboveI[of _ top]) simp
    95 
    96 lemma (in order_bot) bdd_above_bot[simp, intro!]: "bdd_below A"
    97   by (rule bdd_belowI[of _ bot]) simp
    98 
    99 context lattice
   100 begin
   101 
   102 lemma bdd_above_insert [simp]: "bdd_above (insert a A) = bdd_above A"
   103   by (auto simp: bdd_above_def intro: le_supI2 sup_ge1)
   104 
   105 lemma bdd_below_insert [simp]: "bdd_below (insert a A) = bdd_below A"
   106   by (auto simp: bdd_below_def intro: le_infI2 inf_le1)
   107 
   108 lemma bdd_finite [simp]:
   109   assumes "finite A" shows bdd_above_finite: "bdd_above A" and bdd_below_finite: "bdd_below A"
   110   using assms by (induct rule: finite_induct, auto)
   111 
   112 lemma bdd_above_Un [simp]: "bdd_above (A \<union> B) = (bdd_above A \<and> bdd_above B)"
   113 proof
   114   assume "bdd_above (A \<union> B)"
   115   thus "bdd_above A \<and> bdd_above B" unfolding bdd_above_def by auto
   116 next
   117   assume "bdd_above A \<and> bdd_above B"
   118   then obtain a b where "\<forall>x\<in>A. x \<le> a" "\<forall>x\<in>B. x \<le> b" unfolding bdd_above_def by auto
   119   hence "\<forall>x \<in> A \<union> B. x \<le> sup a b" by (auto intro: Un_iff le_supI1 le_supI2)
   120   thus "bdd_above (A \<union> B)" unfolding bdd_above_def ..
   121 qed
   122 
   123 lemma bdd_below_Un [simp]: "bdd_below (A \<union> B) = (bdd_below A \<and> bdd_below B)"
   124 proof
   125   assume "bdd_below (A \<union> B)"
   126   thus "bdd_below A \<and> bdd_below B" unfolding bdd_below_def by auto
   127 next
   128   assume "bdd_below A \<and> bdd_below B"
   129   then obtain a b where "\<forall>x\<in>A. a \<le> x" "\<forall>x\<in>B. b \<le> x" unfolding bdd_below_def by auto
   130   hence "\<forall>x \<in> A \<union> B. inf a b \<le> x" by (auto intro: Un_iff le_infI1 le_infI2)
   131   thus "bdd_below (A \<union> B)" unfolding bdd_below_def ..
   132 qed
   133 
   134 lemma bdd_above_sup[simp]: "bdd_above ((\<lambda>x. sup (f x) (g x)) ` A) \<longleftrightarrow> bdd_above (f`A) \<and> bdd_above (g`A)"
   135   by (auto simp: bdd_above_def intro: le_supI1 le_supI2)
   136 
   137 lemma bdd_below_inf[simp]: "bdd_below ((\<lambda>x. inf (f x) (g x)) ` A) \<longleftrightarrow> bdd_below (f`A) \<and> bdd_below (g`A)"
   138   by (auto simp: bdd_below_def intro: le_infI1 le_infI2)
   139 
   140 end
   141 
   142 
   143 text {*
   144 
   145 To avoid name classes with the @{class complete_lattice}-class we prefix @{const Sup} and
   146 @{const Inf} in theorem names with c.
   147 
   148 *}
   149 
   150 class conditionally_complete_lattice = lattice + Sup + Inf +
   151   assumes cInf_lower: "x \<in> X \<Longrightarrow> bdd_below X \<Longrightarrow> Inf X \<le> x"
   152     and cInf_greatest: "X \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> z \<le> Inf X"
   153   assumes cSup_upper: "x \<in> X \<Longrightarrow> bdd_above X \<Longrightarrow> x \<le> Sup X"
   154     and cSup_least: "X \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup X \<le> z"
   155 begin
   156 
   157 lemma cSup_upper2: "x \<in> X \<Longrightarrow> y \<le> x \<Longrightarrow> bdd_above X \<Longrightarrow> y \<le> Sup X"
   158   by (metis cSup_upper order_trans)
   159 
   160 lemma cInf_lower2: "x \<in> X \<Longrightarrow> x \<le> y \<Longrightarrow> bdd_below X \<Longrightarrow> Inf X \<le> y"
   161   by (metis cInf_lower order_trans)
   162 
   163 lemma cSup_mono: "B \<noteq> {} \<Longrightarrow> bdd_above A \<Longrightarrow> (\<And>b. b \<in> B \<Longrightarrow> \<exists>a\<in>A. b \<le> a) \<Longrightarrow> Sup B \<le> Sup A"
   164   by (metis cSup_least cSup_upper2)
   165 
   166 lemma cInf_mono: "B \<noteq> {} \<Longrightarrow> bdd_below A \<Longrightarrow> (\<And>b. b \<in> B \<Longrightarrow> \<exists>a\<in>A. a \<le> b) \<Longrightarrow> Inf A \<le> Inf B"
   167   by (metis cInf_greatest cInf_lower2)
   168 
   169 lemma cSup_subset_mono: "A \<noteq> {} \<Longrightarrow> bdd_above B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> Sup A \<le> Sup B"
   170   by (metis cSup_least cSup_upper subsetD)
   171 
   172 lemma cInf_superset_mono: "A \<noteq> {} \<Longrightarrow> bdd_below B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> Inf B \<le> Inf A"
   173   by (metis cInf_greatest cInf_lower subsetD)
   174 
   175 lemma cSup_eq_maximum: "z \<in> X \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup X = z"
   176   by (intro antisym cSup_upper[of z X] cSup_least[of X z]) auto
   177 
   178 lemma cInf_eq_minimum: "z \<in> X \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> Inf X = z"
   179   by (intro antisym cInf_lower[of z X] cInf_greatest[of X z]) auto
   180 
   181 lemma cSup_le_iff: "S \<noteq> {} \<Longrightarrow> bdd_above S \<Longrightarrow> Sup S \<le> a \<longleftrightarrow> (\<forall>x\<in>S. x \<le> a)"
   182   by (metis order_trans cSup_upper cSup_least)
   183 
   184 lemma le_cInf_iff: "S \<noteq> {} \<Longrightarrow> bdd_below S \<Longrightarrow> a \<le> Inf S \<longleftrightarrow> (\<forall>x\<in>S. a \<le> x)"
   185   by (metis order_trans cInf_lower cInf_greatest)
   186 
   187 lemma cSup_eq_non_empty:
   188   assumes 1: "X \<noteq> {}"
   189   assumes 2: "\<And>x. x \<in> X \<Longrightarrow> x \<le> a"
   190   assumes 3: "\<And>y. (\<And>x. x \<in> X \<Longrightarrow> x \<le> y) \<Longrightarrow> a \<le> y"
   191   shows "Sup X = a"
   192   by (intro 3 1 antisym cSup_least) (auto intro: 2 1 cSup_upper)
   193 
   194 lemma cInf_eq_non_empty:
   195   assumes 1: "X \<noteq> {}"
   196   assumes 2: "\<And>x. x \<in> X \<Longrightarrow> a \<le> x"
   197   assumes 3: "\<And>y. (\<And>x. x \<in> X \<Longrightarrow> y \<le> x) \<Longrightarrow> y \<le> a"
   198   shows "Inf X = a"
   199   by (intro 3 1 antisym cInf_greatest) (auto intro: 2 1 cInf_lower)
   200 
   201 lemma cInf_cSup: "S \<noteq> {} \<Longrightarrow> bdd_below S \<Longrightarrow> Inf S = Sup {x. \<forall>s\<in>S. x \<le> s}"
   202   by (rule cInf_eq_non_empty) (auto intro!: cSup_upper cSup_least simp: bdd_below_def)
   203 
   204 lemma cSup_cInf: "S \<noteq> {} \<Longrightarrow> bdd_above S \<Longrightarrow> Sup S = Inf {x. \<forall>s\<in>S. s \<le> x}"
   205   by (rule cSup_eq_non_empty) (auto intro!: cInf_lower cInf_greatest simp: bdd_above_def)
   206 
   207 lemma cSup_insert: "X \<noteq> {} \<Longrightarrow> bdd_above X \<Longrightarrow> Sup (insert a X) = sup a (Sup X)"
   208   by (intro cSup_eq_non_empty) (auto intro: le_supI2 cSup_upper cSup_least)
   209 
   210 lemma cInf_insert: "X \<noteq> {} \<Longrightarrow> bdd_below X \<Longrightarrow> Inf (insert a X) = inf a (Inf X)"
   211   by (intro cInf_eq_non_empty) (auto intro: le_infI2 cInf_lower cInf_greatest)
   212 
   213 lemma cSup_singleton [simp]: "Sup {x} = x"
   214   by (intro cSup_eq_maximum) auto
   215 
   216 lemma cInf_singleton [simp]: "Inf {x} = x"
   217   by (intro cInf_eq_minimum) auto
   218 
   219 lemma cSup_insert_If:  "bdd_above X \<Longrightarrow> Sup (insert a X) = (if X = {} then a else sup a (Sup X))"
   220   using cSup_insert[of X] by simp
   221 
   222 lemma cInf_insert_If: "bdd_below X \<Longrightarrow> Inf (insert a X) = (if X = {} then a else inf a (Inf X))"
   223   using cInf_insert[of X] by simp
   224 
   225 lemma le_cSup_finite: "finite X \<Longrightarrow> x \<in> X \<Longrightarrow> x \<le> Sup X"
   226 proof (induct X arbitrary: x rule: finite_induct)
   227   case (insert x X y) then show ?case
   228     by (cases "X = {}") (auto simp: cSup_insert intro: le_supI2)
   229 qed simp
   230 
   231 lemma cInf_le_finite: "finite X \<Longrightarrow> x \<in> X \<Longrightarrow> Inf X \<le> x"
   232 proof (induct X arbitrary: x rule: finite_induct)
   233   case (insert x X y) then show ?case
   234     by (cases "X = {}") (auto simp: cInf_insert intro: le_infI2)
   235 qed simp
   236 
   237 lemma cSup_eq_Sup_fin: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup X = Sup_fin X"
   238   by (induct X rule: finite_ne_induct) (simp_all add: cSup_insert)
   239 
   240 lemma cInf_eq_Inf_fin: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf X = Inf_fin X"
   241   by (induct X rule: finite_ne_induct) (simp_all add: cInf_insert)
   242 
   243 lemma cSup_atMost[simp]: "Sup {..x} = x"
   244   by (auto intro!: cSup_eq_maximum)
   245 
   246 lemma cSup_greaterThanAtMost[simp]: "y < x \<Longrightarrow> Sup {y<..x} = x"
   247   by (auto intro!: cSup_eq_maximum)
   248 
   249 lemma cSup_atLeastAtMost[simp]: "y \<le> x \<Longrightarrow> Sup {y..x} = x"
   250   by (auto intro!: cSup_eq_maximum)
   251 
   252 lemma cInf_atLeast[simp]: "Inf {x..} = x"
   253   by (auto intro!: cInf_eq_minimum)
   254 
   255 lemma cInf_atLeastLessThan[simp]: "y < x \<Longrightarrow> Inf {y..<x} = y"
   256   by (auto intro!: cInf_eq_minimum)
   257 
   258 lemma cInf_atLeastAtMost[simp]: "y \<le> x \<Longrightarrow> Inf {y..x} = y"
   259   by (auto intro!: cInf_eq_minimum)
   260 
   261 lemma cINF_lower: "bdd_below (f ` A) \<Longrightarrow> x \<in> A \<Longrightarrow> INFI A f \<le> f x"
   262   unfolding INF_def by (rule cInf_lower) auto
   263 
   264 lemma cINF_greatest: "A \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> m \<le> f x) \<Longrightarrow> m \<le> INFI A f"
   265   unfolding INF_def by (rule cInf_greatest) auto
   266 
   267 lemma cSUP_upper: "x \<in> A \<Longrightarrow> bdd_above (f ` A) \<Longrightarrow> f x \<le> SUPR A f"
   268   unfolding SUP_def by (rule cSup_upper) auto
   269 
   270 lemma cSUP_least: "A \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x \<le> M) \<Longrightarrow> SUPR A f \<le> M"
   271   unfolding SUP_def by (rule cSup_least) auto
   272 
   273 lemma cINF_lower2: "bdd_below (f ` A) \<Longrightarrow> x \<in> A \<Longrightarrow> f x \<le> u \<Longrightarrow> INFI A f \<le> u"
   274   by (auto intro: cINF_lower assms order_trans)
   275 
   276 lemma cSUP_upper2: "bdd_above (f ` A) \<Longrightarrow> x \<in> A \<Longrightarrow> u \<le> f x \<Longrightarrow> u \<le> SUPR A f"
   277   by (auto intro: cSUP_upper assms order_trans)
   278 
   279 lemma cSUP_const: "A \<noteq> {} \<Longrightarrow> (SUP x:A. c) = c"
   280   by (intro antisym cSUP_least) (auto intro: cSUP_upper)
   281 
   282 lemma cINF_const: "A \<noteq> {} \<Longrightarrow> (INF x:A. c) = c"
   283   by (intro antisym cINF_greatest) (auto intro: cINF_lower)
   284 
   285 lemma le_cINF_iff: "A \<noteq> {} \<Longrightarrow> bdd_below (f ` A) \<Longrightarrow> u \<le> INFI A f \<longleftrightarrow> (\<forall>x\<in>A. u \<le> f x)"
   286   by (metis cINF_greatest cINF_lower assms order_trans)
   287 
   288 lemma cSUP_le_iff: "A \<noteq> {} \<Longrightarrow> bdd_above (f ` A) \<Longrightarrow> SUPR A f \<le> u \<longleftrightarrow> (\<forall>x\<in>A. f x \<le> u)"
   289   by (metis cSUP_least cSUP_upper assms order_trans)
   290 
   291 lemma cINF_insert: "A \<noteq> {} \<Longrightarrow> bdd_below (f ` A) \<Longrightarrow> INFI (insert a A) f = inf (f a) (INFI A f)"
   292   by (metis INF_def cInf_insert assms empty_is_image image_insert)
   293 
   294 lemma cSUP_insert: "A \<noteq> {} \<Longrightarrow> bdd_above (f ` A) \<Longrightarrow> SUPR (insert a A) f = sup (f a) (SUPR A f)"
   295   by (metis SUP_def cSup_insert assms empty_is_image image_insert)
   296 
   297 lemma cINF_mono: "B \<noteq> {} \<Longrightarrow> bdd_below (f ` A) \<Longrightarrow> (\<And>m. m \<in> B \<Longrightarrow> \<exists>n\<in>A. f n \<le> g m) \<Longrightarrow> INFI A f \<le> INFI B g"
   298   unfolding INF_def by (auto intro: cInf_mono)
   299 
   300 lemma cSUP_mono: "A \<noteq> {} \<Longrightarrow> bdd_above (g ` B) \<Longrightarrow> (\<And>n. n \<in> A \<Longrightarrow> \<exists>m\<in>B. f n \<le> g m) \<Longrightarrow> SUPR A f \<le> SUPR B g"
   301   unfolding SUP_def by (auto intro: cSup_mono)
   302 
   303 lemma cINF_superset_mono: "A \<noteq> {} \<Longrightarrow> bdd_below (g ` B) \<Longrightarrow> A \<subseteq> B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> g x \<le> f x) \<Longrightarrow> INFI B g \<le> INFI A f"
   304   by (rule cINF_mono) auto
   305 
   306 lemma cSUP_subset_mono: "A \<noteq> {} \<Longrightarrow> bdd_above (g ` B) \<Longrightarrow> A \<subseteq> B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> f x \<le> g x) \<Longrightarrow> SUPR A f \<le> SUPR B g"
   307   by (rule cSUP_mono) auto
   308 
   309 lemma less_eq_cInf_inter: "bdd_below A \<Longrightarrow> bdd_below B \<Longrightarrow> A \<inter> B \<noteq> {} \<Longrightarrow> inf (Inf A) (Inf B) \<le> Inf (A \<inter> B)"
   310   by (metis cInf_superset_mono lattice_class.inf_sup_ord(1) le_infI1)
   311 
   312 lemma cSup_inter_less_eq: "bdd_above A \<Longrightarrow> bdd_above B \<Longrightarrow> A \<inter> B \<noteq> {} \<Longrightarrow> Sup (A \<inter> B) \<le> sup (Sup A) (Sup B) "
   313   by (metis cSup_subset_mono lattice_class.inf_sup_ord(1) le_supI1)
   314 
   315 lemma cInf_union_distrib: "A \<noteq> {} \<Longrightarrow> bdd_below A \<Longrightarrow> B \<noteq> {} \<Longrightarrow> bdd_below B \<Longrightarrow> Inf (A \<union> B) = inf (Inf A) (Inf B)"
   316   by (intro antisym le_infI cInf_greatest cInf_lower) (auto intro: le_infI1 le_infI2 cInf_lower)
   317 
   318 lemma cINF_union: "A \<noteq> {} \<Longrightarrow> bdd_below (f`A) \<Longrightarrow> B \<noteq> {} \<Longrightarrow> bdd_below (f`B) \<Longrightarrow> INFI (A \<union> B) f = inf (INFI A f) (INFI B f)"
   319   unfolding INF_def using assms by (auto simp add: image_Un intro: cInf_union_distrib)
   320 
   321 lemma cSup_union_distrib: "A \<noteq> {} \<Longrightarrow> bdd_above A \<Longrightarrow> B \<noteq> {} \<Longrightarrow> bdd_above B \<Longrightarrow> Sup (A \<union> B) = sup (Sup A) (Sup B)"
   322   by (intro antisym le_supI cSup_least cSup_upper) (auto intro: le_supI1 le_supI2 cSup_upper)
   323 
   324 lemma cSUP_union: "A \<noteq> {} \<Longrightarrow> bdd_above (f`A) \<Longrightarrow> B \<noteq> {} \<Longrightarrow> bdd_above (f`B) \<Longrightarrow> SUPR (A \<union> B) f = sup (SUPR A f) (SUPR B f)"
   325   unfolding SUP_def by (auto simp add: image_Un intro: cSup_union_distrib)
   326 
   327 lemma cINF_inf_distrib: "A \<noteq> {} \<Longrightarrow> bdd_below (f`A) \<Longrightarrow> bdd_below (g`A) \<Longrightarrow> inf (INFI A f) (INFI A g) = (INF a:A. inf (f a) (g a))"
   328   by (intro antisym le_infI cINF_greatest cINF_lower2)
   329      (auto intro: le_infI1 le_infI2 cINF_greatest cINF_lower le_infI)
   330 
   331 lemma SUP_sup_distrib: "A \<noteq> {} \<Longrightarrow> bdd_above (f`A) \<Longrightarrow> bdd_above (g`A) \<Longrightarrow> sup (SUPR A f) (SUPR A g) = (SUP a:A. sup (f a) (g a))"
   332   by (intro antisym le_supI cSUP_least cSUP_upper2)
   333      (auto intro: le_supI1 le_supI2 cSUP_least cSUP_upper le_supI)
   334 
   335 end
   336 
   337 instance complete_lattice \<subseteq> conditionally_complete_lattice
   338   by default (auto intro: Sup_upper Sup_least Inf_lower Inf_greatest)
   339 
   340 lemma isLub_cSup: 
   341   "(S::'a :: conditionally_complete_lattice set) \<noteq> {} \<Longrightarrow> (\<exists>b. S *<= b) \<Longrightarrow> isLub UNIV S (Sup S)"
   342   by  (auto simp add: isLub_def setle_def leastP_def isUb_def
   343             intro!: setgeI cSup_upper cSup_least)
   344 
   345 lemma cSup_eq:
   346   fixes a :: "'a :: {conditionally_complete_lattice, no_bot}"
   347   assumes upper: "\<And>x. x \<in> X \<Longrightarrow> x \<le> a"
   348   assumes least: "\<And>y. (\<And>x. x \<in> X \<Longrightarrow> x \<le> y) \<Longrightarrow> a \<le> y"
   349   shows "Sup X = a"
   350 proof cases
   351   assume "X = {}" with lt_ex[of a] least show ?thesis by (auto simp: less_le_not_le)
   352 qed (intro cSup_eq_non_empty assms)
   353 
   354 lemma cInf_eq:
   355   fixes a :: "'a :: {conditionally_complete_lattice, no_top}"
   356   assumes upper: "\<And>x. x \<in> X \<Longrightarrow> a \<le> x"
   357   assumes least: "\<And>y. (\<And>x. x \<in> X \<Longrightarrow> y \<le> x) \<Longrightarrow> y \<le> a"
   358   shows "Inf X = a"
   359 proof cases
   360   assume "X = {}" with gt_ex[of a] least show ?thesis by (auto simp: less_le_not_le)
   361 qed (intro cInf_eq_non_empty assms)
   362 
   363 lemma cSup_le: "(S::'a::conditionally_complete_lattice set) \<noteq> {} \<Longrightarrow> S *<= b \<Longrightarrow> Sup S \<le> b"
   364   by (metis cSup_least setle_def)
   365 
   366 lemma cInf_ge: "(S::'a :: conditionally_complete_lattice set) \<noteq> {} \<Longrightarrow> b <=* S \<Longrightarrow> Inf S \<ge> b"
   367   by (metis cInf_greatest setge_def)
   368 
   369 class conditionally_complete_linorder = conditionally_complete_lattice + linorder
   370 begin
   371 
   372 lemma less_cSup_iff : (*REAL_SUP_LE in HOL4*)
   373   "X \<noteq> {} \<Longrightarrow> bdd_above X \<Longrightarrow> y < Sup X \<longleftrightarrow> (\<exists>x\<in>X. y < x)"
   374   by (rule iffI) (metis cSup_least not_less, metis cSup_upper less_le_trans)
   375 
   376 lemma cInf_less_iff: "X \<noteq> {} \<Longrightarrow> bdd_below X \<Longrightarrow> Inf X < y \<longleftrightarrow> (\<exists>x\<in>X. x < y)"
   377   by (rule iffI) (metis cInf_greatest not_less, metis cInf_lower le_less_trans)
   378 
   379 lemma less_cSupE:
   380   assumes "y < Sup X" "X \<noteq> {}" obtains x where "x \<in> X" "y < x"
   381   by (metis cSup_least assms not_le that)
   382 
   383 lemma less_cSupD:
   384   "X \<noteq> {} \<Longrightarrow> z < Sup X \<Longrightarrow> \<exists>x\<in>X. z < x"
   385   by (metis less_cSup_iff not_leE bdd_above_def)
   386 
   387 lemma cInf_lessD:
   388   "X \<noteq> {} \<Longrightarrow> Inf X < z \<Longrightarrow> \<exists>x\<in>X. x < z"
   389   by (metis cInf_less_iff not_leE bdd_below_def)
   390 
   391 lemma complete_interval:
   392   assumes "a < b" and "P a" and "\<not> P b"
   393   shows "\<exists>c. a \<le> c \<and> c \<le> b \<and> (\<forall>x. a \<le> x \<and> x < c \<longrightarrow> P x) \<and>
   394              (\<forall>d. (\<forall>x. a \<le> x \<and> x < d \<longrightarrow> P x) \<longrightarrow> d \<le> c)"
   395 proof (rule exI [where x = "Sup {d. \<forall>x. a \<le> x & x < d --> P x}"], auto)
   396   show "a \<le> Sup {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c}"
   397     by (rule cSup_upper, auto simp: bdd_above_def)
   398        (metis `a < b` `\<not> P b` linear less_le)
   399 next
   400   show "Sup {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c} \<le> b"
   401     apply (rule cSup_least) 
   402     apply auto
   403     apply (metis less_le_not_le)
   404     apply (metis `a<b` `~ P b` linear less_le)
   405     done
   406 next
   407   fix x
   408   assume x: "a \<le> x" and lt: "x < Sup {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c}"
   409   show "P x"
   410     apply (rule less_cSupE [OF lt], auto)
   411     apply (metis less_le_not_le)
   412     apply (metis x) 
   413     done
   414 next
   415   fix d
   416     assume 0: "\<forall>x. a \<le> x \<and> x < d \<longrightarrow> P x"
   417     thus "d \<le> Sup {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c}"
   418       by (rule_tac cSup_upper, auto simp: bdd_above_def)
   419          (metis `a<b` `~ P b` linear less_le)
   420 qed
   421 
   422 end
   423 
   424 lemma cSup_eq_Max: "finite (X::'a::conditionally_complete_linorder set) \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup X = Max X"
   425   using cSup_eq_Sup_fin[of X] Sup_fin_eq_Max[of X] by simp
   426 
   427 lemma cInf_eq_Min: "finite (X::'a::conditionally_complete_linorder set) \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf X = Min X"
   428   using cInf_eq_Inf_fin[of X] Inf_fin_eq_Min[of X] by simp
   429 
   430 lemma cSup_bounds:
   431   fixes S :: "'a :: conditionally_complete_lattice set"
   432   assumes Se: "S \<noteq> {}" and l: "a <=* S" and u: "S *<= b"
   433   shows "a \<le> Sup S \<and> Sup S \<le> b"
   434 proof-
   435   from isLub_cSup[OF Se] u have lub: "isLub UNIV S (Sup S)" by blast
   436   hence b: "Sup S \<le> b" using u 
   437     by (auto simp add: isLub_def leastP_def setle_def setge_def isUb_def) 
   438   from Se obtain y where y: "y \<in> S" by blast
   439   from lub l have "a \<le> Sup S"
   440     by (auto simp add: isLub_def leastP_def setle_def setge_def isUb_def)
   441        (metis le_iff_sup le_sup_iff y)
   442   with b show ?thesis by blast
   443 qed
   444 
   445 lemma cSup_unique: "(S::'a :: {conditionally_complete_linorder, no_bot} set) *<= b \<Longrightarrow> (\<forall>b'<b. \<exists>x\<in>S. b' < x) \<Longrightarrow> Sup S = b"
   446   by (rule cSup_eq) (auto simp: not_le[symmetric] setle_def)
   447 
   448 lemma cInf_unique: "b <=* (S::'a :: {conditionally_complete_linorder, no_top} set) \<Longrightarrow> (\<forall>b'>b. \<exists>x\<in>S. b' > x) \<Longrightarrow> Inf S = b"
   449   by (rule cInf_eq) (auto simp: not_le[symmetric] setge_def)
   450 
   451 lemma cSup_lessThan[simp]: "Sup {..<x::'a::{conditionally_complete_linorder, no_bot, dense_linorder}} = x"
   452   by (auto intro!: cSup_eq_non_empty intro: dense_le)
   453 
   454 lemma cSup_greaterThanLessThan[simp]: "y < x \<Longrightarrow> Sup {y<..<x::'a::{conditionally_complete_linorder, no_bot, dense_linorder}} = x"
   455   by (auto intro!: cSup_eq intro: dense_le_bounded)
   456 
   457 lemma cSup_atLeastLessThan[simp]: "y < x \<Longrightarrow> Sup {y..<x::'a::{conditionally_complete_linorder, no_bot, dense_linorder}} = x"
   458   by (auto intro!: cSup_eq intro: dense_le_bounded)
   459 
   460 lemma cInf_greaterThan[simp]: "Inf {x::'a::{conditionally_complete_linorder, no_top, dense_linorder} <..} = x"
   461   by (auto intro!: cInf_eq intro: dense_ge)
   462 
   463 lemma cInf_greaterThanAtMost[simp]: "y < x \<Longrightarrow> Inf {y<..x::'a::{conditionally_complete_linorder, no_top, dense_linorder}} = y"
   464   by (auto intro!: cInf_eq intro: dense_ge_bounded)
   465 
   466 lemma cInf_greaterThanLessThan[simp]: "y < x \<Longrightarrow> Inf {y<..<x::'a::{conditionally_complete_linorder, no_top, dense_linorder}} = y"
   467   by (auto intro!: cInf_eq intro: dense_ge_bounded)
   468 
   469 class linear_continuum = conditionally_complete_linorder + dense_linorder +
   470   assumes UNIV_not_singleton: "\<exists>a b::'a. a \<noteq> b"
   471 begin
   472 
   473 lemma ex_gt_or_lt: "\<exists>b. a < b \<or> b < a"
   474   by (metis UNIV_not_singleton neq_iff)
   475 
   476 end
   477 
   478 end