| author | desharna | 
| Thu, 19 Dec 2024 16:01:06 +0100 | |
| changeset 81635 | 362b2ff84206 | 
| parent 78698 | 1b9388e6eb75 | 
| permissions | -rw-r--r-- | 
| 52265 | 1 | (* Title: HOL/Conditionally_Complete_Lattices.thy | 
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51475diff
changeset | 2 | Author: Amine Chaieb and L C Paulson, University of Cambridge | 
| 51643 | 3 | Author: Johannes Hölzl, TU München | 
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 4 | Author: Luke S. Serafin, Carnegie Mellon University | 
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51475diff
changeset | 5 | *) | 
| 33269 
3b7e2dbbd684
New theory SupInf of the supremum and infimum operators for sets of reals.
 paulson parents: diff
changeset | 6 | |
| 60758 | 7 | section \<open>Conditionally-complete Lattices\<close> | 
| 33269 
3b7e2dbbd684
New theory SupInf of the supremum and infimum operators for sets of reals.
 paulson parents: diff
changeset | 8 | |
| 51773 | 9 | theory Conditionally_Complete_Lattices | 
| 63331 | 10 | imports Finite_Set Lattices_Big Set_Interval | 
| 33269 
3b7e2dbbd684
New theory SupInf of the supremum and infimum operators for sets of reals.
 paulson parents: diff
changeset | 11 | begin | 
| 
3b7e2dbbd684
New theory SupInf of the supremum and infimum operators for sets of reals.
 paulson parents: diff
changeset | 12 | |
| 73271 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 13 | locale preordering_bdd = preordering | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 14 | begin | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 15 | |
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 16 | definition bdd :: \<open>'a set \<Rightarrow> bool\<close> | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 17 | where unfold: \<open>bdd A \<longleftrightarrow> (\<exists>M. \<forall>x \<in> A. x \<^bold>\<le> M)\<close> | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 18 | |
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 19 | lemma empty [simp, intro]: | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 20 |   \<open>bdd {}\<close>
 | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 21 | by (simp add: unfold) | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 22 | |
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 23 | lemma I [intro]: | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 24 | \<open>bdd A\<close> if \<open>\<And>x. x \<in> A \<Longrightarrow> x \<^bold>\<le> M\<close> | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 25 | using that by (auto simp add: unfold) | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 26 | |
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 27 | lemma E: | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 28 | assumes \<open>bdd A\<close> | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 29 | obtains M where \<open>\<And>x. x \<in> A \<Longrightarrow> x \<^bold>\<le> M\<close> | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 30 | using assms that by (auto simp add: unfold) | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 31 | |
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 32 | lemma I2: | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 33 | \<open>bdd (f ` A)\<close> if \<open>\<And>x. x \<in> A \<Longrightarrow> f x \<^bold>\<le> M\<close> | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 34 | using that by (auto simp add: unfold) | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 35 | |
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 36 | lemma mono: | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 37 | \<open>bdd A\<close> if \<open>bdd B\<close> \<open>A \<subseteq> B\<close> | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 38 | using that by (auto simp add: unfold) | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 39 | |
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 40 | lemma Int1 [simp]: | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 41 | \<open>bdd (A \<inter> B)\<close> if \<open>bdd A\<close> | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 42 | using mono that by auto | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 43 | |
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 44 | lemma Int2 [simp]: | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 45 | \<open>bdd (A \<inter> B)\<close> if \<open>bdd B\<close> | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 46 | using mono that by auto | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 47 | |
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 48 | end | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 49 | |
| 75494 | 50 | subsection \<open>Preorders\<close> | 
| 51 | ||
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 52 | context preorder | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 53 | begin | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 54 | |
| 73271 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 55 | sublocale bdd_above: preordering_bdd \<open>(\<le>)\<close> \<open>(<)\<close> | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 56 | defines bdd_above_primitive_def: bdd_above = bdd_above.bdd .. | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 57 | |
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 58 | sublocale bdd_below: preordering_bdd \<open>(\<ge>)\<close> \<open>(>)\<close> | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 59 | defines bdd_below_primitive_def: bdd_below = bdd_below.bdd .. | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 60 | |
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 61 | lemma bdd_above_def: \<open>bdd_above A \<longleftrightarrow> (\<exists>M. \<forall>x \<in> A. x \<le> M)\<close> | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 62 | by (fact bdd_above.unfold) | 
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 63 | |
| 73271 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 64 | lemma bdd_below_def: \<open>bdd_below A \<longleftrightarrow> (\<exists>M. \<forall>x \<in> A. M \<le> x)\<close> | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 65 | by (fact bdd_below.unfold) | 
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 66 | |
| 73271 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 67 | lemma bdd_aboveI: "(\<And>x. x \<in> A \<Longrightarrow> x \<le> M) \<Longrightarrow> bdd_above A" | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 68 | by (fact bdd_above.I) | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 69 | |
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 70 | lemma bdd_belowI: "(\<And>x. x \<in> A \<Longrightarrow> m \<le> x) \<Longrightarrow> bdd_below A" | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 71 | by (fact bdd_below.I) | 
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 72 | |
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54262diff
changeset | 73 | lemma bdd_aboveI2: "(\<And>x. x \<in> A \<Longrightarrow> f x \<le> M) \<Longrightarrow> bdd_above (f`A)" | 
| 73271 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 74 | by (fact bdd_above.I2) | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54262diff
changeset | 75 | |
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54262diff
changeset | 76 | lemma bdd_belowI2: "(\<And>x. x \<in> A \<Longrightarrow> m \<le> f x) \<Longrightarrow> bdd_below (f`A)" | 
| 73271 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 77 | by (fact bdd_below.I2) | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54262diff
changeset | 78 | |
| 73271 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 79 | lemma bdd_above_empty: "bdd_above {}"
 | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 80 | by (fact bdd_above.empty) | 
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 81 | |
| 73271 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 82 | lemma bdd_below_empty: "bdd_below {}"
 | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 83 | by (fact bdd_below.empty) | 
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 84 | |
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 85 | lemma bdd_above_mono: "bdd_above B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> bdd_above A" | 
| 73271 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 86 | by (fact bdd_above.mono) | 
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 87 | |
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 88 | lemma bdd_below_mono: "bdd_below B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> bdd_below A" | 
| 73271 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 89 | by (fact bdd_below.mono) | 
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 90 | |
| 73271 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 91 | lemma bdd_above_Int1: "bdd_above A \<Longrightarrow> bdd_above (A \<inter> B)" | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 92 | by (fact bdd_above.Int1) | 
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 93 | |
| 73271 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 94 | lemma bdd_above_Int2: "bdd_above B \<Longrightarrow> bdd_above (A \<inter> B)" | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 95 | by (fact bdd_above.Int2) | 
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 96 | |
| 73271 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 97 | lemma bdd_below_Int1: "bdd_below A \<Longrightarrow> bdd_below (A \<inter> B)" | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 98 | by (fact bdd_below.Int1) | 
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 99 | |
| 73271 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 100 | lemma bdd_below_Int2: "bdd_below B \<Longrightarrow> bdd_below (A \<inter> B)" | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 101 | by (fact bdd_below.Int2) | 
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 102 | |
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 103 | lemma bdd_above_Ioo [simp, intro]: "bdd_above {a <..< b}"
 | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 104 | by (auto simp add: bdd_above_def intro!: exI[of _ b] less_imp_le) | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 105 | |
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 106 | lemma bdd_above_Ico [simp, intro]: "bdd_above {a ..< b}"
 | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 107 | by (auto simp add: bdd_above_def intro!: exI[of _ b] less_imp_le) | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 108 | |
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 109 | lemma bdd_above_Iio [simp, intro]: "bdd_above {..< b}"
 | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 110 | by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le) | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 111 | |
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 112 | lemma bdd_above_Ioc [simp, intro]: "bdd_above {a <.. b}"
 | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 113 | by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le) | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 114 | |
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 115 | lemma bdd_above_Icc [simp, intro]: "bdd_above {a .. b}"
 | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 116 | by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le) | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 117 | |
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 118 | lemma bdd_above_Iic [simp, intro]: "bdd_above {.. b}"
 | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 119 | by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le) | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 120 | |
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 121 | lemma bdd_below_Ioo [simp, intro]: "bdd_below {a <..< b}"
 | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 122 | by (auto simp add: bdd_below_def intro!: exI[of _ a] less_imp_le) | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 123 | |
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 124 | lemma bdd_below_Ioc [simp, intro]: "bdd_below {a <.. b}"
 | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 125 | by (auto simp add: bdd_below_def intro!: exI[of _ a] less_imp_le) | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 126 | |
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 127 | lemma bdd_below_Ioi [simp, intro]: "bdd_below {a <..}"
 | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 128 | by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le) | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 129 | |
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 130 | lemma bdd_below_Ico [simp, intro]: "bdd_below {a ..< b}"
 | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 131 | by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le) | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 132 | |
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 133 | lemma bdd_below_Icc [simp, intro]: "bdd_below {a .. b}"
 | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 134 | by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le) | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 135 | |
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 136 | lemma bdd_below_Ici [simp, intro]: "bdd_below {a ..}"
 | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 137 | by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le) | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 138 | |
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 139 | end | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 140 | |
| 73271 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 141 | context order_top | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 142 | begin | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 143 | |
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 144 | lemma bdd_above_top [simp, intro!]: "bdd_above A" | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 145 | by (rule bdd_aboveI [of _ top]) simp | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 146 | |
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 147 | end | 
| 54261 | 148 | |
| 73271 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 149 | context order_bot | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 150 | begin | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 151 | |
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 152 | lemma bdd_below_bot [simp, intro!]: "bdd_below A" | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 153 | by (rule bdd_belowI [of _ bot]) simp | 
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 154 | |
| 
05a873f90655
dedicated locale for preorder and abstract bdd operation
 haftmann parents: 
