| author | wenzelm | 
| Tue, 06 Jun 2023 11:07:49 +0200 | |
| changeset 78134 | a11ebc8c751a | 
| parent 69597 | ff784d5a5bfb | 
| permissions | -rw-r--r-- | 
| 58810 | 1 | (* Title: HOL/Library/Lub_Glb.thy | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 2 | Author: Jacques D. Fleuriot, University of Cambridge | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 3 | Author: Amine Chaieb, University of Cambridge *) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 4 | |
| 60500 | 5 | section \<open>Definitions of Least Upper Bounds and Greatest Lower Bounds\<close> | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 6 | |
| 58810 | 7 | theory Lub_Glb | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 8 | imports Complex_Main | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 9 | begin | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 10 | |
| 60500 | 11 | text \<open>Thanks to suggestions by James Margetson\<close> | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 12 | |
| 69597 | 13 | definition setle :: "'a set \<Rightarrow> 'a::ord \<Rightarrow> bool" (infixl \<open>*<=\<close> 70) | 
| 67613 | 14 | where "S *<= x = (\<forall>y\<in>S. y \<le> x)" | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 15 | |
| 69597 | 16 | definition setge :: "'a::ord \<Rightarrow> 'a set \<Rightarrow> bool" (infixl \<open><=*\<close> 70) | 
| 67613 | 17 | where "x <=* S = (\<forall>y\<in>S. x \<le> y)" | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 18 | |
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 19 | |
| 61585 | 20 | subsection \<open>Rules for the Relations \<open>*<=\<close> and \<open><=*\<close>\<close> | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 21 | |
| 67613 | 22 | lemma setleI: "\<forall>y\<in>S. y \<le> x \<Longrightarrow> S *<= x" | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 23 | by (simp add: setle_def) | 
| 29838 | 24 | |
| 67613 | 25 | lemma setleD: "S *<= x \<Longrightarrow> y\<in>S \<Longrightarrow> y \<le> x" | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 26 | by (simp add: setle_def) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 27 | |
| 67613 | 28 | lemma setgeI: "\<forall>y\<in>S. x \<le> y \<Longrightarrow> x <=* S" | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 29 | by (simp add: setge_def) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 30 | |
| 67613 | 31 | lemma setgeD: "x <=* S \<Longrightarrow> y\<in>S \<Longrightarrow> x \<le> y" | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 32 | by (simp add: setge_def) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 33 | |
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 34 | |
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 35 | definition leastP :: "('a \<Rightarrow> bool) \<Rightarrow> 'a::ord \<Rightarrow> bool"
 | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 36 | where "leastP P x = (P x \<and> x <=* Collect P)" | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 37 | |
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 38 | definition isUb :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a::ord \<Rightarrow> bool" | 
| 67613 | 39 | where "isUb R S x = (S *<= x \<and> x \<in> R)" | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 40 | |
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 41 | definition isLub :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a::ord \<Rightarrow> bool" | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 42 | where "isLub R S x = leastP (isUb R S) x" | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 43 | |
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 44 | definition ubs :: "'a set \<Rightarrow> 'a::ord set \<Rightarrow> 'a set" | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 45 | where "ubs R S = Collect (isUb R S)" | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 46 | |
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 47 | |
| 69593 | 48 | subsection \<open>Rules about the Operators \<^term>\<open>leastP\<close>, \<^term>\<open>ub\<close> and \<^term>\<open>lub\<close>\<close> | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 49 | |
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 50 | lemma leastPD1: "leastP P x \<Longrightarrow> P x" | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 51 | by (simp add: leastP_def) | 
| 29838 | 52 | |
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 53 | lemma leastPD2: "leastP P x \<Longrightarrow> x <=* Collect P" | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 54 | by (simp add: leastP_def) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 55 | |
| 67613 | 56 | lemma leastPD3: "leastP P x \<Longrightarrow> y \<in> Collect P \<Longrightarrow> x \<le> y" | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 57 | by (blast dest!: leastPD2 setgeD) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 58 | |
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 59 | lemma isLubD1: "isLub R S x \<Longrightarrow> S *<= x" | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 60 | by (simp add: isLub_def isUb_def leastP_def) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 61 | |
| 67613 | 62 | lemma isLubD1a: "isLub R S x \<Longrightarrow> x \<in> R" | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 63 | by (simp add: isLub_def isUb_def leastP_def) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 64 | |
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 65 | lemma isLub_isUb: "isLub R S x \<Longrightarrow> isUb R S x" | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 66 | unfolding isUb_def by (blast dest: isLubD1 isLubD1a) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 67 | |
| 67613 | 68 | lemma isLubD2: "isLub R S x \<Longrightarrow> y \<in> S \<Longrightarrow> y \<le> x" | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 69 | by (blast dest!: isLubD1 setleD) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 70 | |
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 71 | lemma isLubD3: "isLub R S x \<Longrightarrow> leastP (isUb R S) x" | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 72 | by (simp add: isLub_def) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 73 | |
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 74 | lemma isLubI1: "leastP(isUb R S) x \<Longrightarrow> isLub R S x" | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 75 | by (simp add: isLub_def) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 76 | |
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 77 | lemma isLubI2: "isUb R S x \<Longrightarrow> x <=* Collect (isUb R S) \<Longrightarrow> isLub R S x" | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 78 | by (simp add: isLub_def leastP_def) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 79 | |
| 67613 | 80 | lemma isUbD: "isUb R S x \<Longrightarrow> y \<in> S \<Longrightarrow> y \<le> x" | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 81 | by (simp add: isUb_def setle_def) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 82 | |
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 83 | lemma isUbD2: "isUb R S x \<Longrightarrow> S *<= x" | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 84 | by (simp add: isUb_def) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 85 | |
| 67613 | 86 | lemma isUbD2a: "isUb R S x \<Longrightarrow> x \<in> R" | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 87 | by (simp add: isUb_def) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 88 | |
| 67613 | 89 | lemma isUbI: "S *<= x \<Longrightarrow> x \<in> R \<Longrightarrow> isUb R S x" | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 90 | by (simp add: isUb_def) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 91 | |
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 92 | lemma isLub_le_isUb: "isLub R S x \<Longrightarrow> isUb R S y \<Longrightarrow> x \<le> y" | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 93 | unfolding isLub_def by (blast intro!: leastPD3) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 94 | |
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 95 | lemma isLub_ubs: "isLub R S x \<Longrightarrow> x <=* ubs R S" | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 96 | unfolding ubs_def isLub_def by (rule leastPD2) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 97 | |
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 98 | lemma isLub_unique: "[| isLub R S x; isLub R S y |] ==> x = (y::'a::linorder)" | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 99 | apply (frule isLub_isUb) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 100 | apply (frule_tac x = y in isLub_isUb) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 101 | apply (blast intro!: order_antisym dest!: isLub_le_isUb) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 102 | done | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 103 | |
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 104 | lemma isUb_UNIV_I: "(\<And>y. y \<in> S \<Longrightarrow> y \<le> u) \<Longrightarrow> isUb UNIV S u" | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 105 | by (simp add: isUbI setleI) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 106 | |
| 29838 | 107 | |
| 46509 | 108 | definition greatestP :: "('a \<Rightarrow> bool) \<Rightarrow> 'a::ord \<Rightarrow> bool"
 | 
