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
7section \<open>Conditionally-complete Lattices\<close>
8
9theory Conditionally_Complete_Lattices
10imports Finite_Set Lattices_Big Set_Interval
11begin
12
13context preorder
14begin
15
16definition "bdd_above A \<longleftrightarrow> (\<exists>M. \<forall>x \<in> A. x \<le> M)"
17definition "bdd_below A \<longleftrightarrow> (\<exists>m. \<forall>x \<in> A. m \<le> x)"
18
19lemma bdd_aboveI[intro]: "(\<And>x. x \<in> A \<Longrightarrow> x \<le> M) \<Longrightarrow> bdd_above A"
20  by (auto simp: bdd_above_def)
21
22lemma bdd_belowI[intro]: "(\<And>x. x \<in> A \<Longrightarrow> m \<le> x) \<Longrightarrow> bdd_below A"
23  by (auto simp: bdd_below_def)
24
25lemma bdd_aboveI2: "(\<And>x. x \<in> A \<Longrightarrow> f x \<le> M) \<Longrightarrow> bdd_above (f`A)"
26  by force
27
28lemma bdd_belowI2: "(\<And>x. x \<in> A \<Longrightarrow> m \<le> f x) \<Longrightarrow> bdd_below (f`A)"
29  by force
30
31lemma bdd_above_empty [simp, intro]: "bdd_above {}"
32  unfolding bdd_above_def by auto
33
34lemma bdd_below_empty [simp, intro]: "bdd_below {}"
35  unfolding bdd_below_def by auto
36
37lemma 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
40lemma 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
43lemma bdd_above_Int1 [simp]: "bdd_above A \<Longrightarrow> bdd_above (A \<inter> B)"
44  using bdd_above_mono by auto
45
46lemma bdd_above_Int2 [simp]: "bdd_above B \<Longrightarrow> bdd_above (A \<inter> B)"
47  using bdd_above_mono by auto
48
49lemma bdd_below_Int1 [simp]: "bdd_below A \<Longrightarrow> bdd_below (A \<inter> B)"
50  using bdd_below_mono by auto
51
52lemma bdd_below_Int2 [simp]: "bdd_below B \<Longrightarrow> bdd_below (A \<inter> B)"
53  using bdd_below_mono by auto
54
55lemma 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
58lemma 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
61lemma bdd_above_Iio [simp, intro]: "bdd_above {..< b}"
62  by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le)
63
64lemma 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
67lemma 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
70lemma bdd_above_Iic [simp, intro]: "bdd_above {.. b}"
71  by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le)
72
73lemma 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
76lemma 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
79lemma bdd_below_Ioi [simp, intro]: "bdd_below {a <..}"
80  by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le)
81
82lemma 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
85lemma 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
88lemma bdd_below_Ici [simp, intro]: "bdd_below {a ..}"
89  by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le)
90
91end
92
93lemma (in order_top) bdd_above_top[simp, intro!]: "bdd_above A"
94  by (rule bdd_aboveI[of _ top]) simp
95
96lemma (in order_bot) bdd_above_bot[simp, intro!]: "bdd_below A"
97  by (rule bdd_belowI[of _ bot]) simp
98
99lemma bdd_above_image_mono: "mono f \<Longrightarrow> bdd_above A \<Longrightarrow> bdd_above (f`A)"
100  by (auto simp: bdd_above_def mono_def)
101
102lemma bdd_below_image_mono: "mono f \<Longrightarrow> bdd_below A \<Longrightarrow> bdd_below (f`A)"
103  by (auto simp: bdd_below_def mono_def)
104
105lemma bdd_above_image_antimono: "antimono f \<Longrightarrow> bdd_below A \<Longrightarrow> bdd_above (f`A)"
106  by (auto simp: bdd_above_def bdd_below_def antimono_def)
107
108lemma bdd_below_image_antimono: "antimono f \<Longrightarrow> bdd_above A \<Longrightarrow> bdd_below (f`A)"
109  by (auto simp: bdd_above_def bdd_below_def antimono_def)
110
111lemma
112  fixes X :: "'a::ordered_ab_group_add set"
113  shows bdd_above_uminus[simp]: "bdd_above (uminus ` X) \<longleftrightarrow> bdd_below X"
114    and bdd_below_uminus[simp]: "bdd_below (uminus ` X) \<longleftrightarrow> bdd_above X"
115  using bdd_above_image_antimono[of uminus X] bdd_below_image_antimono[of uminus "uminus`X"]
116  using bdd_below_image_antimono[of uminus X] bdd_above_image_antimono[of uminus "uminus`X"]
117  by (auto simp: antimono_def image_image)
118
119context lattice
120begin
121
122lemma bdd_above_insert [simp]: "bdd_above (insert a A) = bdd_above A"
123  by (auto simp: bdd_above_def intro: le_supI2 sup_ge1)
124
125lemma bdd_below_insert [simp]: "bdd_below (insert a A) = bdd_below A"
126  by (auto simp: bdd_below_def intro: le_infI2 inf_le1)
127
128lemma bdd_finite [simp]:
129  assumes "finite A" shows bdd_above_finite: "bdd_above A" and bdd_below_finite: "bdd_below A"
130  using assms by (induct rule: finite_induct, auto)
131
132lemma bdd_above_Un [simp]: "bdd_above (A \<union> B) = (bdd_above A \<and> bdd_above B)"
133proof
134  assume "bdd_above (A \<union> B)"
135  thus "bdd_above A \<and> bdd_above B" unfolding bdd_above_def by auto
136next
137  assume "bdd_above A \<and> bdd_above B"
138  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
139  hence "\<forall>x \<in> A \<union> B. x \<le> sup a b" by (auto intro: Un_iff le_supI1 le_supI2)
140  thus "bdd_above (A \<union> B)" unfolding bdd_above_def ..
141qed
142
143lemma bdd_below_Un [simp]: "bdd_below (A \<union> B) = (bdd_below A \<and> bdd_below B)"
144proof
145  assume "bdd_below (A \<union> B)"
146  thus "bdd_below A \<and> bdd_below B" unfolding bdd_below_def by auto
147next
148  assume "bdd_below A \<and> bdd_below B"
149  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
150  hence "\<forall>x \<in> A \<union> B. inf a b \<le> x" by (auto intro: Un_iff le_infI1 le_infI2)
151  thus "bdd_below (A \<union> B)" unfolding bdd_below_def ..
152qed
153
154lemma bdd_above_image_sup[simp]:
155  "bdd_above ((\<lambda>x. sup (f x) (g x)) ` A) \<longleftrightarrow> bdd_above (f`A) \<and> bdd_above (g`A)"
156by (auto simp: bdd_above_def intro: le_supI1 le_supI2)
157
158lemma bdd_below_image_inf[simp]:
159  "bdd_below ((\<lambda>x. inf (f x) (g x)) ` A) \<longleftrightarrow> bdd_below (f`A) \<and> bdd_below (g`A)"
160by (auto simp: bdd_below_def intro: le_infI1 le_infI2)
161
162lemma bdd_below_UN[simp]: "finite I \<Longrightarrow> bdd_below (\<Union>i\<in>I. A i) = (\<forall>i \<in> I. bdd_below (A i))"
163by (induction I rule: finite.induct) auto
164
165lemma bdd_above_UN[simp]: "finite I \<Longrightarrow> bdd_above (\<Union>i\<in>I. A i) = (\<forall>i \<in> I. bdd_above (A i))"
166by (induction I rule: finite.induct) auto
167
168end
169
170
171text \<open>
172
173To avoid name classes with the \<^class>\<open>complete_lattice\<close>-class we prefix \<^const>\<open>Sup\<close> and
174\<^const>\<open>Inf\<close> in theorem names with c.
175
176\<close>
177
178class conditionally_complete_lattice = lattice + Sup + Inf +
179  assumes cInf_lower: "x \<in> X \<Longrightarrow> bdd_below X \<Longrightarrow> Inf X \<le> x"
180    and cInf_greatest: "X \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> z \<le> Inf X"
181  assumes cSup_upper: "x \<in> X \<Longrightarrow> bdd_above X \<Longrightarrow> x \<le> Sup X"
182    and cSup_least: "X \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup X \<le> z"
183begin
184
185lemma cSup_upper2: "x \<in> X \<Longrightarrow> y \<le> x \<Longrightarrow> bdd_above X \<Longrightarrow> y \<le> Sup X"
186  by (metis cSup_upper order_trans)
187
188lemma cInf_lower2: "x \<in> X \<Longrightarrow> x \<le> y \<Longrightarrow> bdd_below X \<Longrightarrow> Inf X \<le> y"
189  by (metis cInf_lower order_trans)
190
191lemma 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"
192  by (metis cSup_least cSup_upper2)
193
194lemma 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"
195  by (metis cInf_greatest cInf_lower2)
196
197lemma cSup_subset_mono: "A \<noteq> {} \<Longrightarrow> bdd_above B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> Sup A \<le> Sup B"
198  by (metis cSup_least cSup_upper subsetD)
199
200lemma cInf_superset_mono: "A \<noteq> {} \<Longrightarrow> bdd_below B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> Inf B \<le> Inf A"
201  by (metis cInf_greatest cInf_lower subsetD)
202
203lemma cSup_eq_maximum: "z \<in> X \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup X = z"
204  by (intro antisym cSup_upper[of z X] cSup_least[of X z]) auto
205
206lemma cInf_eq_minimum: "z \<in> X \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> Inf X = z"
207  by (intro antisym cInf_lower[of z X] cInf_greatest[of X z]) auto
208
209lemma cSup_le_iff: "S \<noteq> {} \<Longrightarrow> bdd_above S \<Longrightarrow> Sup S \<le> a \<longleftrightarrow> (\<forall>x\<in>S. x \<le> a)"
210  by (metis order_trans cSup_upper cSup_least)
211
212lemma le_cInf_iff: "S \<noteq> {} \<Longrightarrow> bdd_below S \<Longrightarrow> a \<le> Inf S \<longleftrightarrow> (\<forall>x\<in>S. a \<le> x)"
213  by (metis order_trans cInf_lower cInf_greatest)
214
215lemma cSup_eq_non_empty:
216  assumes 1: "X \<noteq> {}"
217  assumes 2: "\<And>x. x \<in> X \<Longrightarrow> x \<le> a"
218  assumes 3: "\<And>y. (\<And>x. x \<in> X \<Longrightarrow> x \<le> y) \<Longrightarrow> a \<le> y"
219  shows "Sup X = a"
220  by (intro 3 1 antisym cSup_least) (auto intro: 2 1 cSup_upper)
221
222lemma cInf_eq_non_empty:
223  assumes 1: "X \<noteq> {}"
224  assumes 2: "\<And>x. x \<in> X \<Longrightarrow> a \<le> x"
225  assumes 3: "\<And>y. (\<And>x. x \<in> X \<Longrightarrow> y \<le> x) \<Longrightarrow> y \<le> a"
226  shows "Inf X = a"
227  by (intro 3 1 antisym cInf_greatest) (auto intro: 2 1 cInf_lower)
228
229lemma cInf_cSup: "S \<noteq> {} \<Longrightarrow> bdd_below S \<Longrightarrow> Inf S = Sup {x. \<forall>s\<in>S. x \<le> s}"
230  by (rule cInf_eq_non_empty) (auto intro!: cSup_upper cSup_least simp: bdd_below_def)
231
232lemma cSup_cInf: "S \<noteq> {} \<Longrightarrow> bdd_above S \<Longrightarrow> Sup S = Inf {x. \<forall>s\<in>S. s \<le> x}"
233  by (rule cSup_eq_non_empty) (auto intro!: cInf_lower cInf_greatest simp: bdd_above_def)
234
235lemma cSup_insert: "X \<noteq> {} \<Longrightarrow> bdd_above X \<Longrightarrow> Sup (insert a X) = sup a (Sup X)"
236  by (intro cSup_eq_non_empty) (auto intro: le_supI2 cSup_upper cSup_least)
237
238lemma cInf_insert: "X \<noteq> {} \<Longrightarrow> bdd_below X \<Longrightarrow> Inf (insert a X) = inf a (Inf X)"
239  by (intro cInf_eq_non_empty) (auto intro: le_infI2 cInf_lower cInf_greatest)
240
241lemma cSup_singleton [simp]: "Sup {x} = x"
242  by (intro cSup_eq_maximum) auto
243
244lemma cInf_singleton [simp]: "Inf {x} = x"
245  by (intro cInf_eq_minimum) auto
246
247lemma cSup_insert_If:  "bdd_above X \<Longrightarrow> Sup (insert a X) = (if X = {} then a else sup a (Sup X))"
248  using cSup_insert[of X] by simp
249
250lemma cInf_insert_If: "bdd_below X \<Longrightarrow> Inf (insert a X) = (if X = {} then a else inf a (Inf X))"
251  using cInf_insert[of X] by simp
252
253lemma le_cSup_finite: "finite X \<Longrightarrow> x \<in> X \<Longrightarrow> x \<le> Sup X"
254proof (induct X arbitrary: x rule: finite_induct)
255  case (insert x X y) then show ?case
256    by (cases "X = {}") (auto simp: cSup_insert intro: le_supI2)
257qed simp
258
259lemma cInf_le_finite: "finite X \<Longrightarrow> x \<in> X \<Longrightarrow> Inf X \<le> x"
260proof (induct X arbitrary: x rule: finite_induct)
261  case (insert x X y) then show ?case
262    by (cases "X = {}") (auto simp: cInf_insert intro: le_infI2)
263qed simp
264
265lemma cSup_eq_Sup_fin: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup X = Sup_fin X"
266  by (induct X rule: finite_ne_induct) (simp_all add: cSup_insert)
267
268lemma cInf_eq_Inf_fin: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf X = Inf_fin X"
269  by (induct X rule: finite_ne_induct) (simp_all add: cInf_insert)
270
271lemma cSup_atMost[simp]: "Sup {..x} = x"
272  by (auto intro!: cSup_eq_maximum)
273
274lemma cSup_greaterThanAtMost[simp]: "y < x \<Longrightarrow> Sup {y<..x} = x"
275  by (auto intro!: cSup_eq_maximum)
276
277lemma cSup_atLeastAtMost[simp]: "y \<le> x \<Longrightarrow> Sup {y..x} = x"
278  by (auto intro!: cSup_eq_maximum)
279
280lemma cInf_atLeast[simp]: "Inf {x..} = x"
281  by (auto intro!: cInf_eq_minimum)
282
283lemma cInf_atLeastLessThan[simp]: "y < x \<Longrightarrow> Inf {y..<x} = y"
284  by (auto intro!: cInf_eq_minimum)
285
286lemma cInf_atLeastAtMost[simp]: "y \<le> x \<Longrightarrow> Inf {y..x} = y"
287  by (auto intro!: cInf_eq_minimum)
288
289lemma cINF_lower: "bdd_below (f ` A) \<Longrightarrow> x \<in> A \<Longrightarrow> \<Sqinter>(f ` A) \<le> f x"
290  using cInf_lower [of _ "f ` A"] by simp
291
292lemma cINF_greatest: "A \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> m \<le> f x) \<Longrightarrow> m \<le> \<Sqinter>(f ` A)"
293  using cInf_greatest [of "f ` A"] by auto
294
295lemma cSUP_upper: "x \<in> A \<Longrightarrow> bdd_above (f ` A) \<Longrightarrow> f x \<le> \<Squnion>(f ` A)"
296  using cSup_upper [of _ "f ` A"] by simp
297
298lemma cSUP_least: "A \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x \<le> M) \<Longrightarrow> \<Squnion>(f ` A) \<le> M"
299  using cSup_least [of "f ` A"] by auto
300
301lemma cINF_lower2: "bdd_below (f ` A) \<Longrightarrow> x \<in> A \<Longrightarrow> f x \<le> u \<Longrightarrow> \<Sqinter>(f ` A) \<le> u"
302  by (auto intro: cINF_lower order_trans)
303
304lemma cSUP_upper2: "bdd_above (f ` A) \<Longrightarrow> x \<in> A \<Longrightarrow> u \<le> f x \<Longrightarrow> u \<le> \<Squnion>(f ` A)"
305  by (auto intro: cSUP_upper order_trans)
306
307lemma cSUP_const [simp]: "A \<noteq> {} \<Longrightarrow> (\<Squnion>x\<in>A. c) = c"
308  by (intro antisym cSUP_least) (auto intro: cSUP_upper)
309
310lemma cINF_const [simp]: "A \<noteq> {} \<Longrightarrow> (\<Sqinter>x\<in>A. c) = c"
311  by (intro antisym cINF_greatest) (auto intro: cINF_lower)
312
313lemma le_cINF_iff: "A \<noteq> {} \<Longrightarrow> bdd_below (f ` A) \<Longrightarrow> u \<le> \<Sqinter>(f ` A) \<longleftrightarrow> (\<forall>x\<in>A. u \<le> f x)"
314  by (metis cINF_greatest cINF_lower order_trans)
315
316lemma cSUP_le_iff: "A \<noteq> {} \<Longrightarrow> bdd_above (f ` A) \<Longrightarrow> \<Squnion>(f ` A) \<le> u \<longleftrightarrow> (\<forall>x\<in>A. f x \<le> u)"
317  by (metis cSUP_least cSUP_upper order_trans)
318
319lemma less_cINF_D: "bdd_below (f`A) \<Longrightarrow> y < (\<Sqinter>i\<in>A. f i) \<Longrightarrow> i \<in> A \<Longrightarrow> y < f i"
320  by (metis cINF_lower less_le_trans)
321
322lemma cSUP_lessD: "bdd_above (f`A) \<Longrightarrow> (\<Squnion>i\<in>A. f i) < y \<Longrightarrow> i \<in> A \<Longrightarrow> f i < y"
323  by (metis cSUP_upper le_less_trans)
324
325lemma cINF_insert: "A \<noteq> {} \<Longrightarrow> bdd_below (f ` A) \<Longrightarrow> \<Sqinter>(f ` insert a A) = inf (f a) (\<Sqinter>(f ` A))"
326  by (simp add: cInf_insert)
327
328lemma cSUP_insert: "A \<noteq> {} \<Longrightarrow> bdd_above (f ` A) \<Longrightarrow> \<Squnion>(f ` insert a A) = sup (f a) (\<Squnion>(f ` A))"
329  by (simp add: cSup_insert)
330
331lemma 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> \<Sqinter>(f ` A) \<le> \<Sqinter>(g ` B)"
332  using cInf_mono [of "g ` B" "f ` A"] by auto
333
334lemma 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> \<Squnion>(f ` A) \<le> \<Squnion>(g ` B)"
335  using cSup_mono [of "f ` A" "g ` B"] by auto
336
337lemma 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> \<Sqinter>(g ` B) \<le> \<Sqinter>(f ` A)"
338  by (rule cINF_mono) auto
339
340lemma cSUP_subset_mono: 
341  "\<lbrakk>A \<noteq> {}; bdd_above (g ` B); A \<subseteq> B; \<And>x. x \<in> A \<Longrightarrow> f x \<le> g x\<rbrakk> \<Longrightarrow> \<Squnion> (f ` A) \<le> \<Squnion> (g ` B)"
342  by (rule cSUP_mono) auto
343
344lemma 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)"
345  by (metis cInf_superset_mono lattice_class.inf_sup_ord(1) le_infI1)
346
347lemma 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) "
348  by (metis cSup_subset_mono lattice_class.inf_sup_ord(1) le_supI1)
349
350lemma 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)"
351  by (intro antisym le_infI cInf_greatest cInf_lower) (auto intro: le_infI1 le_infI2 cInf_lower)
352
353lemma cINF_union: "A \<noteq> {} \<Longrightarrow> bdd_below (f ` A) \<Longrightarrow> B \<noteq> {} \<Longrightarrow> bdd_below (f ` B) \<Longrightarrow> \<Sqinter> (f ` (A \<union> B)) = \<Sqinter> (f ` A) \<sqinter> \<Sqinter> (f ` B)"
354  using cInf_union_distrib [of "f ` A" "f ` B"] by (simp add: image_Un)
355
356lemma 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)"
357  by (intro antisym le_supI cSup_least cSup_upper) (auto intro: le_supI1 le_supI2 cSup_upper)
358
359lemma cSUP_union: "A \<noteq> {} \<Longrightarrow> bdd_above (f ` A) \<Longrightarrow> B \<noteq> {} \<Longrightarrow> bdd_above (f ` B) \<Longrightarrow> \<Squnion> (f ` (A \<union> B)) = \<Squnion> (f ` A) \<squnion> \<Squnion> (f ` B)"
360  using cSup_union_distrib [of "f ` A" "f ` B"] by (simp add: image_Un)
361
362lemma cINF_inf_distrib: "A \<noteq> {} \<Longrightarrow> bdd_below (f`A) \<Longrightarrow> bdd_below (g`A) \<Longrightarrow> \<Sqinter> (f ` A) \<sqinter> \<Sqinter> (g ` A) = (\<Sqinter>a\<in>A. inf (f a) (g a))"
363  by (intro antisym le_infI cINF_greatest cINF_lower2)
364     (auto intro: le_infI1 le_infI2 cINF_greatest cINF_lower le_infI)
365
366lemma SUP_sup_distrib: "A \<noteq> {} \<Longrightarrow> bdd_above (f`A) \<Longrightarrow> bdd_above (g`A) \<Longrightarrow> \<Squnion> (f ` A) \<squnion> \<Squnion> (g ` A) = (\<Squnion>a\<in>A. sup (f a) (g a))"
367  by (intro antisym le_supI cSUP_least cSUP_upper2)
368     (auto intro: le_supI1 le_supI2 cSUP_least cSUP_upper le_supI)
369
370lemma cInf_le_cSup:
371  "A \<noteq> {} \<Longrightarrow> bdd_above A \<Longrightarrow> bdd_below A \<Longrightarrow> Inf A \<le> Sup A"
372  by (auto intro!: cSup_upper2[of "SOME a. a \<in> A"] intro: someI cInf_lower)
373
374end
375
376instance complete_lattice \<subseteq> conditionally_complete_lattice
377  by standard (auto intro: Sup_upper Sup_least Inf_lower Inf_greatest)
378
379lemma cSup_eq:
380  fixes a :: "'a :: {conditionally_complete_lattice, no_bot}"
381  assumes upper: "\<And>x. x \<in> X \<Longrightarrow> x \<le> a"
382  assumes least: "\<And>y. (\<And>x. x \<in> X \<Longrightarrow> x \<le> y) \<Longrightarrow> a \<le> y"
383  shows "Sup X = a"
384proof cases
385  assume "X = {}" with lt_ex[of a] least show ?thesis by (auto simp: less_le_not_le)
386qed (intro cSup_eq_non_empty assms)
387
388lemma cInf_eq:
389  fixes a :: "'a :: {conditionally_complete_lattice, no_top}"
390  assumes upper: "\<And>x. x \<in> X \<Longrightarrow> a \<le> x"
391  assumes least: "\<And>y. (\<And>x. x \<in> X \<Longrightarrow> y \<le> x) \<Longrightarrow> y \<le> a"
392  shows "Inf X = a"
393proof cases
394  assume "X = {}" with gt_ex[of a] least show ?thesis by (auto simp: less_le_not_le)
395qed (intro cInf_eq_non_empty assms)
396
397class conditionally_complete_linorder = conditionally_complete_lattice + linorder
398begin
399
400lemma less_cSup_iff:
401  "X \<noteq> {} \<Longrightarrow> bdd_above X \<Longrightarrow> y < Sup X \<longleftrightarrow> (\<exists>x\<in>X. y < x)"
402  by (rule iffI) (metis cSup_least not_less, metis cSup_upper less_le_trans)
403
404lemma cInf_less_iff: "X \<noteq> {} \<Longrightarrow> bdd_below X \<Longrightarrow> Inf X < y \<longleftrightarrow> (\<exists>x\<in>X. x < y)"
405  by (rule iffI) (metis cInf_greatest not_less, metis cInf_lower le_less_trans)
406
407lemma cINF_less_iff: "A \<noteq> {} \<Longrightarrow> bdd_below (f`A) \<Longrightarrow> (\<Sqinter>i\<in>A. f i) < a \<longleftrightarrow> (\<exists>x\<in>A. f x < a)"
408  using cInf_less_iff[of "f`A"] by auto
409
410lemma less_cSUP_iff: "A \<noteq> {} \<Longrightarrow> bdd_above (f`A) \<Longrightarrow> a < (\<Squnion>i\<in>A. f i) \<longleftrightarrow> (\<exists>x\<in>A. a < f x)"
411  using less_cSup_iff[of "f`A"] by auto
412
413lemma less_cSupE:
414  assumes "y < Sup X" "X \<noteq> {}" obtains x where "x \<in> X" "y < x"
415  by (metis cSup_least assms not_le that)
416
417lemma less_cSupD:
418  "X \<noteq> {} \<Longrightarrow> z < Sup X \<Longrightarrow> \<exists>x\<in>X. z < x"
419  by (metis less_cSup_iff not_le_imp_less bdd_above_def)
420
421lemma cInf_lessD:
422  "X \<noteq> {} \<Longrightarrow> Inf X < z \<Longrightarrow> \<exists>x\<in>X. x < z"
423  by (metis cInf_less_iff not_le_imp_less bdd_below_def)
424
425lemma complete_interval:
426  assumes "a < b" and "P a" and "\<not> P b"
427  shows "\<exists>c. a \<le> c \<and> c \<le> b \<and> (\<forall>x. a \<le> x \<and> x < c \<longrightarrow> P x) \<and>
428             (\<forall>d. (\<forall>x. a \<le> x \<and> x < d \<longrightarrow> P x) \<longrightarrow> d \<le> c)"
429proof (rule exI [where x = "Sup {d. \<forall>x. a \<le> x \<and> x < d \<longrightarrow> P x}"], auto)
430  show "a \<le> Sup {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c}"
431    by (rule cSup_upper, auto simp: bdd_above_def)
432       (metis \<open>a < b\<close> \<open>\<not> P b\<close> linear less_le)
433next
434  show "Sup {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c} \<le> b"
435    apply (rule cSup_least)
436    apply auto
437    apply (metis less_le_not_le)
438    apply (metis \<open>a<b\<close> \<open>\<not> P b\<close> linear less_le)
439    done
440next
441  fix x
442  assume x: "a \<le> x" and lt: "x < Sup {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c}"
443  show "P x"
444    apply (rule less_cSupE [OF lt], auto)
445    apply (metis less_le_not_le)
446    apply (metis x)
447    done
448next
449  fix d
450    assume 0: "\<forall>x. a \<le> x \<and> x < d \<longrightarrow> P x"
451    thus "d \<le> Sup {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c}"
452      by (rule_tac cSup_upper, auto simp: bdd_above_def)
453         (metis \<open>a<b\<close> \<open>\<not> P b\<close> linear less_le)
454qed
455
456end
457
458instance complete_linorder < conditionally_complete_linorder
459  ..
460
461lemma cSup_eq_Max: "finite (X::'a::conditionally_complete_linorder set) \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup X = Max X"
462  using cSup_eq_Sup_fin[of X] by (simp add: Sup_fin_Max)
463
464lemma cInf_eq_Min: "finite (X::'a::conditionally_complete_linorder set) \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf X = Min X"
465  using cInf_eq_Inf_fin[of X] by (simp add: Inf_fin_Min)
466
467lemma cSup_lessThan[simp]: "Sup {..<x::'a::{conditionally_complete_linorder, no_bot, dense_linorder}} = x"
468  by (auto intro!: cSup_eq_non_empty intro: dense_le)
469
470lemma cSup_greaterThanLessThan[simp]: "y < x \<Longrightarrow> Sup {y<..<x::'a::{conditionally_complete_linorder, dense_linorder}} = x"
471  by (auto intro!: cSup_eq_non_empty intro: dense_le_bounded)
472
473lemma cSup_atLeastLessThan[simp]: "y < x \<Longrightarrow> Sup {y..<x::'a::{conditionally_complete_linorder, dense_linorder}} = x"
474  by (auto intro!: cSup_eq_non_empty intro: dense_le_bounded)
475
476lemma cInf_greaterThan[simp]: "Inf {x::'a::{conditionally_complete_linorder, no_top, dense_linorder} <..} = x"
477  by (auto intro!: cInf_eq_non_empty intro: dense_ge)
478
479lemma cInf_greaterThanAtMost[simp]: "y < x \<Longrightarrow> Inf {y<..x::'a::{conditionally_complete_linorder, dense_linorder}} = y"
480  by (auto intro!: cInf_eq_non_empty intro: dense_ge_bounded)
481
482lemma cInf_greaterThanLessThan[simp]: "y < x \<Longrightarrow> Inf {y<..<x::'a::{conditionally_complete_linorder, dense_linorder}} = y"
483  by (auto intro!: cInf_eq_non_empty intro: dense_ge_bounded)
484
485lemma Inf_insert_finite:
486  fixes S :: "'a::conditionally_complete_linorder set"
487  shows "finite S \<Longrightarrow> Inf (insert x S) = (if S = {} then x else min x (Inf S))"
488  by (simp add: cInf_eq_Min)
489
490lemma Sup_insert_finite:
491  fixes S :: "'a::conditionally_complete_linorder set"
492  shows "finite S \<Longrightarrow> Sup (insert x S) = (if S = {} then x else max x (Sup S))"
493  by (simp add: cSup_insert sup_max)
494
495lemma finite_imp_less_Inf:
496  fixes a :: "'a::conditionally_complete_linorder"
497  shows "\<lbrakk>finite X; x \<in> X; \<And>x. x\<in>X \<Longrightarrow> a < x\<rbrakk> \<Longrightarrow> a < Inf X"
498  by (induction X rule: finite_induct) (simp_all add: cInf_eq_Min Inf_insert_finite)
499
500lemma finite_less_Inf_iff:
501  fixes a :: "'a :: conditionally_complete_linorder"
502  shows "\<lbrakk>finite X; X \<noteq> {}\<rbrakk> \<Longrightarrow> a < Inf X \<longleftrightarrow> (\<forall>x \<in> X. a < x)"
503  by (auto simp: cInf_eq_Min)
504
505lemma finite_imp_Sup_less:
506  fixes a :: "'a::conditionally_complete_linorder"
507  shows "\<lbrakk>finite X; x \<in> X; \<And>x. x\<in>X \<Longrightarrow> a > x\<rbrakk> \<Longrightarrow> a > Sup X"
508  by (induction X rule: finite_induct) (simp_all add: cSup_eq_Max Sup_insert_finite)
509
510lemma finite_Sup_less_iff:
511  fixes a :: "'a :: conditionally_complete_linorder"
512  shows "\<lbrakk>finite X; X \<noteq> {}\<rbrakk> \<Longrightarrow> a > Sup X \<longleftrightarrow> (\<forall>x \<in> X. a > x)"
513  by (auto simp: cSup_eq_Max)
514
515class linear_continuum = conditionally_complete_linorder + dense_linorder +
516  assumes UNIV_not_singleton: "\<exists>a b::'a. a \<noteq> b"
517begin
518
519lemma ex_gt_or_lt: "\<exists>b. a < b \<or> b < a"
520  by (metis UNIV_not_singleton neq_iff)
521
522end
523
524instantiation nat :: conditionally_complete_linorder
525begin
526
527definition "Sup (X::nat set) = Max X"
528definition "Inf (X::nat set) = (LEAST n. n \<in> X)"
529
530lemma bdd_above_nat: "bdd_above X \<longleftrightarrow> finite (X::nat set)"
531proof
532  assume "bdd_above X"
533  then obtain z where "X \<subseteq> {.. z}"
534    by (auto simp: bdd_above_def)
535  then show "finite X"
536    by (rule finite_subset) simp
537qed simp
538
539instance
540proof
541  fix x :: nat
542  fix X :: "nat set"
543  show "Inf X \<le> x" if "x \<in> X" "bdd_below X"
544    using that by (simp add: Inf_nat_def Least_le)
545  show "x \<le> Inf X" if "X \<noteq> {}" "\<And>y. y \<in> X \<Longrightarrow> x \<le> y"
546    using that unfolding Inf_nat_def ex_in_conv[symmetric] by (rule LeastI2_ex)
547  show "x \<le> Sup X" if "x \<in> X" "bdd_above X"
548    using that by (simp add: Sup_nat_def bdd_above_nat)
549  show "Sup X \<le> x" if "X \<noteq> {}" "\<And>y. y \<in> X \<Longrightarrow> y \<le> x"
550  proof -
551    from that have "bdd_above X"
552      by (auto simp: bdd_above_def)
553    with that show ?thesis 
554      by (simp add: Sup_nat_def bdd_above_nat)
555  qed
556qed
557
558end
559
560lemma Inf_nat_def1:
561  fixes K::"nat set"
562  assumes "K \<noteq> {}"
563  shows "Inf K \<in> K"
564by (auto simp add: Min_def Inf_nat_def) (meson LeastI assms bot.extremum_unique subsetI)
565
566
567instantiation int :: conditionally_complete_linorder
568begin
569
570definition "Sup (X::int set) = (THE x. x \<in> X \<and> (\<forall>y\<in>X. y \<le> x))"
571definition "Inf (X::int set) = - (Sup (uminus ` X))"
572
573instance
574proof
575  { fix x :: int and X :: "int set" assume "X \<noteq> {}" "bdd_above X"
576    then obtain x y where "X \<subseteq> {..y}" "x \<in> X"
577      by (auto simp: bdd_above_def)
578    then have *: "finite (X \<inter> {x..y})" "X \<inter> {x..y} \<noteq> {}" and "x \<le> y"
579      by (auto simp: subset_eq)
580    have "\<exists>!x\<in>X. (\<forall>y\<in>X. y \<le> x)"
581    proof
582      { fix z assume "z \<in> X"
583        have "z \<le> Max (X \<inter> {x..y})"
584        proof cases
585          assume "x \<le> z" with \<open>z \<in> X\<close> \<open>X \<subseteq> {..y}\<close> *(1) show ?thesis
586            by (auto intro!: Max_ge)
587        next
588          assume "\<not> x \<le> z"
589          then have "z < x" by simp
590          also have "x \<le> Max (X \<inter> {x..y})"
591            using \<open>x \<in> X\<close> *(1) \<open>x \<le> y\<close> by (intro Max_ge) auto
592          finally show ?thesis by simp
593        qed }
594      note le = this
595      with Max_in[OF *] show ex: "Max (X \<inter> {x..y}) \<in> X \<and> (\<forall>z\<in>X. z \<le> Max (X \<inter> {x..y}))" by auto
596
597      fix z assume *: "z \<in> X \<and> (\<forall>y\<in>X. y \<le> z)"
598      with le have "z \<le> Max (X \<inter> {x..y})"
599        by auto
600      moreover have "Max (X \<inter> {x..y}) \<le> z"
601        using * ex by auto
602      ultimately show "z = Max (X \<inter> {x..y})"
603        by auto
604    qed
605    then have "Sup X \<in> X \<and> (\<forall>y\<in>X. y \<le> Sup X)"
606      unfolding Sup_int_def by (rule theI') }
607  note Sup_int = this
608
609  { fix x :: int and X :: "int set" assume "x \<in> X" "bdd_above X" then show "x \<le> Sup X"
610      using Sup_int[of X] by auto }
611  note le_Sup = this
612  { fix x :: int and X :: "int set" assume "X \<noteq> {}" "\<And>y. y \<in> X \<Longrightarrow> y \<le> x" then show "Sup X \<le> x"
613      using Sup_int[of X] by (auto simp: bdd_above_def) }
614  note Sup_le = this
615
616  { fix x :: int and X :: "int set" assume "x \<in> X" "bdd_below X" then show "Inf X \<le> x"
617      using le_Sup[of "-x" "uminus ` X"] by (auto simp: Inf_int_def) }
618  { fix x :: int and X :: "int set" assume "X \<noteq> {}" "\<And>y. y \<in> X \<Longrightarrow> x \<le> y" then show "x \<le> Inf X"
619      using Sup_le[of "uminus ` X" "-x"] by (force simp: Inf_int_def) }
620qed
621end
622
623lemma interval_cases:
624  fixes S :: "'a :: conditionally_complete_linorder set"
625  assumes ivl: "\<And>a b x. a \<in> S \<Longrightarrow> b \<in> S \<Longrightarrow> a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> x \<in> S"
626  shows "\<exists>a b. S = {} \<or>
627    S = UNIV \<or>
628    S = {..<b} \<or>
629    S = {..b} \<or>
630    S = {a<..} \<or>
631    S = {a..} \<or>
632    S = {a<..<b} \<or>
633    S = {a<..b} \<or>
634    S = {a..<b} \<or>
635    S = {a..b}"
636proof -
637  define lower upper where "lower = {x. \<exists>s\<in>S. s \<le> x}" and "upper = {x. \<exists>s\<in>S. x \<le> s}"
638  with ivl have "S = lower \<inter> upper"
639    by auto
640  moreover
641  have "\<exists>a. upper = UNIV \<or> upper = {} \<or> upper = {.. a} \<or> upper = {..< a}"
642  proof cases
643    assume *: "bdd_above S \<and> S \<noteq> {}"
644    from * have "upper \<subseteq> {.. Sup S}"
645      by (auto simp: upper_def intro: cSup_upper2)
646    moreover from * have "{..< Sup S} \<subseteq> upper"
647      by (force simp add: less_cSup_iff upper_def subset_eq Ball_def)
648    ultimately have "upper = {.. Sup S} \<or> upper = {..< Sup S}"
649      unfolding ivl_disj_un(2)[symmetric] by auto
650    then show ?thesis by auto
651  next
652    assume "\<not> (bdd_above S \<and> S \<noteq> {})"
653    then have "upper = UNIV \<or> upper = {}"
654      by (auto simp: upper_def bdd_above_def not_le dest: less_imp_le)
655    then show ?thesis
656      by auto
657  qed
658  moreover
659  have "\<exists>b. lower = UNIV \<or> lower = {} \<or> lower = {b ..} \<or> lower = {b <..}"
660  proof cases
661    assume *: "bdd_below S \<and> S \<noteq> {}"
662    from * have "lower \<subseteq> {Inf S ..}"
663      by (auto simp: lower_def intro: cInf_lower2)
664    moreover from * have "{Inf S <..} \<subseteq> lower"
665      by (force simp add: cInf_less_iff lower_def subset_eq Ball_def)
666    ultimately have "lower = {Inf S ..} \<or> lower = {Inf S <..}"
667      unfolding ivl_disj_un(1)[symmetric] by auto
668    then show ?thesis by auto
669  next
670    assume "\<not> (bdd_below S \<and> S \<noteq> {})"
671    then have "lower = UNIV \<or> lower = {}"
672      by (auto simp: lower_def bdd_below_def not_le dest: less_imp_le)
673    then show ?thesis
674      by auto
675  qed
676  ultimately show ?thesis
677    unfolding greaterThanAtMost_def greaterThanLessThan_def atLeastAtMost_def atLeastLessThan_def
678    by (metis inf_bot_left inf_bot_right inf_top.left_neutral inf_top.right_neutral)
679qed
680
681lemma cSUP_eq_cINF_D:
682  fixes f :: "_ \<Rightarrow> 'b::conditionally_complete_lattice"
683  assumes eq: "(\<Squnion>x\<in>A. f x) = (\<Sqinter>x\<in>A. f x)"
684     and bdd: "bdd_above (f ` A)" "bdd_below (f ` A)"
685     and a: "a \<in> A"
686  shows "f a = (\<Sqinter>x\<in>A. f x)"
687apply (rule antisym)
688using a bdd
689apply (auto simp: cINF_lower)
690apply (metis eq cSUP_upper)
691done
692
693lemma cSUP_UNION:
694  fixes f :: "_ \<Rightarrow> 'b::conditionally_complete_lattice"
695  assumes ne: "A \<noteq> {}" "\<And>x. x \<in> A \<Longrightarrow> B(x) \<noteq> {}"
696      and bdd_UN: "bdd_above (\<Union>x\<in>A. f ` B x)"
697  shows "(\<Squnion>z \<in> \<Union>x\<in>A. B x. f z) = (\<Squnion>x\<in>A. \<Squnion>z\<in>B x. f z)"
698proof -
699  have bdd: "\<And>x. x \<in> A \<Longrightarrow> bdd_above (f ` B x)"
700    using bdd_UN by (meson UN_upper bdd_above_mono)
701  obtain M where "\<And>x y. x \<in> A \<Longrightarrow> y \<in> B(x) \<Longrightarrow> f y \<le> M"
702    using bdd_UN by (auto simp: bdd_above_def)
703  then have bdd2: "bdd_above ((\<lambda>x. \<Squnion>z\<in>B x. f z) ` A)"
704    unfolding bdd_above_def by (force simp: bdd cSUP_le_iff ne(2))
705  have "(\<Squnion>z \<in> \<Union>x\<in>A. B x. f z) \<le> (\<Squnion>x\<in>A. \<Squnion>z\<in>B x. f z)"
706    using assms by (fastforce simp add: intro!: cSUP_least intro: cSUP_upper2 simp: bdd2 bdd)
707  moreover have "(\<Squnion>x\<in>A. \<Squnion>z\<in>B x. f z) \<le> (\<Squnion> z \<in> \<Union>x\<in>A. B x. f z)"
708    using assms by (fastforce simp add: intro!: cSUP_least intro: cSUP_upper simp: image_UN bdd_UN)
709  ultimately show ?thesis
710    by (rule order_antisym)
711qed
712
713lemma cINF_UNION:
714  fixes f :: "_ \<Rightarrow> 'b::conditionally_complete_lattice"
715  assumes ne: "A \<noteq> {}" "\<And>x. x \<in> A \<Longrightarrow> B(x) \<noteq> {}"
716      and bdd_UN: "bdd_below (\<Union>x\<in>A. f ` B x)"
717  shows "(\<Sqinter>z \<in> \<Union>x\<in>A. B x. f z) = (\<Sqinter>x\<in>A. \<Sqinter>z\<in>B x. f z)"
718proof -
719  have bdd: "\<And>x. x \<in> A \<Longrightarrow> bdd_below (f ` B x)"
720    using bdd_UN by (meson UN_upper bdd_below_mono)
721  obtain M where "\<And>x y. x \<in> A \<Longrightarrow> y \<in> B(x) \<Longrightarrow> f y \<ge> M"
722    using bdd_UN by (auto simp: bdd_below_def)
723  then have bdd2: "bdd_below ((\<lambda>x. \<Sqinter>z\<in>B x. f z) ` A)"
724    unfolding bdd_below_def by (force simp: bdd le_cINF_iff ne(2))
725  have "(\<Sqinter>z \<in> \<Union>x\<in>A. B x. f z) \<le> (\<Sqinter>x\<in>A. \<Sqinter>z\<in>B x. f z)"
726    using assms by (fastforce simp add: intro!: cINF_greatest intro: cINF_lower simp: bdd2 bdd)
727  moreover have "(\<Sqinter>x\<in>A. \<Sqinter>z\<in>B x. f z) \<le> (\<Sqinter>z \<in> \<Union>x\<in>A. B x. f z)"
728    using assms  by (fastforce simp add: intro!: cINF_greatest intro: cINF_lower2  simp: bdd bdd_UN bdd2)
729  ultimately show ?thesis
730    by (rule order_antisym)
731qed
732
733lemma cSup_abs_le:
734  fixes S :: "('a::{linordered_idom,conditionally_complete_linorder}) set"
735  shows "S \<noteq> {} \<Longrightarrow> (\<And>x. x\<in>S \<Longrightarrow> \<bar>x\<bar> \<le> a) \<Longrightarrow> \<bar>Sup S\<bar> \<le> a"
736  apply (auto simp add: abs_le_iff intro: cSup_least)
737  by (metis bdd_aboveI cSup_upper neg_le_iff_le order_trans)
738
739end
740