71834diff
changeset | 155 | end | 
| 54261 | 156 | |
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
58889diff
changeset | 157 | lemma bdd_above_image_mono: "mono f \<Longrightarrow> bdd_above A \<Longrightarrow> bdd_above (f`A)" | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
58889diff
changeset | 158 | by (auto simp: bdd_above_def mono_def) | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
58889diff
changeset | 159 | |
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
58889diff
changeset | 160 | lemma bdd_below_image_mono: "mono f \<Longrightarrow> bdd_below A \<Longrightarrow> bdd_below (f`A)" | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
58889diff
changeset | 161 | by (auto simp: bdd_below_def mono_def) | 
| 63331 | 162 | |
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
58889diff
changeset | 163 | lemma bdd_above_image_antimono: "antimono f \<Longrightarrow> bdd_below A \<Longrightarrow> bdd_above (f`A)" | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
58889diff
changeset | 164 | by (auto simp: bdd_above_def bdd_below_def antimono_def) | 
| 54262 
326fd7103cb4
generalize bdd_above/below_uminus to ordered_ab_group_add
 hoelzl parents: 
54261diff
changeset | 165 | |
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
58889diff
changeset | 166 | lemma bdd_below_image_antimono: "antimono f \<Longrightarrow> bdd_above A \<Longrightarrow> bdd_below (f`A)" | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
58889diff
changeset | 167 | by (auto simp: bdd_above_def bdd_below_def antimono_def) | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
58889diff
changeset | 168 | |
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
58889diff
changeset | 169 | lemma | 
| 54262 
326fd7103cb4
generalize bdd_above/below_uminus to ordered_ab_group_add
 hoelzl parents: 
54261diff
changeset | 170 | fixes X :: "'a::ordered_ab_group_add set" | 
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
58889diff
changeset | 171 | shows bdd_above_uminus[simp]: "bdd_above (uminus ` X) \<longleftrightarrow> bdd_below X" | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
58889diff
changeset | 172 | and bdd_below_uminus[simp]: "bdd_below (uminus ` X) \<longleftrightarrow> bdd_above X" | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
58889diff
changeset | 173 | using bdd_above_image_antimono[of uminus X] bdd_below_image_antimono[of uminus "uminus`X"] | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
58889diff
changeset | 174 | using bdd_below_image_antimono[of uminus X] bdd_above_image_antimono[of uminus "uminus`X"] | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
58889diff
changeset | 175 | by (auto simp: antimono_def image_image) | 
| 54262 
326fd7103cb4
generalize bdd_above/below_uminus to ordered_ab_group_add
 hoelzl parents: 
54261diff
changeset | 176 | |
| 75494 | 177 | subsection \<open>Lattices\<close> | 
| 178 | ||
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 179 | context lattice | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 180 | begin | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 181 | |
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 182 | lemma bdd_above_insert [simp]: "bdd_above (insert a A) = bdd_above A" | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 183 | by (auto simp: bdd_above_def intro: le_supI2 sup_ge1) | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 184 | |
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 185 | lemma bdd_below_insert [simp]: "bdd_below (insert a A) = bdd_below A" | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 186 | by (auto simp: bdd_below_def intro: le_infI2 inf_le1) | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 187 | |
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 188 | lemma bdd_finite [simp]: | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 189 | assumes "finite A" shows bdd_above_finite: "bdd_above A" and bdd_below_finite: "bdd_below A" | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 190 | using assms by (induct rule: finite_induct, auto) | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 191 | |
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 192 | lemma bdd_above_Un [simp]: "bdd_above (A \<union> B) = (bdd_above A \<and> bdd_above B)" | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 193 | proof | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 194 | assume "bdd_above (A \<union> B)" | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 195 | thus "bdd_above A \<and> bdd_above B" unfolding bdd_above_def by auto | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 196 | next | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 197 | assume "bdd_above A \<and> bdd_above B" | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 198 | 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 | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 199 | hence "\<forall>x \<in> A \<union> B. x \<le> sup a b" by (auto intro: Un_iff le_supI1 le_supI2) | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 200 | thus "bdd_above (A \<union> B)" unfolding bdd_above_def .. | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 201 | qed | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 202 | |
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 203 | lemma bdd_below_Un [simp]: "bdd_below (A \<union> B) = (bdd_below A \<and> bdd_below B)" | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 204 | proof | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 205 | assume "bdd_below (A \<union> B)" | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 206 | thus "bdd_below A \<and> bdd_below B" unfolding bdd_below_def by auto | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 207 | next | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 208 | assume "bdd_below A \<and> bdd_below B" | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 209 | 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 | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 210 | hence "\<forall>x \<in> A \<union> B. inf a b \<le> x" by (auto intro: Un_iff le_infI1 le_infI2) | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 211 | thus "bdd_below (A \<union> B)" unfolding bdd_below_def .. | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 212 | qed | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 213 | |
| 67458 | 214 | lemma bdd_above_image_sup[simp]: | 
| 215 | "bdd_above ((\<lambda>x. sup (f x) (g x)) ` A) \<longleftrightarrow> bdd_above (f`A) \<and> bdd_above (g`A)" | |
| 216 | by (auto simp: bdd_above_def intro: le_supI1 le_supI2) | |
| 54259 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 217 | |
| 67458 | 218 | lemma bdd_below_image_inf[simp]: | 
| 219 | "bdd_below ((\<lambda>x. inf (f x) (g x)) ` A) \<longleftrightarrow> bdd_below (f`A) \<and> bdd_below (g`A)" | |
| 220 | by (auto simp: bdd_below_def intro: le_infI1 le_infI2) | |
| 221 | ||
| 222 | lemma bdd_below_UN[simp]: "finite I \<Longrightarrow> bdd_below (\<Union>i\<in>I. A i) = (\<forall>i \<in> I. bdd_below (A i))" | |
| 223 | by (induction I rule: finite.induct) auto | |
| 224 | ||
| 225 | lemma bdd_above_UN[simp]: "finite I \<Longrightarrow> bdd_above (\<Union>i\<in>I. A i) = (\<forall>i \<in> I. bdd_above (A i))" | |
| 226 | by (induction I rule: finite.induct) auto | |
| 54259 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 227 | |
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 228 | end | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 229 | |
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 230 | |
| 60758 | 231 | text \<open> | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 232 | |
| 69593 | 233 | To avoid name classes with the \<^class>\<open>complete_lattice\<close>-class we prefix \<^const>\<open>Sup\<close> and | 
| 234 | \<^const>\<open>Inf\<close> in theorem names with c. | |
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 235 | |
| 60758 | 236 | \<close> | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 237 | |
| 75494 | 238 | subsection \<open>Conditionally complete lattices\<close> | 
| 239 | ||
| 51773 | 240 | class conditionally_complete_lattice = lattice + Sup + Inf + | 
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 241 | assumes cInf_lower: "x \<in> X \<Longrightarrow> bdd_below X \<Longrightarrow> Inf X \<le> x" | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 242 |     and cInf_greatest: "X \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> z \<le> Inf X"
 | 
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 243 | assumes cSup_upper: "x \<in> X \<Longrightarrow> bdd_above X \<Longrightarrow> x \<le> Sup X" | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 244 |     and cSup_least: "X \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup X \<le> z"
 | 
| 33269 
3b7e2dbbd684
New theory SupInf of the supremum and infimum operators for sets of reals.
 paulson parents: diff
changeset | 245 | begin | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 246 | |
| 54259 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 247 | lemma cSup_upper2: "x \<in> X \<Longrightarrow> y \<le> x \<Longrightarrow> bdd_above X \<Longrightarrow> y \<le> Sup X" | 
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 248 | by (metis cSup_upper order_trans) | 
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 249 | |
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 250 | lemma cInf_lower2: "x \<in> X \<Longrightarrow> x \<le> y \<Longrightarrow> bdd_below X \<Longrightarrow> Inf X \<le> y" | 
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 251 | by (metis cInf_lower order_trans) | 
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 252 | |
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 253 | 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"
 | 
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 254 | by (metis cSup_least cSup_upper2) | 
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 255 | |
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 256 | 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"
 | 
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 257 | by (metis cInf_greatest cInf_lower2) | 
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 258 | |
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 259 | lemma cSup_subset_mono: "A \<noteq> {} \<Longrightarrow> bdd_above B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> Sup A \<le> Sup B"
 | 
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 260 | by (metis cSup_least cSup_upper subsetD) | 
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 261 | |
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 262 | lemma cInf_superset_mono: "A \<noteq> {} \<Longrightarrow> bdd_below B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> Inf B \<le> Inf A"
 | 
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 263 | by (metis cInf_greatest cInf_lower subsetD) | 
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 264 | |
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 265 | lemma cSup_eq_maximum: "z \<in> X \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup X = z" | 
| 73411 | 266 | by (intro order.antisym cSup_upper[of z X] cSup_least[of X z]) auto | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 267 | |
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 268 | lemma cInf_eq_minimum: "z \<in> X \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> Inf X = z" | 
| 73411 | 269 | by (intro order.antisym cInf_lower[of z X] cInf_greatest[of X z]) auto | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 270 | |
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 271 | lemma cSup_le_iff: "S \<noteq> {} \<Longrightarrow> bdd_above S \<Longrightarrow> Sup S \<le> a \<longleftrightarrow> (\<forall>x\<in>S. x \<le> a)"
 | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 272 | by (metis order_trans cSup_upper cSup_least) | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 273 | |
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 274 | lemma le_cInf_iff: "S \<noteq> {} \<Longrightarrow> bdd_below S \<Longrightarrow> a \<le> Inf S \<longleftrightarrow> (\<forall>x\<in>S. a \<le> x)"
 | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 275 | by (metis order_trans cInf_lower cInf_greatest) | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 276 | |
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 277 | lemma cSup_eq_non_empty: | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 278 |   assumes 1: "X \<noteq> {}"
 | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 279 | assumes 2: "\<And>x. x \<in> X \<Longrightarrow> x \<le> a" | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 280 | assumes 3: "\<And>y. (\<And>x. x \<in> X \<Longrightarrow> x \<le> y) \<Longrightarrow> a \<le> y" | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 281 | shows "Sup X = a" | 