| 109 | where "greatestP P x = (P x \<and> Collect P *<= x)" | |
| 29838 | 110 | |
| 46509 | 111 | definition isLb :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a::ord \<Rightarrow> bool" | 
| 67613 | 112 | where "isLb R S x = (x <=* S \<and> x \<in> R)" | 
| 29838 | 113 | |
| 46509 | 114 | definition isGlb :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a::ord \<Rightarrow> bool" | 
| 115 | where "isGlb R S x = greatestP (isLb R S) x" | |
| 29838 | 116 | |
| 46509 | 117 | definition lbs :: "'a set \<Rightarrow> 'a::ord set \<Rightarrow> 'a set" | 
| 118 | where "lbs R S = Collect (isLb R S)" | |
| 119 | ||
| 29838 | 120 | |
| 69593 | 121 | subsection \<open>Rules about the Operators \<^term>\<open>greatestP\<close>, \<^term>\<open>isLb\<close> and \<^term>\<open>isGlb\<close>\<close> | 
| 29838 | 122 | |
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 123 | lemma greatestPD1: "greatestP P x \<Longrightarrow> P x" | 
| 46509 | 124 | by (simp add: greatestP_def) | 
| 29838 | 125 | |
| 46509 | 126 | lemma greatestPD2: "greatestP P x \<Longrightarrow> Collect P *<= x" | 
| 127 | by (simp add: greatestP_def) | |
| 29838 | 128 | |
| 67613 | 129 | lemma greatestPD3: "greatestP P x \<Longrightarrow> y \<in> Collect P \<Longrightarrow> x \<ge> y" | 
| 46509 | 130 | by (blast dest!: greatestPD2 setleD) | 
| 29838 | 131 | |
| 46509 | 132 | lemma isGlbD1: "isGlb R S x \<Longrightarrow> x <=* S" | 
| 133 | by (simp add: isGlb_def isLb_def greatestP_def) | |
| 29838 | 134 | |
| 67613 | 135 | lemma isGlbD1a: "isGlb R S x \<Longrightarrow> x \<in> R" | 
| 46509 | 136 | by (simp add: isGlb_def isLb_def greatestP_def) | 
| 29838 | 137 | |
| 46509 | 138 | lemma isGlb_isLb: "isGlb R S x \<Longrightarrow> isLb R S x" | 
| 139 | unfolding isLb_def by (blast dest: isGlbD1 isGlbD1a) | |
| 29838 | 140 | |
| 67613 | 141 | lemma isGlbD2: "isGlb R S x \<Longrightarrow> y \<in> S \<Longrightarrow> y \<ge> x" | 
| 46509 | 142 | by (blast dest!: isGlbD1 setgeD) | 
| 29838 | 143 | |
| 46509 | 144 | lemma isGlbD3: "isGlb R S x \<Longrightarrow> greatestP (isLb R S) x" | 
| 145 | by (simp add: isGlb_def) | |
| 29838 | 146 | |
| 46509 | 147 | lemma isGlbI1: "greatestP (isLb R S) x \<Longrightarrow> isGlb R S x" | 
| 148 | by (simp add: isGlb_def) | |
| 29838 | 149 | |
| 46509 | 150 | lemma isGlbI2: "isLb R S x \<Longrightarrow> Collect (isLb R S) *<= x \<Longrightarrow> isGlb R S x" | 
| 151 | by (simp add: isGlb_def greatestP_def) | |
| 29838 | 152 | |
| 67613 | 153 | lemma isLbD: "isLb R S x \<Longrightarrow> y \<in> S \<Longrightarrow> y \<ge> x" | 
| 46509 | 154 | by (simp add: isLb_def setge_def) | 
| 29838 | 155 | |
| 46509 | 156 | lemma isLbD2: "isLb R S x \<Longrightarrow> x <=* S " | 
| 157 | by (simp add: isLb_def) | |
| 29838 | 158 | |
| 67613 | 159 | lemma isLbD2a: "isLb R S x \<Longrightarrow> x \<in> R" | 
| 46509 | 160 | by (simp add: isLb_def) | 
| 29838 | 161 | |
| 67613 | 162 | lemma isLbI: "x <=* S \<Longrightarrow> x \<in> R \<Longrightarrow> isLb R S x" | 
| 46509 | 163 | by (simp add: isLb_def) | 
| 29838 | 164 | |
| 46509 | 165 | lemma isGlb_le_isLb: "isGlb R S x \<Longrightarrow> isLb R S y \<Longrightarrow> x \<ge> y" | 
| 166 | unfolding isGlb_def by (blast intro!: greatestPD3) | |
| 29838 | 167 | |
| 46509 | 168 | lemma isGlb_ubs: "isGlb R S x \<Longrightarrow> lbs R S *<= x" | 
| 169 | unfolding lbs_def isGlb_def by (rule greatestPD2) | |
| 29838 | 170 | |
| 51342 | 171 | lemma isGlb_unique: "[| isGlb R S x; isGlb R S y |] ==> x = (y::'a::linorder)" | 
| 172 | apply (frule isGlb_isLb) | |
| 173 | apply (frule_tac x = y in isGlb_isLb) | |
| 174 | apply (blast intro!: order_antisym dest!: isGlb_le_isLb) | |
| 175 | done | |
| 176 | ||
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 177 | lemma bdd_above_setle: "bdd_above A \<longleftrightarrow> (\<exists>a. A *<= a)" | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 178 | by (auto simp: bdd_above_def setle_def) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 179 | |
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 180 | lemma bdd_below_setge: "bdd_below A \<longleftrightarrow> (\<exists>a. a <=* A)" | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 181 | by (auto simp: bdd_below_def setge_def) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 182 | |
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 183 | lemma isLub_cSup: | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 184 |   "(S::'a :: conditionally_complete_lattice set) \<noteq> {} \<Longrightarrow> (\<exists>b. S *<= b) \<Longrightarrow> isLub UNIV S (Sup S)"
 | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 185 | by (auto simp add: isLub_def setle_def leastP_def isUb_def | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 186 | intro!: setgeI cSup_upper cSup_least) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 187 | |
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 188 | lemma isGlb_cInf: | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 189 |   "(S::'a :: conditionally_complete_lattice set) \<noteq> {} \<Longrightarrow> (\<exists>b. b <=* S) \<Longrightarrow> isGlb UNIV S (Inf S)"
 | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 190 | by (auto simp add: isGlb_def setge_def greatestP_def isLb_def | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 191 | intro!: setleI cInf_lower cInf_greatest) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 192 | |
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 193 | lemma cSup_le: "(S::'a::conditionally_complete_lattice set) \<noteq> {} \<Longrightarrow> S *<= b \<Longrightarrow> Sup S \<le> b"
 | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 194 | by (metis cSup_least setle_def) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 195 | |
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 196 | lemma cInf_ge: "(S::'a :: conditionally_complete_lattice set) \<noteq> {} \<Longrightarrow> b <=* S \<Longrightarrow> Inf S \<ge> b"
 | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 197 | by (metis cInf_greatest setge_def) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 198 | |
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 199 | lemma cSup_bounds: | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 200 | fixes S :: "'a :: conditionally_complete_lattice set" | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 201 |   shows "S \<noteq> {} \<Longrightarrow> a <=* S \<Longrightarrow> S *<= b \<Longrightarrow> a \<le> Sup S \<and> Sup S \<le> b"
 | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 202 | using cSup_least[of S b] cSup_upper2[of _ S a] | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 203 | by (auto simp: bdd_above_setle setge_def setle_def) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 204 | |
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 205 | lemma cSup_unique: "(S::'a :: {conditionally_complete_linorder, no_bot} set) *<= b \<Longrightarrow> (\<forall>b'<b. \<exists>x\<in>S. b' < x) \<Longrightarrow> Sup S = b"
 | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 206 | by (rule cSup_eq) (auto simp: not_le[symmetric] setle_def) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 207 | |
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 208 | lemma cInf_unique: "b <=* (S::'a :: {conditionally_complete_linorder, no_top} set) \<Longrightarrow> (\<forall>b'>b. \<exists>x\<in>S. b' > x) \<Longrightarrow> Inf S = b"
 | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 209 | by (rule cInf_eq) (auto simp: not_le[symmetric] setge_def) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 210 | |