| 73411 | 282 | by (intro 3 1 order.antisym cSup_least) (auto intro: 2 1 cSup_upper) | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 283 | |
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 284 | lemma cInf_eq_non_empty: | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 285 |   assumes 1: "X \<noteq> {}"
 | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 286 | assumes 2: "\<And>x. x \<in> X \<Longrightarrow> a \<le> x" | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 287 | assumes 3: "\<And>y. (\<And>x. x \<in> X \<Longrightarrow> y \<le> x) \<Longrightarrow> y \<le> a" | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 288 | shows "Inf X = a" | 
| 73411 | 289 | by (intro 3 1 order.antisym cInf_greatest) (auto intro: 2 1 cInf_lower) | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 290 | |
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 291 | lemma cInf_cSup: "S \<noteq> {} \<Longrightarrow> bdd_below S \<Longrightarrow> Inf S = Sup {x. \<forall>s\<in>S. x \<le> s}"
 | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 292 | by (rule cInf_eq_non_empty) (auto intro!: cSup_upper cSup_least simp: bdd_below_def) | 
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51475diff
changeset | 293 | |
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 294 | lemma cSup_cInf: "S \<noteq> {} \<Longrightarrow> bdd_above S \<Longrightarrow> Sup S = Inf {x. \<forall>s\<in>S. s \<le> x}"
 | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 295 | by (rule cSup_eq_non_empty) (auto intro!: cInf_lower cInf_greatest simp: bdd_above_def) | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 296 | |
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 297 | lemma cSup_insert: "X \<noteq> {} \<Longrightarrow> bdd_above X \<Longrightarrow> Sup (insert a X) = sup a (Sup X)"
 | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 298 | by (intro cSup_eq_non_empty) (auto intro: le_supI2 cSup_upper cSup_least) | 
| 33269 
3b7e2dbbd684
New theory SupInf of the supremum and infimum operators for sets of reals.
 paulson parents: diff
changeset | 299 | |
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 300 | lemma cInf_insert: "X \<noteq> {} \<Longrightarrow> bdd_below X \<Longrightarrow> Inf (insert a X) = inf a (Inf X)"
 | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 301 | by (intro cInf_eq_non_empty) (auto intro: le_infI2 cInf_lower cInf_greatest) | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 302 | |
| 54259 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 303 | lemma cSup_singleton [simp]: "Sup {x} = x"
 | 
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 304 | by (intro cSup_eq_maximum) auto | 
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 305 | |
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 306 | lemma cInf_singleton [simp]: "Inf {x} = x"
 | 
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 307 | by (intro cInf_eq_minimum) auto | 
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 308 | |
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 309 | lemma cSup_insert_If:  "bdd_above X \<Longrightarrow> Sup (insert a X) = (if X = {} then a else sup a (Sup X))"
 | 
| 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 310 | using cSup_insert[of X] by simp | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 311 | |
| 54259 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 312 | lemma cInf_insert_If: "bdd_below X \<Longrightarrow> Inf (insert a X) = (if X = {} then a else inf a (Inf X))"
 | 
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 313 | using cInf_insert[of X] by simp | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 314 | |
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 315 | lemma le_cSup_finite: "finite X \<Longrightarrow> x \<in> X \<Longrightarrow> x \<le> Sup X" | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 316 | proof (induct X arbitrary: x rule: finite_induct) | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 317 | case (insert x X y) then show ?case | 
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 318 |     by (cases "X = {}") (auto simp: cSup_insert intro: le_supI2)
 | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 319 | qed simp | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 320 | |
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 321 | lemma cInf_le_finite: "finite X \<Longrightarrow> x \<in> X \<Longrightarrow> Inf X \<le> x" | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 322 | proof (induct X arbitrary: x rule: finite_induct) | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 323 | case (insert x X y) then show ?case | 
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 324 |     by (cases "X = {}") (auto simp: cInf_insert intro: le_infI2)
 | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 325 | qed simp | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 326 | |
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 327 | lemma cSup_eq_Sup_fin: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup X = Sup_fin X"
 | 
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 328 | by (induct X rule: finite_ne_induct) (simp_all add: cSup_insert) | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 329 | |
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 330 | lemma cInf_eq_Inf_fin: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf X = Inf_fin X"
 | 
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 331 | by (induct X rule: finite_ne_induct) (simp_all add: cInf_insert) | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 332 | |
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 333 | lemma cSup_atMost[simp]: "Sup {..x} = x"
 | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 334 | by (auto intro!: cSup_eq_maximum) | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 335 | |
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 336 | lemma cSup_greaterThanAtMost[simp]: "y < x \<Longrightarrow> Sup {y<..x} = x"
 | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 337 | by (auto intro!: cSup_eq_maximum) | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 338 | |
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 339 | lemma cSup_atLeastAtMost[simp]: "y \<le> x \<Longrightarrow> Sup {y..x} = x"
 | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 340 | by (auto intro!: cSup_eq_maximum) | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 341 | |
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 342 | lemma cInf_atLeast[simp]: "Inf {x..} = x"
 | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 343 | by (auto intro!: cInf_eq_minimum) | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 344 | |
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 345 | lemma cInf_atLeastLessThan[simp]: "y < x \<Longrightarrow> Inf {y..<x} = y"
 | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 346 | by (auto intro!: cInf_eq_minimum) | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 347 | |
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 348 | lemma cInf_atLeastAtMost[simp]: "y \<le> x \<Longrightarrow> Inf {y..x} = y"
 | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 349 | by (auto intro!: cInf_eq_minimum) | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 350 | |
| 69275 | 351 | lemma cINF_lower: "bdd_below (f ` A) \<Longrightarrow> x \<in> A \<Longrightarrow> \<Sqinter>(f ` A) \<le> f x" | 
| 56166 | 352 | using cInf_lower [of _ "f ` A"] by simp | 
| 54259 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 353 | |
| 69275 | 354 | lemma cINF_greatest: "A \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> m \<le> f x) \<Longrightarrow> m \<le> \<Sqinter>(f ` A)"
 | 
| 56166 | 355 | using cInf_greatest [of "f ` A"] by auto | 
| 54259 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 356 | |
| 69275 | 357 | lemma cSUP_upper: "x \<in> A \<Longrightarrow> bdd_above (f ` A) \<Longrightarrow> f x \<le> \<Squnion>(f ` A)" | 
| 56166 | 358 | using cSup_upper [of _ "f ` A"] by simp | 
| 54259 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 359 | |
| 69275 | 360 | lemma cSUP_least: "A \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x \<le> M) \<Longrightarrow> \<Squnion>(f ` A) \<le> M"
 | 
| 56166 | 361 | using cSup_least [of "f ` A"] by auto | 
| 54259 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 362 | |
| 69275 | 363 | lemma cINF_lower2: "bdd_below (f ` A) \<Longrightarrow> x \<in> A \<Longrightarrow> f x \<le> u \<Longrightarrow> \<Sqinter>(f ` A) \<le> u" | 
| 63092 | 364 | by (auto intro: cINF_lower order_trans) | 
| 54259 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 365 | |
| 69275 | 366 | lemma cSUP_upper2: "bdd_above (f ` A) \<Longrightarrow> x \<in> A \<Longrightarrow> u \<le> f x \<Longrightarrow> u \<le> \<Squnion>(f ` A)" | 
| 63092 | 367 | by (auto intro: cSUP_upper order_trans) | 
| 54259 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 368 | |
| 67613 | 369 | lemma cSUP_const [simp]: "A \<noteq> {} \<Longrightarrow> (\<Squnion>x\<in>A. c) = c"
 | 
| 73411 | 370 | by (intro order.antisym cSUP_least) (auto intro: cSUP_upper) | 
| 54261 | 371 | |
| 67613 | 372 | lemma cINF_const [simp]: "A \<noteq> {} \<Longrightarrow> (\<Sqinter>x\<in>A. c) = c"
 | 
| 73411 | 373 | by (intro order.antisym cINF_greatest) (auto intro: cINF_lower) | 
| 54261 | 374 | |
| 69275 | 375 | lemma 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)"
 | 
| 63092 | 376 | by (metis cINF_greatest cINF_lower order_trans) | 
| 54259 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 377 | |
| 69275 | 378 | lemma 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)"
 | 
| 63092 | 379 | by (metis cSUP_least cSUP_upper order_trans) | 
| 54259 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 380 | |
| 67613 | 381 | lemma less_cINF_D: "bdd_below (f`A) \<Longrightarrow> y < (\<Sqinter>i\<in>A. f i) \<Longrightarrow> i \<in> A \<Longrightarrow> y < f i" | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54262diff
changeset | 382 | by (metis cINF_lower less_le_trans) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54262diff
changeset | 383 | |
| 67613 | 384 | lemma cSUP_lessD: "bdd_above (f`A) \<Longrightarrow> (\<Squnion>i\<in>A. f i) < y \<Longrightarrow> i \<in> A \<Longrightarrow> f i < y" | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54262diff
changeset | 385 | by (metis cSUP_upper le_less_trans) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54262diff
changeset | 386 | |
| 69275 | 387 | lemma cINF_insert: "A \<noteq> {} \<Longrightarrow> bdd_below (f ` A) \<Longrightarrow> \<Sqinter>(f ` insert a A) = inf (f a) (\<Sqinter>(f ` A))"
 | 
| 71238 | 388 | by (simp add: cInf_insert) | 
| 54259 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 389 | |
| 69275 | 390 | lemma cSUP_insert: "A \<noteq> {} \<Longrightarrow> bdd_above (f ` A) \<Longrightarrow> \<Squnion>(f ` insert a A) = sup (f a) (\<Squnion>(f ` A))"
 | 
| 71238 | 391 | by (simp add: cSup_insert) | 
| 54259 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 392 | |
| 69275 | 393 | 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> \<Sqinter>(f ` A) \<le> \<Sqinter>(g ` B)"
 | 
| 56166 | 394 | using cInf_mono [of "g ` B" "f ` A"] by auto | 
| 54259 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 395 | |
| 69275 | 396 | 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> \<Squnion>(f ` A) \<le> \<Squnion>(g ` B)"
 | 
| 56166 | 397 | using cSup_mono [of "f ` A" "g ` B"] by auto | 
| 54259 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 398 | |
| 69275 | 399 | 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> \<Sqinter>(g ` B) \<le> \<Sqinter>(f ` A)"
 | 
| 54259 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 400 | by (rule cINF_mono) auto | 
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 401 | |
| 71096 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69611diff
changeset | 402 | lemma cSUP_subset_mono: | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69611diff
changeset | 403 |   "\<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)"
 | 
| 54259 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 404 | by (rule cSUP_mono) auto | 
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 405 | |
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 406 | 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)"
 | 
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 407 | by (metis cInf_superset_mono lattice_class.inf_sup_ord(1) le_infI1) | 
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 408 | |
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 409 | 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) "
 | 
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 410 | by (metis cSup_subset_mono lattice_class.inf_sup_ord(1) le_supI1) | 
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 411 | |
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 412 | 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)"
 | 
| 73411 | 413 | by (intro order.antisym le_infI cInf_greatest cInf_lower) (auto intro: le_infI1 le_infI2 cInf_lower) | 
| 54259 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 414 | |
| 69275 | 415 | lemma 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)"
 | 
| 71238 | 416 | using cInf_union_distrib [of "f ` A" "f ` B"] by (simp add: image_Un) | 
| 54259 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 417 | |
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 418 | 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)"
 | 
| 73411 | 419 | by (intro order.antisym le_supI cSup_least cSup_upper) (auto intro: le_supI1 le_supI2 cSup_upper) | 
| 54259 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 420 | |
| 69275 | 421 | lemma 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)"
 | 
| 71238 | 422 | using cSup_union_distrib [of "f ` A" "f ` B"] by (simp add: image_Un) | 
| 54259 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 423 | |
| 69275 | 424 | lemma 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))"
 | 
| 73411 | 425 | by (intro order.antisym le_infI cINF_greatest cINF_lower2) | 
| 54259 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 426 | (auto intro: le_infI1 le_infI2 cINF_greatest cINF_lower le_infI) | 
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 427 | |
| 69275 | 428 | lemma 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))"
 | 
| 73411 | 429 | by (intro order.antisym le_supI cSUP_least cSUP_upper2) | 
| 54259 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 430 | (auto intro: le_supI1 le_supI2 cSUP_least cSUP_upper le_supI) | 
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 431 | |
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57275diff
changeset | 432 | lemma cInf_le_cSup: | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57275diff
changeset | 433 |   "A \<noteq> {} \<Longrightarrow> bdd_above A \<Longrightarrow> bdd_below A \<Longrightarrow> Inf A \<le> Sup A"
 | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57275diff
changeset | 434 | by (auto intro!: cSup_upper2[of "SOME a. a \<in> A"] intro: someI cInf_lower) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57275diff
changeset | 435 | |
| 74007 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73411diff
changeset | 436 | context | 
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73411diff
changeset | 437 | fixes f :: "'a \<Rightarrow> 'b::conditionally_complete_lattice" | 
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73411diff
changeset | 438 | assumes "mono f" | 
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73411diff
changeset | 439 | begin | 
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73411diff
changeset | 440 | |
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73411diff
changeset | 441 |   lemma mono_cInf: "\<lbrakk>bdd_below A; A\<noteq>{}\<rbrakk> \<Longrightarrow> f (Inf A) \<le> (INF x\<in>A. f x)"
 | 
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73411diff
changeset | 442 | by (simp add: \<open>mono f\<close> conditionally_complete_lattice_class.cINF_greatest cInf_lower monoD) | 
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73411diff
changeset | 443 | |
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73411diff
changeset | 444 |   lemma mono_cSup: "\<lbrakk>bdd_above A; A\<noteq>{}\<rbrakk> \<Longrightarrow> (SUP x\<in>A. f x) \<le> f (Sup A)"
 | 
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73411diff
changeset | 445 | by (simp add: \<open>mono f\<close> conditionally_complete_lattice_class.cSUP_least cSup_upper monoD) | 
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73411diff
changeset | 446 | |
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73411diff
changeset | 447 |   lemma mono_cINF: "\<lbrakk>bdd_below (A`I); I\<noteq>{}\<rbrakk> \<Longrightarrow> f (INF i\<in>I. A i) \<le> (INF x\<in>I. f (A x))"
 | 
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73411diff
changeset | 448 | by (simp add: \<open>mono f\<close> conditionally_complete_lattice_class.cINF_greatest cINF_lower monoD) | 
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73411diff
changeset | 449 | |
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73411diff
changeset | 450 |   lemma mono_cSUP: "\<lbrakk>bdd_above (A`I); I\<noteq>{}\<rbrakk> \<Longrightarrow> (SUP x\<in>I. f (A x)) \<le> f (SUP i\<in>I. A i)"
 | 
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73411diff
changeset | 451 | by (simp add: \<open>mono f\<close> conditionally_complete_lattice_class.cSUP_least cSUP_upper monoD) | 
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73411diff
changeset | 452 | |
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73411diff
changeset | 453 | end | 
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73411diff
changeset | 454 | |
| 33269 
3b7e2dbbd684
New theory SupInf of the supremum and infimum operators for sets of reals.
 paulson parents: diff
changeset | 455 | end | 
| 
3b7e2dbbd684
New theory SupInf of the supremum and infimum operators for sets of reals.
 paulson parents: diff
changeset | 456 | |
| 75494 | 457 | text \<open>The special case of well-orderings\<close> | 
| 458 | ||
| 459 | lemma wellorder_InfI: | |
| 460 |   fixes k :: "'a::{wellorder,conditionally_complete_lattice}"
 | |
| 461 | assumes "k \<in> A" shows "Inf A \<in> A" | |
| 462 | using wellorder_class.LeastI [of "\<lambda>x. x \<in> A" k] | |
| 463 | by (simp add: Least_le assms cInf_eq_minimum) | |
| 464 | ||
| 465 | lemma wellorder_Inf_le1: | |
| 466 |   fixes k :: "'a::{wellorder,conditionally_complete_lattice}"
 | |
| 467 | assumes "k \<in> A" shows "Inf A \<le> k" | |
| 468 | by (meson Least_le assms bdd_below.I cInf_lower) | |
| 469 | ||
| 470 | subsection \<open>Complete lattices\<close> | |
| 471 | ||
| 51773 | 472 | instance complete_lattice \<subseteq> conditionally_complete_lattice | 
| 61169 | 473 | by standard (auto intro: Sup_upper Sup_least Inf_lower Inf_greatest) | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 474 | |
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 475 | lemma cSup_eq: | 
| 51773 | 476 |   fixes a :: "'a :: {conditionally_complete_lattice, no_bot}"
 | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 477 | assumes upper: "\<And>x. x \<in> X \<Longrightarrow> x \<le> a" | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 478 | assumes least: "\<And>y. (\<And>x. x \<in> X \<Longrightarrow> x \<le> y) \<Longrightarrow> a \<le> y" | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 479 | shows "Sup X = a" | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 480 | proof cases | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 481 |   assume "X = {}" with lt_ex[of a] least show ?thesis by (auto simp: less_le_not_le)
 | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 482 | qed (intro cSup_eq_non_empty assms) | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 483 | |
| 77934 | 484 | lemma cSup_unique: | 
| 485 |   fixes b :: "'a :: {conditionally_complete_lattice, no_bot}"
 | |
| 486 | assumes "\<And>c. (\<forall>x \<in> s. x \<le> c) \<longleftrightarrow> b \<le> c" | |
| 487 | shows "Sup s = b" | |
| 488 | by (metis assms cSup_eq order.refl) | |
| 489 | ||
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 490 | lemma cInf_eq: | 
| 51773 | 491 |   fixes a :: "'a :: {conditionally_complete_lattice, no_top}"
 | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 492 | assumes upper: "\<And>x. x \<in> X \<Longrightarrow> a \<le> x" | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 493 | assumes least: "\<And>y. (\<And>x. x \<in> X \<Longrightarrow> y \<le> x) \<Longrightarrow> y \<le> a" | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 494 | shows "Inf X = a" | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 495 | proof cases | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 496 |   assume "X = {}" with gt_ex[of a] least show ?thesis by (auto simp: less_le_not_le)
 | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 497 | qed (intro cInf_eq_non_empty assms) | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 498 | |
| 77934 | 499 | lemma cInf_unique: | 
| 500 |   fixes b :: "'a :: {conditionally_complete_lattice, no_top}"
 | |
| 501 | assumes "\<And>c. (\<forall>x \<in> s. x \<ge> c) \<longleftrightarrow> b \<ge> c" | |
| 502 | shows "Inf s = b" | |
| 503 | by (meson assms cInf_eq order.refl) | |
| 504 | ||
| 51773 | 505 | class conditionally_complete_linorder = conditionally_complete_lattice + linorder | 
| 33269 
3b7e2dbbd684
New theory SupInf of the supremum and infimum operators for sets of reals.
 paulson parents: diff