| 60500 | 211 | text\<open>Use completeness of reals (supremum property) to show that any bounded sequence has a least upper bound\<close> | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 212 | |
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 213 | lemma reals_complete: "\<exists>X. X \<in> S \<Longrightarrow> \<exists>Y. isUb (UNIV::real set) S Y \<Longrightarrow> \<exists>t. isLub (UNIV :: real set) S t" | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 214 | by (intro exI[of _ "Sup S"] isLub_cSup) (auto simp: setle_def isUb_def intro!: cSup_upper) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 215 | |
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 216 | lemma Bseq_isUb: "\<And>X :: nat \<Rightarrow> real. Bseq X \<Longrightarrow> \<exists>U. isUb (UNIV::real set) {x. \<exists>n. X n = x} U"
 | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 217 | by (auto intro: isUbI setleI simp add: Bseq_def abs_le_iff) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 218 | |
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 219 | lemma Bseq_isLub: "\<And>X :: nat \<Rightarrow> real. Bseq X \<Longrightarrow> \<exists>U. isLub (UNIV::real set) {x. \<exists>n. X n = x} U"
 | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 220 | by (blast intro: reals_complete Bseq_isUb) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 221 | |
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 222 | lemma isLub_mono_imp_LIMSEQ: | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 223 | fixes X :: "nat \<Rightarrow> real" | 
| 68356 | 224 |   assumes u: "isLub UNIV {x. \<exists>n. X n = x} u" (* FIXME: use \<open>range X\<close> *)
 | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 225 | assumes X: "\<forall>m n. m \<le> n \<longrightarrow> X m \<le> X n" | 