changeset | 506 | begin | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 507 | |
| 63331 | 508 | lemma less_cSup_iff: | 
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 509 |   "X \<noteq> {} \<Longrightarrow> bdd_above X \<Longrightarrow> y < Sup X \<longleftrightarrow> (\<exists>x\<in>X. y < x)"
 | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 510 | by (rule iffI) (metis cSup_least not_less, metis cSup_upper less_le_trans) | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 511 | |
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 512 | lemma cInf_less_iff: "X \<noteq> {} \<Longrightarrow> bdd_below X \<Longrightarrow> Inf X < y \<longleftrightarrow> (\<exists>x\<in>X. x < y)"
 | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 513 | by (rule iffI) (metis cInf_greatest not_less, metis cInf_lower le_less_trans) | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 514 | |
| 67613 | 515 | lemma 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)"
 | 
| 56166 | 516 | using cInf_less_iff[of "f`A"] by auto | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54262diff
changeset | 517 | |
| 67613 | 518 | lemma 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)"
 | 
| 56166 | 519 | using less_cSup_iff[of "f`A"] by auto | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
54262diff
changeset | 520 | |
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 521 | lemma less_cSupE: | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 522 |   assumes "y < Sup X" "X \<noteq> {}" obtains x where "x \<in> X" "y < x"
 | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 523 | by (metis cSup_least assms not_le that) | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 524 | |
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51475diff
changeset | 525 | lemma less_cSupD: | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51475diff
changeset | 526 |   "X \<noteq> {} \<Longrightarrow> z < Sup X \<Longrightarrow> \<exists>x\<in>X. z < x"
 | 
| 61824 
dcbe9f756ae0
not_leE -> not_le_imp_less and other tidying
 paulson <lp15@cam.ac.uk> parents: 
61169diff
changeset | 527 | by (metis less_cSup_iff not_le_imp_less bdd_above_def) | 
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51475diff
changeset | 528 | |
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51475diff
changeset | 529 | lemma cInf_lessD: | 
| 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51475diff
changeset | 530 |   "X \<noteq> {} \<Longrightarrow> Inf X < z \<Longrightarrow> \<exists>x\<in>X. x < z"
 | 
| 61824 
dcbe9f756ae0
not_leE -> not_le_imp_less and other tidying
 paulson <lp15@cam.ac.uk> parents: 
61169diff
changeset | 531 | by (metis cInf_less_iff not_le_imp_less bdd_below_def) | 
| 51518 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 hoelzl parents: 
51475diff
changeset | 532 | |
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 533 | lemma complete_interval: | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 534 | assumes "a < b" and "P a" and "\<not> P b" | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 535 | shows "\<exists>c. a \<le> c \<and> c \<le> b \<and> (\<forall>x. a \<le> x \<and> x < c \<longrightarrow> P x) \<and> | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 536 | (\<forall>d. (\<forall>x. a \<le> x \<and> x < d \<longrightarrow> P x) \<longrightarrow> d \<le> c)" | 
| 75669 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75494diff
changeset | 537 | proof (rule exI [where x = "Sup {d. \<forall>x. a \<le> x \<and> x < d \<longrightarrow> P x}"], safe)
 | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 538 |   show "a \<le> Sup {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c}"
 | 
| 54258 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 hoelzl parents: 
54257diff
changeset | 539 | by (rule cSup_upper, auto simp: bdd_above_def) | 
| 60758 | 540 | (metis \<open>a < b\<close> \<open>\<not> P b\<close> linear less_le) | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 541 | next | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 542 |   show "Sup {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c} \<le> b"
 | 
| 75669 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75494diff
changeset | 543 | by (rule cSup_least) | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75494diff
changeset | 544 | (use \<open>a<b\<close> \<open>\<not> P b\<close> in \<open>auto simp add: less_le_not_le\<close>) | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 545 | next | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 546 | fix x | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 547 |   assume x: "a \<le> x" and lt: "x < Sup {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c}"
 | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 548 | show "P x" | 
| 75669 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75494diff
changeset | 549 | by (rule less_cSupE [OF lt]) (use less_le_not_le x in \<open>auto\<close>) | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 550 | next | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 551 | fix d | 
| 75669 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75494diff
changeset | 552 | assume 0: "\<forall>x. a \<le> x \<and> x < d \<longrightarrow> P x" | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75494diff
changeset | 553 |   then have "d \<in> {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c}"
 | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75494diff
changeset | 554 | by auto | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75494diff
changeset | 555 |   moreover have "bdd_above {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c}"
 | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75494diff
changeset | 556 | unfolding bdd_above_def using \<open>a<b\<close> \<open>\<not> P b\<close> linear | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75494diff
changeset | 557 | by (simp add: less_le) blast | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75494diff
changeset | 558 |   ultimately show "d \<le> Sup {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c}"
 | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
75494diff
changeset | 559 | by (auto simp: cSup_upper) | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 560 | qed | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 561 | |
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 562 | end | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 563 | |
| 75494 | 564 | subsection \<open>Instances\<close> | 
| 565 | ||
| 60172 
423273355b55
rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
 hoelzl parents: 
59452diff
changeset | 566 | instance complete_linorder < conditionally_complete_linorder | 
| 
423273355b55
rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
 hoelzl parents: 
59452diff
changeset | 567 | .. | 
| 
423273355b55
rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
 hoelzl parents: 
59452diff
changeset | 568 | |
| 54259 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 569 | lemma cSup_eq_Max: "finite (X::'a::conditionally_complete_linorder set) \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup X = Max X"
 | 
| 67484 | 570 | using cSup_eq_Sup_fin[of X] by (simp add: Sup_fin_Max) | 
| 51775 
408d937c9486
revert #916271d52466; add non-topological linear_continuum type class; show linear_continuum_topology is a perfect_space
 hoelzl parents: 
51773diff
changeset | 571 | |
| 54259 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 572 | lemma cInf_eq_Min: "finite (X::'a::conditionally_complete_linorder set) \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf X = Min X"
 | 
| 67484 | 573 | using cInf_eq_Inf_fin[of X] by (simp add: Inf_fin_Min) | 
| 51775 
408d937c9486
revert #916271d52466; add non-topological linear_continuum type class; show linear_continuum_topology is a perfect_space
 hoelzl parents: 
51773diff
changeset | 574 | |
| 54257 
5c7a3b6b05a9
generalize SUP and INF to the syntactic type classes Sup and Inf
 hoelzl parents: 
53216diff
changeset | 575 | lemma cSup_lessThan[simp]: "Sup {..<x::'a::{conditionally_complete_linorder, no_bot, dense_linorder}} = x"
 | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 576 | by (auto intro!: cSup_eq_non_empty intro: dense_le) | 
| 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 577 | |
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57275diff
changeset | 578 | lemma cSup_greaterThanLessThan[simp]: "y < x \<Longrightarrow> Sup {y<..<x::'a::{conditionally_complete_linorder, dense_linorder}} = x"
 | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57275diff
changeset | 579 | by (auto intro!: cSup_eq_non_empty intro: dense_le_bounded) | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 580 | |
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57275diff
changeset | 581 | lemma cSup_atLeastLessThan[simp]: "y < x \<Longrightarrow> Sup {y..<x::'a::{conditionally_complete_linorder, dense_linorder}} = x"
 | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57275diff
changeset | 582 | by (auto intro!: cSup_eq_non_empty intro: dense_le_bounded) | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 583 | |
| 54257 
5c7a3b6b05a9
generalize SUP and INF to the syntactic type classes Sup and Inf
 hoelzl parents: 
53216diff
changeset | 584 | lemma cInf_greaterThan[simp]: "Inf {x::'a::{conditionally_complete_linorder, no_top, dense_linorder} <..} = x"
 | 
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57275diff
changeset | 585 | by (auto intro!: cInf_eq_non_empty intro: dense_ge) | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 586 | |
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57275diff
changeset | 587 | lemma cInf_greaterThanAtMost[simp]: "y < x \<Longrightarrow> Inf {y<..x::'a::{conditionally_complete_linorder, dense_linorder}} = y"
 | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57275diff
changeset | 588 | by (auto intro!: cInf_eq_non_empty intro: dense_ge_bounded) | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 589 | |
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57275diff
changeset | 590 | lemma cInf_greaterThanLessThan[simp]: "y < x \<Longrightarrow> Inf {y<..<x::'a::{conditionally_complete_linorder, dense_linorder}} = y"
 | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57275diff
changeset | 591 | by (auto intro!: cInf_eq_non_empty intro: dense_ge_bounded) | 
| 51475 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 hoelzl parents: 
46757diff
changeset | 592 | |
| 78698 | 593 | lemma Sup_inverse_eq_inverse_Inf: | 
| 594 |   fixes f::"'b \<Rightarrow> 'a::{conditionally_complete_linorder,linordered_field}"
 | |
| 595 | assumes "bdd_above (range f)" "L > 0" and geL: "\<And>x. f x \<ge> L" | |
| 596 | shows "(SUP x. 1 / f x) = 1 / (INF x. f x)" | |
| 597 | proof (rule antisym) | |
| 598 | have bdd_f: "bdd_below (range f)" | |
| 599 | by (meson assms bdd_belowI2) | |
| 600 | have "Inf (range f) \<ge> L" | |
| 601 | by (simp add: cINF_greatest geL) | |
| 602 | have bdd_invf: "bdd_above (range (\<lambda>x. 1 / f x))" | |
| 603 | proof (rule bdd_aboveI2) | |
| 604 | show "\<And>x. 1 / f x \<le> 1/L" | |
| 605 | using assms by (auto simp: divide_simps) | |
| 606 | qed | |
| 607 | moreover have le_inverse_Inf: "1 / f x \<le> 1 / Inf (range f)" for x | |
| 608 | proof - | |
| 609 | have "Inf (range f) \<le> f x" | |
| 610 | by (simp add: bdd_f cInf_lower) | |
| 611 | then show ?thesis | |
| 612 | using assms \<open>L \<le> Inf (range f)\<close> by (auto simp: divide_simps) | |
| 613 | qed | |
| 614 | ultimately show *: "(SUP x. 1 / f x) \<le> 1 / Inf (range f)" | |
| 615 | by (auto simp: cSup_le_iff cINF_lower) | |
| 616 | have "1 / (SUP x. 1 / f x) \<le> f y" for y | |
| 617 | proof (cases "(SUP x. 1 / f x) < 0") | |
| 618 | case True | |
| 619 | with assms show ?thesis | |
| 620 | by (meson less_asym' order_trans linorder_not_le zero_le_divide_1_iff) | |
| 621 | next | |
| 622 | case False | |
| 623 | have "1 / f y \<le> (SUP x. 1 / f x)" | |
| 624 | by (simp add: bdd_invf cSup_upper) | |
| 625 | with False assms show ?thesis | |
| 626 | by (metis (no_types) div_by_1 divide_divide_eq_right dual_order.strict_trans1 inverse_eq_divide | |
| 627 | inverse_le_imp_le mult.left_neutral) | |
| 628 | qed | |
| 629 | then have "1 / (SUP x. 1 / f x) \<le> Inf (range f)" | |
| 630 | using bdd_f by (simp add: le_cInf_iff) | |
| 631 | moreover have "(SUP x. 1 / f x) > 0" | |
| 632 | using assms cSUP_upper [OF _ bdd_invf] by (meson UNIV_I less_le_trans zero_less_divide_1_iff) | |
| 633 | ultimately show "1 / Inf (range f) \<le> (SUP t. 1 / f t)" | |
| 634 | using \<open>L \<le> Inf (range f)\<close> \<open>L>0\<close> by (auto simp: field_simps) | |
| 635 | qed | |
| 636 | ||
| 637 | lemma Inf_inverse_eq_inverse_Sup: | |
| 638 |   fixes f::"'b \<Rightarrow> 'a::{conditionally_complete_linorder,linordered_field}"
 | |
| 639 | assumes "bdd_above (range f)" "L > 0" and geL: "\<And>x. f x \<ge> L" | |
| 640 | shows "(INF x. 1 / f x) = 1 / (SUP x. f x)" | |
| 641 | proof - | |
| 642 | obtain M where "M>0" and M: "\<And>x. f x \<le> M" | |
| 643 | by (meson assms cSup_upper dual_order.strict_trans1 rangeI) | |
| 644 | have bdd: "bdd_above (range (inverse \<circ> f))" | |
| 645 | using assms le_imp_inverse_le by (auto simp: bdd_above_def) | |
| 646 | have "f x > 0" for x | |
| 647 | using \<open>L>0\<close> geL order_less_le_trans by blast | |
| 648 | then have [simp]: "1 / inverse(f x) = f x" "1 / M \<le> 1 / f x" for x | |
| 649 | using M \<open>M>0\<close> by (auto simp: divide_simps) | |
| 650 | show ?thesis | |
| 651 | using Sup_inverse_eq_inverse_Inf [OF bdd, of "inverse M"] \<open>M>0\<close> | |
| 652 | by (simp add: inverse_eq_divide) | |
| 653 | qed | |
| 654 | ||
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69593diff
changeset | 655 | lemma Inf_insert_finite: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69593diff
changeset | 656 | fixes S :: "'a::conditionally_complete_linorder set" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69593diff
changeset | 657 |   shows "finite S \<Longrightarrow> Inf (insert x S) = (if S = {} then x else min x (Inf S))"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69593diff
changeset | 658 | by (simp add: cInf_eq_Min) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69593diff
changeset | 659 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69593diff
changeset | 660 | lemma Sup_insert_finite: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69593diff
changeset | 661 | fixes S :: "'a::conditionally_complete_linorder set" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69593diff
changeset | 662 |   shows "finite S \<Longrightarrow> Sup (insert x S) = (if S = {} then x else max x (Sup S))"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69593diff
changeset | 663 | by (simp add: cSup_insert sup_max) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69593diff
changeset | 664 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69593diff
changeset | 665 | lemma finite_imp_less_Inf: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69593diff
changeset | 666 | fixes a :: "'a::conditionally_complete_linorder" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69593diff
changeset | 667 | shows "\<lbrakk>finite X; x \<in> X; \<And>x. x\<in>X \<Longrightarrow> a < x\<rbrakk> \<Longrightarrow> a < Inf X" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69593diff
changeset | 668 | by (induction X rule: finite_induct) (simp_all add: cInf_eq_Min Inf_insert_finite) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69593diff
changeset | 669 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69593diff
changeset | 670 | lemma finite_less_Inf_iff: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69593diff
changeset | 671 | fixes a :: "'a :: conditionally_complete_linorder" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69593diff
changeset | 672 |   shows "\<lbrakk>finite X; X \<noteq> {}\<rbrakk> \<Longrightarrow> a < Inf X \<longleftrightarrow> (\<forall>x \<in> X. a < x)"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69593diff
changeset | 673 | by (auto simp: cInf_eq_Min) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69593diff
changeset | 674 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69593diff
changeset | 675 | lemma finite_imp_Sup_less: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69593diff
changeset | 676 | fixes a :: "'a::conditionally_complete_linorder" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69593diff
changeset | 677 | shows "\<lbrakk>finite X; x \<in> X; \<And>x. x\<in>X \<Longrightarrow> a > x\<rbrakk> \<Longrightarrow> a > Sup X" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69593diff
changeset | 678 | by (induction X rule: finite_induct) (simp_all add: cSup_eq_Max Sup_insert_finite) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69593diff
changeset | 679 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69593diff
changeset | 680 | lemma finite_Sup_less_iff: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69593diff
changeset | 681 | fixes a :: "'a :: conditionally_complete_linorder" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69593diff
changeset | 682 |   shows "\<lbrakk>finite X; X \<noteq> {}\<rbrakk> \<Longrightarrow> a > Sup X \<longleftrightarrow> (\<forall>x \<in> X. a > x)"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69593diff
changeset | 683 | by (auto simp: cSup_eq_Max) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69593diff
changeset | 684 | |
| 54259 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 685 | class linear_continuum = conditionally_complete_linorder + dense_linorder + | 
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 686 | assumes UNIV_not_singleton: "\<exists>a b::'a. a \<noteq> b" | 
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 687 | begin | 
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 688 | |
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 689 | lemma ex_gt_or_lt: "\<exists>b. a < b \<or> b < a" | 
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 690 | by (metis UNIV_not_singleton neq_iff) | 
| 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 691 | |
| 33269 
3b7e2dbbd684
New theory SupInf of the supremum and infimum operators for sets of reals.
 paulson parents: diff
changeset | 692 | end | 
| 54259 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 693 | |
| 76770 | 694 | |
| 695 | context | |
| 696 |   fixes f::"'a \<Rightarrow> 'b::{conditionally_complete_linorder,ordered_ab_group_add}"
 | |
| 697 | begin | |
| 698 | ||
| 699 | lemma bdd_above_uminus_image: "bdd_above ((\<lambda>x. - f x) ` A) \<longleftrightarrow> bdd_below (f ` A)" | |
| 700 | by (metis bdd_above_uminus image_image) | |
| 701 | ||
| 702 | lemma bdd_below_uminus_image: "bdd_below ((\<lambda>x. - f x) ` A) \<longleftrightarrow> bdd_above (f ` A)" | |
| 703 | by (metis bdd_below_uminus image_image) | |
| 704 | ||
| 705 | lemma uminus_cSUP: | |
| 706 |   assumes "bdd_above (f ` A)" "A \<noteq> {}"
 | |
| 707 | shows "- (SUP x\<in>A. f x) = (INF x\<in>A. - f x)" | |
| 708 | proof (rule antisym) | |
| 709 | show "(INF x\<in>A. - f x) \<le> - Sup (f ` A)" | |
| 710 | by (metis cINF_lower cSUP_least bdd_below_uminus_image assms le_minus_iff) | |
| 711 | have *: "\<And>x. x \<in>A \<Longrightarrow> f x \<le> Sup (f ` A)" | |
| 712 | by (simp add: assms cSup_upper) | |
| 713 | then show "- Sup (f ` A) \<le> (INF x\<in>A. - f x)" | |
| 714 | by (simp add: assms cINF_greatest) | |
| 715 | qed | |
| 716 | ||
| 717 | end | |
| 718 | ||
| 719 | context | |
| 720 |   fixes f::"'a \<Rightarrow> 'b::{conditionally_complete_linorder,ordered_ab_group_add}"
 | |
| 721 | begin | |
| 722 | ||
| 723 | lemma uminus_cINF: | |
| 724 |   assumes "bdd_below (f ` A)" "A \<noteq> {}"
 | |
| 725 | shows "- (INF x\<in>A. f x) = (SUP x\<in>A. - f x)" | |
| 726 | by (metis (mono_tags, lifting) INF_cong uminus_cSUP assms bdd_above_uminus_image minus_equation_iff) | |
| 727 | ||
| 728 | lemma Sup_add_eq: | |
| 729 |   assumes "bdd_above (f ` A)" "A \<noteq> {}"
 | |
| 730 | shows "(SUP x\<in>A. a + f x) = a + (SUP x\<in>A. f x)" (is "?L=?R") | |
| 731 | proof (rule antisym) | |
| 732 | have bdd: "bdd_above ((\<lambda>x. a + f x) ` A)" | |
| 733 | by (metis assms bdd_above_image_mono image_image mono_add) | |
| 734 | with assms show "?L \<le> ?R" | |
| 735 | by (simp add: assms cSup_le_iff cSUP_upper) | |
| 736 | have "\<And>x. x \<in> A \<Longrightarrow> f x \<le> (SUP x\<in>A. a + f x) - a" | |
| 737 | by (simp add: bdd cSup_upper le_diff_eq) | |
| 738 |   with \<open>A \<noteq> {}\<close> have "\<Squnion> (f ` A) \<le> (\<Squnion>x\<in>A. a + f x) - a"
 | |
| 739 | by (simp add: cSUP_least) | |
| 740 | then show "?R \<le> ?L" | |
| 741 | by (metis add.commute le_diff_eq) | |
| 742 | qed | |
| 743 | ||
| 744 | lemma Inf_add_eq: \<comment>\<open>you don't get a shorter proof by duality\<close> | |
| 745 |   assumes "bdd_below (f ` A)" "A \<noteq> {}"
 | |
| 746 | shows "(INF x\<in>A. a + f x) = a + (INF x\<in>A. f x)" (is "?L=?R") | |
| 747 | proof (rule antisym) | |
| 748 | show "?R \<le> ?L" | |
| 749 | using assms mono_add mono_cINF by blast | |
| 750 | have bdd: "bdd_below ((\<lambda>x. a + f x) ` A)" | |
| 751 | by (metis add_left_mono assms(1) bdd_below.E bdd_below.I2 imageI) | |
| 752 | with assms have "\<And>x. x \<in> A \<Longrightarrow> f x \<ge> (INF x\<in>A. a + f x) - a" | |
| 753 | by (simp add: cInf_lower diff_le_eq) | |
| 754 |   with \<open>A \<noteq> {}\<close> have "(\<Sqinter>x\<in>A. a + f x) - a \<le> \<Sqinter> (f ` A)"
 | |
| 755 | by (simp add: cINF_greatest) | |
| 756 | with assms show "?L \<le> ?R" | |
| 757 | by (metis add.commute diff_le_eq) | |
| 758 | qed | |
| 759 | ||
| 760 | end | |
| 761 | ||
| 54281 | 762 | instantiation nat :: conditionally_complete_linorder | 
| 763 | begin | |
| 764 | ||
| 71833 
ff6f3b09b8b4
abbrevs for the Uniq quantifier; trying Sup_nat_def to allow zero (experimentally)
 paulson <lp15@cam.ac.uk> parents: 
71238diff
changeset | 765 | definition "Sup (X::nat set) = (if X={} then 0 else Max X)"
 | 
| 54281 | 766 | definition "Inf (X::nat set) = (LEAST n. n \<in> X)" | 
| 767 | ||
| 768 | lemma bdd_above_nat: "bdd_above X \<longleftrightarrow> finite (X::nat set)" | |
| 769 | proof | |
| 770 | assume "bdd_above X" | |
| 771 |   then obtain z where "X \<subseteq> {.. z}"
 | |
| 772 | by (auto simp: bdd_above_def) | |
| 773 | then show "finite X" | |
| 774 | by (rule finite_subset) simp | |
| 775 | qed simp | |
| 776 | ||
| 777 | instance | |
| 778 | proof | |
| 63540 | 779 | fix x :: nat | 
| 780 | fix X :: "nat set" | |
| 781 | show "Inf X \<le> x" if "x \<in> X" "bdd_below X" | |
| 782 | using that by (simp add: Inf_nat_def Least_le) | |
| 783 |   show "x \<le> Inf X" if "X \<noteq> {}" "\<And>y. y \<in> X \<Longrightarrow> x \<le> y"
 | |
| 784 | using that unfolding Inf_nat_def ex_in_conv[symmetric] by (rule LeastI2_ex) | |
| 785 | show "x \<le> Sup X" if "x \<in> X" "bdd_above X" | |
| 71833 
ff6f3b09b8b4
abbrevs for the Uniq quantifier; trying Sup_nat_def to allow zero (experimentally)
 paulson <lp15@cam.ac.uk> parents: 
71238diff
changeset | 786 | using that by (auto simp add: Sup_nat_def bdd_above_nat) | 
| 63540 | 787 |   show "Sup X \<le> x" if "X \<noteq> {}" "\<And>y. y \<in> X \<Longrightarrow> y \<le> x"
 | 
| 788 | proof - | |
| 789 | from that have "bdd_above X" | |
| 54281 | 790 | by (auto simp: bdd_above_def) | 
| 63540 | 791 | with that show ?thesis | 
| 792 | by (simp add: Sup_nat_def bdd_above_nat) | |
| 793 | qed | |
| 54281 | 794 | qed | 
| 63540 | 795 | |
| 54259 
71c701dc5bf9
add SUP and INF for conditionally complete lattices
 hoelzl parents: 
54258diff
changeset | 796 | end | 
| 54281 | 797 | |
| 68610 | 798 | lemma Inf_nat_def1: | 
| 799 | fixes K::"nat set" | |
| 800 |   assumes "K \<noteq> {}"
 | |
| 801 | shows "Inf K \<in> K" | |
| 802 | by (auto simp add: Min_def Inf_nat_def) (meson LeastI assms bot.extremum_unique subsetI) | |
| 803 | ||
| 71834 | 804 | lemma Sup_nat_empty [simp]: "Sup {} = (0::nat)"
 | 
| 805 | by (auto simp add: Sup_nat_def) | |
| 806 | ||
| 807 | ||
| 68610 | 808 | |
| 54281 | 809 | instantiation int :: conditionally_complete_linorder | 
| 810 | begin | |
| 811 | ||
| 812 | definition "Sup (X::int set) = (THE x. x \<in> X \<and> (\<forall>y\<in>X. y \<le> x))" | |
| 813 | definition "Inf (X::int set) = - (Sup (uminus ` X))" | |
| 814 | ||
| 815 | instance | |
| 816 | proof | |
| 817 |   { fix x :: int and X :: "int set" assume "X \<noteq> {}" "bdd_above X"
 | |
| 818 |     then obtain x y where "X \<subseteq> {..y}" "x \<in> X"
 | |
| 819 | by (auto simp: bdd_above_def) | |
| 820 |     then have *: "finite (X \<inter> {x..y})" "X \<inter> {x..y} \<noteq> {}" and "x \<le> y"
 | |
| 821 | by (auto simp: subset_eq) | |
| 822 | have "\<exists>!x\<in>X. (\<forall>y\<in>X. y \<le> x)" | |
| 823 | proof | |
| 824 |       { fix z assume "z \<in> X"
 | |
| 825 |         have "z \<le> Max (X \<inter> {x..y})"
 | |
| 826 | proof cases | |
| 60758 | 827 |           assume "x \<le> z" with \<open>z \<in> X\<close> \<open>X \<subseteq> {..y}\<close> *(1) show ?thesis
 | 
| 54281 | 828 | by (auto intro!: Max_ge) | 
| 829 | next | |
| 830 | assume "\<not> x \<le> z" | |
| 831 | then have "z < x" by simp | |
| 832 |           also have "x \<le> Max (X \<inter> {x..y})"
 | |
| 60758 | 833 | using \<open>x \<in> X\<close> *(1) \<open>x \<le> y\<close> by (intro Max_ge) auto | 
| 54281 | 834 | finally show ?thesis by simp | 
| 835 | qed } | |
| 836 | note le = this | |
| 837 |       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
 | |
| 838 | ||
| 839 | fix z assume *: "z \<in> X \<and> (\<forall>y\<in>X. y \<le> z)" | |
| 840 |       with le have "z \<le> Max (X \<inter> {x..y})"
 | |
| 841 | by auto | |
| 842 |       moreover have "Max (X \<inter> {x..y}) \<le> z"
 | |
| 843 | using * ex by auto | |
| 844 |       ultimately show "z = Max (X \<inter> {x..y})"
 | |
| 845 | by auto | |
| 846 | qed | |
| 847 | then have "Sup X \<in> X \<and> (\<forall>y\<in>X. y \<le> Sup X)" | |
| 848 | unfolding Sup_int_def by (rule theI') } | |
| 849 | note Sup_int = this | |
| 850 | ||
| 851 |   { fix x :: int and X :: "int set" assume "x \<in> X" "bdd_above X" then show "x \<le> Sup X"
 | |
| 852 | using Sup_int[of X] by auto } | |
| 853 | note le_Sup = this | |
| 854 |   { 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"
 | |
| 855 | using Sup_int[of X] by (auto simp: bdd_above_def) } | |
| 856 | note Sup_le = this | |
| 857 | ||
| 858 |   { fix x :: int and X :: "int set" assume "x \<in> X" "bdd_below X" then show "Inf X \<le> x"
 | |
| 859 | using le_Sup[of "-x" "uminus ` X"] by (auto simp: Inf_int_def) } | |
| 860 |   { 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"
 | |
| 861 | using Sup_le[of "uminus ` X" "-x"] by (force simp: Inf_int_def) } | |
| 862 | qed | |
| 863 | end | |
| 864 | ||
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 865 | lemma interval_cases: | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 866 | fixes S :: "'a :: conditionally_complete_linorder set" | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 867 | 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" | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 868 |   shows "\<exists>a b. S = {} \<or>
 | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 869 | S = UNIV \<or> | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 870 |     S = {..<b} \<or>
 | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 871 |     S = {..b} \<or>
 | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 872 |     S = {a<..} \<or>
 | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 873 |     S = {a..} \<or>
 | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 874 |     S = {a<..<b} \<or>
 | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 875 |     S = {a<..b} \<or>
 | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 876 |     S = {a..<b} \<or>
 | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 877 |     S = {a..b}"
 | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 878 | proof - | 
| 63040 | 879 |   define lower upper where "lower = {x. \<exists>s\<in>S. s \<le> x}" and "upper = {x. \<exists>s\<in>S. x \<le> s}"
 | 
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 880 | with ivl have "S = lower \<inter> upper" | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 881 | by auto | 
| 63331 | 882 | moreover | 
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 883 |   have "\<exists>a. upper = UNIV \<or> upper = {} \<or> upper = {.. a} \<or> upper = {..< a}"
 | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 884 | proof cases | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 885 |     assume *: "bdd_above S \<and> S \<noteq> {}"
 | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 886 |     from * have "upper \<subseteq> {.. Sup S}"
 | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 887 | by (auto simp: upper_def intro: cSup_upper2) | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 888 |     moreover from * have "{..< Sup S} \<subseteq> upper"
 | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 889 | by (force simp add: less_cSup_iff upper_def subset_eq Ball_def) | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 890 |     ultimately have "upper = {.. Sup S} \<or> upper = {..< Sup S}"
 | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 891 | unfolding ivl_disj_un(2)[symmetric] by auto | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 892 | then show ?thesis by auto | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 893 | next | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 894 |     assume "\<not> (bdd_above S \<and> S \<noteq> {})"
 | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 895 |     then have "upper = UNIV \<or> upper = {}"
 | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 896 | by (auto simp: upper_def bdd_above_def not_le dest: less_imp_le) | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 897 | then show ?thesis | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 898 | by auto | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 899 | qed | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 900 | moreover | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 901 |   have "\<exists>b. lower = UNIV \<or> lower = {} \<or> lower = {b ..} \<or> lower = {b <..}"
 | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 902 | proof cases | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 903 |     assume *: "bdd_below S \<and> S \<noteq> {}"
 | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 904 |     from * have "lower \<subseteq> {Inf S ..}"
 | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 905 | by (auto simp: lower_def intro: cInf_lower2) | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 906 |     moreover from * have "{Inf S <..} \<subseteq> lower"
 | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 907 | by (force simp add: cInf_less_iff lower_def subset_eq Ball_def) | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 908 |     ultimately have "lower = {Inf S ..} \<or> lower = {Inf S <..}"
 | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 909 | unfolding ivl_disj_un(1)[symmetric] by auto | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 910 | then show ?thesis by auto | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 911 | next | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 912 |     assume "\<not> (bdd_below S \<and> S \<noteq> {})"
 | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 913 |     then have "lower = UNIV \<or> lower = {}"
 | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 914 | by (auto simp: lower_def bdd_below_def not_le dest: less_imp_le) | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 915 | then show ?thesis | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 916 | by auto | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 917 | qed | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 918 | ultimately show ?thesis | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 919 | unfolding greaterThanAtMost_def greaterThanLessThan_def atLeastAtMost_def atLeastLessThan_def | 
| 63171 | 920 | by (metis inf_bot_left inf_bot_right inf_top.left_neutral inf_top.right_neutral) | 
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 921 | qed | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
56218diff
changeset | 922 | |
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60172diff
changeset | 923 | lemma cSUP_eq_cINF_D: | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60172diff
changeset | 924 | fixes f :: "_ \<Rightarrow> 'b::conditionally_complete_lattice" | 
| 67613 | 925 | assumes eq: "(\<Squnion>x\<in>A. f x) = (\<Sqinter>x\<in>A. f x)" | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60172diff
changeset | 926 | and bdd: "bdd_above (f ` A)" "bdd_below (f ` A)" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60172diff
changeset | 927 | and a: "a \<in> A" | 
| 67613 | 928 | shows "f a = (\<Sqinter>x\<in>A. f x)" | 
| 75494 | 929 | proof (rule antisym) | 
| 930 | show "f a \<le> \<Sqinter> (f ` A)" | |
| 931 | by (metis a bdd(1) eq cSUP_upper) | |
| 932 | show "\<Sqinter> (f ` A) \<le> f a" | |
| 933 | using a bdd by (auto simp: cINF_lower) | |
| 934 | qed | |
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60172diff
changeset | 935 | |
| 62379 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62343diff
changeset | 936 | lemma cSUP_UNION: | 
| 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62343diff
changeset | 937 | fixes f :: "_ \<Rightarrow> 'b::conditionally_complete_lattice" | 
| 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62343diff
changeset | 938 |   assumes ne: "A \<noteq> {}" "\<And>x. x \<in> A \<Longrightarrow> B(x) \<noteq> {}"
 | 
| 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62343diff
changeset | 939 | and bdd_UN: "bdd_above (\<Union>x\<in>A. f ` B x)" | 
| 67613 | 940 | shows "(\<Squnion>z \<in> \<Union>x\<in>A. B x. f z) = (\<Squnion>x\<in>A. \<Squnion>z\<in>B x. f z)" | 
| 62379 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62343diff
changeset | 941 | proof - | 
| 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62343diff
changeset | 942 | have bdd: "\<And>x. x \<in> A \<Longrightarrow> bdd_above (f ` B x)" | 
| 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62343diff
changeset | 943 | using bdd_UN by (meson UN_upper bdd_above_mono) | 
| 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62343diff
changeset | 944 | obtain M where "\<And>x y. x \<in> A \<Longrightarrow> y \<in> B(x) \<Longrightarrow> f y \<le> M" | 
| 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62343diff
changeset | 945 | using bdd_UN by (auto simp: bdd_above_def) | 
| 67613 | 946 | then have bdd2: "bdd_above ((\<lambda>x. \<Squnion>z\<in>B x. f z) ` A)" | 
| 62379 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62343diff
changeset | 947 | unfolding bdd_above_def by (force simp: bdd cSUP_le_iff ne(2)) | 
| 67613 | 948 | have "(\<Squnion>z \<in> \<Union>x\<in>A. B x. f z) \<le> (\<Squnion>x\<in>A. \<Squnion>z\<in>B x. f z)" | 
| 62379 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62343diff
changeset | 949 | using assms by (fastforce simp add: intro!: cSUP_least intro: cSUP_upper2 simp: bdd2 bdd) | 
| 67613 | 950 | 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)" | 
| 62379 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62343diff
changeset | 951 | using assms by (fastforce simp add: intro!: cSUP_least intro: cSUP_upper simp: image_UN bdd_UN) | 
| 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62343diff
changeset | 952 | ultimately show ?thesis | 
| 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62343diff
changeset | 953 | by (rule order_antisym) | 
| 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62343diff
changeset | 954 | qed | 
| 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62343diff
changeset | 955 | |
| 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62343diff
changeset | 956 | lemma cINF_UNION: | 
| 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62343diff
changeset | 957 | fixes f :: "_ \<Rightarrow> 'b::conditionally_complete_lattice" | 
| 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62343diff
changeset | 958 |   assumes ne: "A \<noteq> {}" "\<And>x. x \<in> A \<Longrightarrow> B(x) \<noteq> {}"
 | 
| 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62343diff
changeset | 959 | and bdd_UN: "bdd_below (\<Union>x\<in>A. f ` B x)" | 
| 67613 | 960 | shows "(\<Sqinter>z \<in> \<Union>x\<in>A. B x. f z) = (\<Sqinter>x\<in>A. \<Sqinter>z\<in>B x. f z)" | 
| 62379 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62343diff
changeset | 961 | proof - | 
| 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62343diff
changeset | 962 | have bdd: "\<And>x. x \<in> A \<Longrightarrow> bdd_below (f ` B x)" | 
| 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62343diff
changeset | 963 | using bdd_UN by (meson UN_upper bdd_below_mono) | 
| 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62343diff
changeset | 964 | obtain M where "\<And>x y. x \<in> A \<Longrightarrow> y \<in> B(x) \<Longrightarrow> f y \<ge> M" | 
| 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62343diff
changeset | 965 | using bdd_UN by (auto simp: bdd_below_def) | 
| 67613 | 966 | then have bdd2: "bdd_below ((\<lambda>x. \<Sqinter>z\<in>B x. f z) ` A)" | 
| 62379 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62343diff
changeset | 967 | unfolding bdd_below_def by (force simp: bdd le_cINF_iff ne(2)) | 
| 67613 | 968 | have "(\<Sqinter>z \<in> \<Union>x\<in>A. B x. f z) \<le> (\<Sqinter>x\<in>A. \<Sqinter>z\<in>B x. f z)" | 
| 62379 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62343diff
changeset | 969 | using assms by (fastforce simp add: intro!: cINF_greatest intro: cINF_lower simp: bdd2 bdd) | 
| 67613 | 970 | 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)" | 
| 62379 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62343diff
changeset | 971 | using assms by (fastforce simp add: intro!: cINF_greatest intro: cINF_lower2 simp: bdd bdd_UN bdd2) | 
| 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62343diff
changeset | 972 | ultimately show ?thesis | 
| 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62343diff
changeset | 973 | by (rule order_antisym) | 
| 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62343diff
changeset | 974 | qed | 
| 
340738057c8c
An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule!
 paulson <lp15@cam.ac.uk> parents: 
62343diff
changeset | 975 | |
| 63331 | 976 | lemma cSup_abs_le: | 
| 62626 
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
 paulson <lp15@cam.ac.uk> parents: 
62379diff
changeset | 977 |   fixes S :: "('a::{linordered_idom,conditionally_complete_linorder}) set"
 | 
| 
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
 paulson <lp15@cam.ac.uk> parents: 
62379diff
changeset | 978 |   shows "S \<noteq> {} \<Longrightarrow> (\<And>x. x\<in>S \<Longrightarrow> \<bar>x\<bar> \<le> a) \<Longrightarrow> \<bar>Sup S\<bar> \<le> a"
 | 
| 
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
 paulson <lp15@cam.ac.uk> parents: 
62379diff
changeset | 979 | apply (auto simp add: abs_le_iff intro: cSup_least) | 
| 
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
 paulson <lp15@cam.ac.uk> parents: 
62379diff
changeset | 980 | by (metis bdd_aboveI cSup_upper neg_le_iff_le order_trans) | 
| 
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
 paulson <lp15@cam.ac.uk> parents: 
62379diff
changeset | 981 | |
| 54281 | 982 | end |