| 61969 | 226 | shows "X \<longlonglongrightarrow> u" | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 227 | proof - | 
| 61969 | 228 | have "X \<longlonglongrightarrow> (SUP i. X i)" | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 229 | using u[THEN isLubD1] X | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 230 | by (intro LIMSEQ_incseq_SUP) (auto simp: incseq_def image_def eq_commute bdd_above_setle) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 231 | also have "(SUP i. X i) = u" | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 232 | using isLub_cSup[of "range X"] u[THEN isLubD1] | 
| 62343 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61969diff
changeset | 233 | by (intro isLub_unique[OF _ u]) (auto simp add: image_def eq_commute) | 
| 54263 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 234 | finally show ?thesis . | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 235 | qed | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 236 | |
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 237 | lemmas real_isGlb_unique = isGlb_unique[where 'a=real] | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 238 | |
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 239 | lemma real_le_inf_subset: "t \<noteq> {} \<Longrightarrow> t \<subseteq> s \<Longrightarrow> \<exists>b. b <=* s \<Longrightarrow> Inf s \<le> Inf (t::real set)"
 | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 240 | by (rule cInf_superset_mono) (auto simp: bdd_below_setge) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 241 | |
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 242 | lemma real_ge_sup_subset: "t \<noteq> {} \<Longrightarrow> t \<subseteq> s \<Longrightarrow> \<exists>b. s *<= b \<Longrightarrow> Sup s \<ge> Sup (t::real set)"
 | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 243 | by (rule cSup_subset_mono) (auto simp: bdd_above_setle) | 
| 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 hoelzl parents: 
51342diff
changeset | 244 | |
| 29838 | 245 | end |