| author | wenzelm | 
| Fri, 14 Feb 2025 15:23:00 +0100 | |
| changeset 82161 | 843c2205b076 | 
| parent 80914 | d97fdabd9e2b | 
| child 82248 | e8c96013ea8a | 
| permissions | -rw-r--r-- | 
| 13383 | 1 | (* Title: HOL/ex/Tarski.thy | 
| 40945 | 2 | Author: Florian Kammüller, Cambridge University Computer Laboratory | 
| 13383 | 3 | *) | 
| 7112 | 4 | |
| 61343 | 5 | section \<open>The Full Theorem of Tarski\<close> | 
| 7112 | 6 | |
| 27681 | 7 | theory Tarski | 
| 68188 
2af1f142f855
move FuncSet back to HOL-Library (amending 493b818e8e10)
 immler parents: 
68072diff
changeset | 8 | imports Main "HOL-Library.FuncSet" | 
| 27681 | 9 | begin | 
| 7112 | 10 | |
| 61343 | 11 | text \<open> | 
| 13383 | 12 | Minimal version of lattice theory plus the full theorem of Tarski: | 
| 13 | The fixedpoints of a complete lattice themselves form a complete | |
| 14 | lattice. | |
| 15 | ||
| 16 | Illustrates first-class theories, using the Sigma representation of | |
| 17 | structures. Tidied and converted to Isar by lcp. | |
| 61343 | 18 | \<close> | 
| 13383 | 19 | |
| 20 | record 'a potype = | |
| 7112 | 21 | pset :: "'a set" | 
| 64915 | 22 |   order :: "('a \<times> 'a) set"
 | 
| 7112 | 23 | |
| 64915 | 24 | definition monotone :: "['a \<Rightarrow> 'a, 'a set, ('a \<times> 'a) set] \<Rightarrow> bool"
 | 
| 25 | where "monotone f A r \<longleftrightarrow> (\<forall>x\<in>A. \<forall>y\<in>A. (x, y) \<in> r \<longrightarrow> (f x, f y) \<in> r)" | |
| 7112 | 26 | |
| 64915 | 27 | definition least :: "['a \<Rightarrow> bool, 'a potype] \<Rightarrow> 'a" | 
| 28 | where "least P po = (SOME x. x \<in> pset po \<and> P x \<and> (\<forall>y \<in> pset po. P y \<longrightarrow> (x, y) \<in> order po))" | |
| 7112 | 29 | |
| 64915 | 30 | definition greatest :: "['a \<Rightarrow> bool, 'a potype] \<Rightarrow> 'a" | 
| 31 | where "greatest P po = (SOME x. x \<in> pset po \<and> P x \<and> (\<forall>y \<in> pset po. P y \<longrightarrow> (y, x) \<in> order po))" | |
| 7112 | 32 | |
| 64915 | 33 | definition lub :: "['a set, 'a potype] \<Rightarrow> 'a" | 
| 34 | where "lub S po = least (\<lambda>x. \<forall>y\<in>S. (y, x) \<in> order po) po" | |
| 7112 | 35 | |
| 64915 | 36 | definition glb :: "['a set, 'a potype] \<Rightarrow> 'a" | 
| 37 | where "glb S po = greatest (\<lambda>x. \<forall>y\<in>S. (x, y) \<in> order po) po" | |
| 7112 | 38 | |
| 64915 | 39 | definition isLub :: "['a set, 'a potype, 'a] \<Rightarrow> bool" | 
| 40 | where "isLub S po = | |
| 41 | (\<lambda>L. L \<in> pset po \<and> (\<forall>y\<in>S. (y, L) \<in> order po) \<and> | |
| 42 | (\<forall>z\<in>pset po. (\<forall>y\<in>S. (y, z) \<in> order po) \<longrightarrow> (L, z) \<in> order po))" | |
| 7112 | 43 | |
| 64915 | 44 | definition isGlb :: "['a set, 'a potype, 'a] \<Rightarrow> bool" | 
| 45 | where "isGlb S po = | |
| 46 | (\<lambda>G. (G \<in> pset po \<and> (\<forall>y\<in>S. (G, y) \<in> order po) \<and> | |
| 47 | (\<forall>z \<in> pset po. (\<forall>y\<in>S. (z, y) \<in> order po) \<longrightarrow> (z, G) \<in> order po)))" | |
| 7112 | 48 | |
| 64915 | 49 | definition "fix" :: "['a \<Rightarrow> 'a, 'a set] \<Rightarrow> 'a set" | 
| 50 |   where "fix f A  = {x. x \<in> A \<and> f x = x}"
 | |
| 7112 | 51 | |
| 64915 | 52 | definition interval :: "[('a \<times> 'a) set, 'a, 'a] \<Rightarrow> 'a set"
 | 
| 53 |   where "interval r a b = {x. (a, x) \<in> r \<and> (x, b) \<in> r}"
 | |
| 7112 | 54 | |
| 64915 | 55 | definition Bot :: "'a potype \<Rightarrow> 'a" | 
| 56 | where "Bot po = least (\<lambda>x. True) po" | |
| 7112 | 57 | |
| 64915 | 58 | definition Top :: "'a potype \<Rightarrow> 'a" | 
| 59 | where "Top po = greatest (\<lambda>x. True) po" | |
| 7112 | 60 | |
| 64915 | 61 | definition PartialOrder :: "'a potype set" | 
| 62 |   where "PartialOrder = {P. refl_on (pset P) (order P) \<and> antisym (order P) \<and> trans (order P)}"
 | |
| 7112 | 63 | |
| 64915 | 64 | definition CompleteLattice :: "'a potype set" | 
| 65 | where "CompleteLattice = | |
| 66 |     {cl. cl \<in> PartialOrder \<and>
 | |
| 67 | (\<forall>S. S \<subseteq> pset cl \<longrightarrow> (\<exists>L. isLub S cl L)) \<and> | |
| 68 | (\<forall>S. S \<subseteq> pset cl \<longrightarrow> (\<exists>G. isGlb S cl G))}" | |
| 7112 | 69 | |
| 64915 | 70 | definition CLF_set :: "('a potype \<times> ('a \<Rightarrow> 'a)) set"
 | 
| 71 | where "CLF_set = | |
| 72 | (SIGMA cl : CompleteLattice. | |
| 73 |       {f. f \<in> pset cl \<rightarrow> pset cl \<and> monotone f (pset cl) (order cl)})"
 | |
| 13383 | 74 | |
| 64915 | 75 | definition induced :: "['a set, ('a \<times> 'a) set] \<Rightarrow> ('a \<times> 'a) set"
 | 
| 76 |   where "induced A r = {(a, b). a \<in> A \<and> b \<in> A \<and> (a, b) \<in> r}"
 | |
| 7112 | 77 | |
| 64915 | 78 | definition sublattice :: "('a potype \<times> 'a set) set"
 | 
| 79 | where "sublattice = | |
| 80 | (SIGMA cl : CompleteLattice. | |
| 81 |       {S. S \<subseteq> pset cl \<and> \<lparr>pset = S, order = induced S (order cl)\<rparr> \<in> CompleteLattice})"
 | |
| 7112 | 82 | |
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
70202diff
changeset | 83 | abbreviation sublat :: "['a set, 'a potype] \<Rightarrow> bool" (\<open>_ <<= _\<close> [51, 50] 50) | 
| 64915 | 84 |   where "S <<= cl \<equiv> S \<in> sublattice `` {cl}"
 | 
| 7112 | 85 | |
| 64915 | 86 | definition dual :: "'a potype \<Rightarrow> 'a potype" | 
| 87 | where "dual po = \<lparr>pset = pset po, order = converse (order po)\<rparr>" | |
| 7112 | 88 | |
| 27681 | 89 | locale S = | 
| 13115 | 90 | fixes cl :: "'a potype" | 
| 64915 | 91 | and A :: "'a set" | 
| 92 |     and r :: "('a \<times> 'a) set"
 | |
| 93 | defines A_def: "A \<equiv> pset cl" | |
| 94 | and r_def: "r \<equiv> order cl" | |
| 7112 | 95 | |
| 27681 | 96 | locale PO = S + | 
| 64915 | 97 | assumes cl_po: "cl \<in> PartialOrder" | 
| 27681 | 98 | |
| 99 | locale CL = S + | |
| 64915 | 100 | assumes cl_co: "cl \<in> CompleteLattice" | 
| 7112 | 101 | |
| 61565 
352c73a689da
Qualifiers in locale expressions default to mandatory regardless of the command.
 ballarin parents: 
61384diff
changeset | 102 | sublocale CL < po?: PO | 
| 70194 | 103 | unfolding A_def r_def | 
| 104 | using CompleteLattice_def PO.intro cl_co by fastforce | |
| 27681 | 105 | |
| 106 | locale CLF = S + | |
| 64915 | 107 | fixes f :: "'a \<Rightarrow> 'a" | 
| 13115 | 108 | and P :: "'a set" | 
| 70194 | 109 | assumes f_cl: "(cl, f) \<in> CLF_set" | 
| 64915 | 110 | defines P_def: "P \<equiv> fix f A" | 
| 7112 | 111 | |
| 61565 
352c73a689da
Qualifiers in locale expressions default to mandatory regardless of the command.
 ballarin parents: 
61384diff
changeset | 112 | sublocale CLF < cl?: CL | 
| 70194 | 113 | unfolding A_def r_def CL_def | 
| 114 | using CLF_set_def f_cl by blast | |
| 7112 | 115 | |
| 27681 | 116 | locale Tarski = CLF + | 
| 64915 | 117 | fixes Y :: "'a set" | 
| 13115 | 118 | and intY1 :: "'a set" | 
| 64915 | 119 | and v :: "'a" | 
| 120 | assumes Y_ss: "Y \<subseteq> P" | |
| 121 | defines intY1_def: "intY1 \<equiv> interval r (lub Y cl) (Top cl)" | |
| 122 | and v_def: "v \<equiv> | |
| 123 |       glb {x. ((\<lambda>x \<in> intY1. f x) x, x) \<in> induced intY1 r \<and> x \<in> intY1}
 | |
| 124 | \<lparr>pset = intY1, order = induced intY1 r\<rparr>" | |
| 13115 | 125 | |
| 126 | ||
| 61343 | 127 | subsection \<open>Partial Order\<close> | 
| 13115 | 128 | |
| 64916 | 129 | context PO | 
| 130 | begin | |
| 131 | ||
| 132 | lemma dual: "PO (dual cl)" | |
| 70194 | 133 | proof | 
| 134 | show "dual cl \<in> PartialOrder" | |
| 135 | using cl_po unfolding PartialOrder_def dual_def by auto | |
| 136 | qed | |
| 27681 | 137 | |
| 64916 | 138 | lemma PO_imp_refl_on [simp]: "refl_on A r" | 
| 64915 | 139 | using cl_po by (simp add: PartialOrder_def A_def r_def) | 
| 13115 | 140 | |
| 64916 | 141 | lemma PO_imp_sym [simp]: "antisym r" | 
| 64915 | 142 | using cl_po by (simp add: PartialOrder_def r_def) | 
| 13115 | 143 | |
| 64916 | 144 | lemma PO_imp_trans [simp]: "trans r" | 
| 64915 | 145 | using cl_po by (simp add: PartialOrder_def r_def) | 
| 13115 | 146 | |
| 64916 | 147 | lemma reflE: "x \<in> A \<Longrightarrow> (x, x) \<in> r" | 
| 64915 | 148 | using cl_po by (simp add: PartialOrder_def refl_on_def A_def r_def) | 
| 13115 | 149 | |
| 64916 | 150 | lemma antisymE: "\<lbrakk>(a, b) \<in> r; (b, a) \<in> r\<rbrakk> \<Longrightarrow> a = b" | 
| 64915 | 151 | using cl_po by (simp add: PartialOrder_def antisym_def r_def) | 
| 13115 | 152 | |
| 64916 | 153 | lemma transE: "\<lbrakk>(a, b) \<in> r; (b, c) \<in> r\<rbrakk> \<Longrightarrow> (a, c) \<in> r" | 
| 64915 | 154 | using cl_po by (simp add: PartialOrder_def r_def) (unfold trans_def, fast) | 
| 13115 | 155 | |
| 64916 | 156 | lemma monotoneE: "\<lbrakk>monotone f A r; x \<in> A; y \<in> A; (x, y) \<in> r\<rbrakk> \<Longrightarrow> (f x, f y) \<in> r" | 
| 64915 | 157 | by (simp add: monotone_def) | 
| 13115 | 158 | |
| 70194 | 159 | lemma po_subset_po: | 
| 160 | assumes "S \<subseteq> A" shows "\<lparr>pset = S, order = induced S r\<rparr> \<in> PartialOrder" | |
| 161 | proof - | |
| 162 | have "refl_on S (induced S r)" | |
| 163 | using \<open>S \<subseteq> A\<close> by (auto simp: refl_on_def induced_def intro: reflE) | |
| 164 | moreover | |
| 165 | have "antisym (induced S r)" | |
| 166 | by (auto simp add: antisym_def induced_def intro: antisymE) | |
| 167 | moreover | |
| 168 | have "trans (induced S r)" | |
| 169 | by (auto simp add: trans_def induced_def intro: transE) | |
| 170 | ultimately show ?thesis | |
| 171 | by (simp add: PartialOrder_def) | |
| 172 | qed | |
| 13115 | 173 | |
| 64916 | 174 | lemma indE: "\<lbrakk>(x, y) \<in> induced S r; S \<subseteq> A\<rbrakk> \<Longrightarrow> (x, y) \<in> r" | 
| 64915 | 175 | by (simp add: induced_def) | 
| 13115 | 176 | |
| 64916 | 177 | lemma indI: "\<lbrakk>(x, y) \<in> r; x \<in> S; y \<in> S\<rbrakk> \<Longrightarrow> (x, y) \<in> induced S r" | 
| 64915 | 178 | by (simp add: induced_def) | 
| 13115 | 179 | |
| 64916 | 180 | end | 
| 181 | ||
| 64915 | 182 | lemma (in CL) CL_imp_ex_isLub: "S \<subseteq> A \<Longrightarrow> \<exists>L. isLub S cl L" | 
| 183 | using cl_co by (simp add: CompleteLattice_def A_def) | |
| 13115 | 184 | |
| 185 | declare (in CL) cl_co [simp] | |
| 186 | ||
| 64915 | 187 | lemma isLub_lub: "(\<exists>L. isLub S cl L) \<longleftrightarrow> isLub S cl (lub S cl)" | 
| 188 | by (simp add: lub_def least_def isLub_def some_eq_ex [symmetric]) | |
| 13115 | 189 | |
| 64915 | 190 | lemma isGlb_glb: "(\<exists>G. isGlb S cl G) \<longleftrightarrow> isGlb S cl (glb S cl)" | 
| 191 | by (simp add: glb_def greatest_def isGlb_def some_eq_ex [symmetric]) | |
| 13115 | 192 | |
| 193 | lemma isGlb_dual_isLub: "isGlb S cl = isLub S (dual cl)" | |
| 64915 | 194 | by (simp add: isLub_def isGlb_def dual_def converse_unfold) | 
| 13115 | 195 | |
| 196 | lemma isLub_dual_isGlb: "isLub S cl = isGlb S (dual cl)" | |
| 64915 | 197 | by (simp add: isLub_def isGlb_def dual_def converse_unfold) | 
| 13115 | 198 | |
| 199 | lemma (in PO) dualPO: "dual cl \<in> PartialOrder" | |
| 64915 | 200 | using cl_po by (simp add: PartialOrder_def dual_def) | 
| 13115 | 201 | |
| 202 | lemma Rdual: | |
| 70194 | 203 | assumes major: "\<And>S. S \<subseteq> A \<Longrightarrow> \<exists>L. isLub S po L" and "S \<subseteq> A" and "A = pset po" | 
| 204 | shows "\<exists>G. isGlb S po G" | |
| 205 | proof | |
| 206 |   show "isGlb S po (lub {y \<in> A. \<forall>k\<in>S. (y, k) \<in> order po} po)"
 | |
| 207 |     using major [of "{y. y \<in> A \<and> (\<forall>k \<in> S. (y, k) \<in> order po)}"] \<open>S \<subseteq> A\<close> \<open>A = pset po\<close>
 | |
| 208 | apply (simp add: isLub_lub isGlb_def) | |
| 209 | apply (auto simp add: isLub_def) | |
| 210 | done | |
| 211 | qed | |
| 13115 | 212 | |
| 213 | lemma lub_dual_glb: "lub S cl = glb S (dual cl)" | |
| 64915 | 214 | by (simp add: lub_def glb_def least_def greatest_def dual_def converse_unfold) | 
| 13115 | 215 | |
| 216 | lemma glb_dual_lub: "glb S cl = lub S (dual cl)" | |
| 64915 | 217 | by (simp add: lub_def glb_def least_def greatest_def dual_def converse_unfold) | 
| 13115 | 218 | |
| 17841 | 219 | lemma CL_subset_PO: "CompleteLattice \<subseteq> PartialOrder" | 
| 64915 | 220 | by (auto simp: PartialOrder_def CompleteLattice_def) | 
| 13115 | 221 | |
| 222 | lemmas CL_imp_PO = CL_subset_PO [THEN subsetD] | |
| 223 | ||
| 64916 | 224 | context CL | 
| 225 | begin | |
| 226 | ||
| 227 | lemma CO_refl_on: "refl_on A r" | |
| 64915 | 228 | by (rule PO_imp_refl_on) | 
| 13115 | 229 | |
| 64916 | 230 | lemma CO_antisym: "antisym r" | 
| 64915 | 231 | by (rule PO_imp_sym) | 
| 13115 | 232 | |
| 64916 | 233 | lemma CO_trans: "trans r" | 
| 64915 | 234 | by (rule PO_imp_trans) | 
| 13115 | 235 | |
| 64916 | 236 | end | 
| 237 | ||
| 13115 | 238 | lemma CompleteLatticeI: | 
| 64915 | 239 | "\<lbrakk>po \<in> PartialOrder; \<forall>S. S \<subseteq> pset po \<longrightarrow> (\<exists>L. isLub S po L); | 
| 64916 | 240 | \<forall>S. S \<subseteq> pset po \<longrightarrow> (\<exists>G. isGlb S po G)\<rbrakk> | 
| 64915 | 241 | \<Longrightarrow> po \<in> CompleteLattice" | 
| 242 | unfolding CompleteLattice_def by blast | |
| 13115 | 243 | |
| 244 | lemma (in CL) CL_dualCL: "dual cl \<in> CompleteLattice" | |
| 64915 | 245 | using cl_co | 
| 246 | apply (simp add: CompleteLattice_def dual_def) | |
| 70194 | 247 | apply (simp add: dualPO flip: dual_def isLub_dual_isGlb isGlb_dual_isLub) | 
| 64915 | 248 | done | 
| 13115 | 249 | |
| 64916 | 250 | context PO | 
| 251 | begin | |
| 252 | ||
| 70194 | 253 | lemma dualA_iff [simp]: "pset (dual cl) = pset cl" | 
| 64915 | 254 | by (simp add: dual_def) | 
| 13115 | 255 | |
| 70194 | 256 | lemma dualr_iff [simp]: "(x, y) \<in> (order (dual cl)) \<longleftrightarrow> (y, x) \<in> order cl" | 
| 64915 | 257 | by (simp add: dual_def) | 
| 13115 | 258 | |
| 64916 | 259 | lemma monotone_dual: | 
| 64915 | 260 | "monotone f (pset cl) (order cl) \<Longrightarrow> monotone f (pset (dual cl)) (order(dual cl))" | 
| 70194 | 261 | by (simp add: monotone_def) | 
| 13115 | 262 | |
| 64916 | 263 | lemma interval_dual: "\<lbrakk>x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow> interval r x y = interval (order(dual cl)) y x" | 
| 70194 | 264 | unfolding interval_def dualr_iff by (auto simp flip: r_def) | 
| 13115 | 265 | |
| 64916 | 266 | lemma interval_not_empty: "interval r a b \<noteq> {} \<Longrightarrow> (a, b) \<in> r"
 | 
| 70194 | 267 | by (simp add: interval_def) (use transE in blast) | 
| 64915 | 268 | |
| 64916 | 269 | lemma interval_imp_mem: "x \<in> interval r a b \<Longrightarrow> (a, x) \<in> r" | 
| 64915 | 270 | by (simp add: interval_def) | 
| 13115 | 271 | |
| 64916 | 272 | lemma left_in_interval: "\<lbrakk>a \<in> A; b \<in> A; interval r a b \<noteq> {}\<rbrakk> \<Longrightarrow> a \<in> interval r a b"
 | 
| 70194 | 273 | using interval_def interval_not_empty reflE by fastforce | 
| 13115 | 274 | |
| 64916 | 275 | lemma right_in_interval: "\<lbrakk>a \<in> A; b \<in> A; interval r a b \<noteq> {}\<rbrakk> \<Longrightarrow> b \<in> interval r a b"
 | 
| 70194 | 276 | by (simp add: A_def PO.dual PO.left_in_interval PO_axioms interval_dual) | 
| 13115 | 277 | |
| 64916 | 278 | end | 
| 279 | ||
| 13383 | 280 | |
| 61343 | 281 | subsection \<open>sublattice\<close> | 
| 13383 | 282 | |
| 13115 | 283 | lemma (in PO) sublattice_imp_CL: | 
| 64915 | 284 | "S <<= cl \<Longrightarrow> \<lparr>pset = S, order = induced S r\<rparr> \<in> CompleteLattice" | 
| 285 | by (simp add: sublattice_def CompleteLattice_def r_def) | |
| 13115 | 286 | |
| 287 | lemma (in CL) sublatticeI: | |
| 64915 | 288 | "\<lbrakk>S \<subseteq> A; \<lparr>pset = S, order = induced S r\<rparr> \<in> CompleteLattice\<rbrakk> \<Longrightarrow> S <<= cl" | 
| 289 | by (simp add: sublattice_def A_def r_def) | |
| 13115 | 290 | |
| 64915 | 291 | lemma (in CL) dual: "CL (dual cl)" | 
| 70194 | 292 | proof | 
| 293 | show "dual cl \<in> CompleteLattice" | |
| 64915 | 294 | using cl_co | 
| 70194 | 295 | by (simp add: CompleteLattice_def dualPO flip: isGlb_dual_isLub isLub_dual_isGlb) | 
| 296 | qed | |
| 13383 | 297 | |
| 61343 | 298 | subsection \<open>lub\<close> | 
| 13383 | 299 | |
| 64916 | 300 | context CL | 
| 301 | begin | |
| 302 | ||
| 303 | lemma lub_unique: "\<lbrakk>S \<subseteq> A; isLub S cl x; isLub S cl L\<rbrakk> \<Longrightarrow> x = L" | |
| 64915 | 304 | by (rule antisymE) (auto simp add: isLub_def r_def) | 
| 13115 | 305 | |
| 70194 | 306 | lemma lub_upper: | 
| 307 | assumes "S \<subseteq> A" "x \<in> S" shows "(x, lub S cl) \<in> r" | |
| 308 | proof - | |
| 309 | obtain L where "isLub S cl L" | |
| 310 | using CL_imp_ex_isLub \<open>S \<subseteq> A\<close> by auto | |
| 311 | then show ?thesis | |
| 312 | by (metis assms(2) isLub_def isLub_lub r_def) | |
| 313 | qed | |
| 13115 | 314 | |
| 70194 | 315 | lemma lub_least: | 
| 316 | assumes "S \<subseteq> A" and L: "L \<in> A" "\<forall>x \<in> S. (x, L) \<in> r" shows "(lub S cl, L) \<in> r" | |
| 317 | proof - | |
| 318 | obtain L' where "isLub S cl L'" | |
| 319 | using CL_imp_ex_isLub \<open>S \<subseteq> A\<close> by auto | |
| 320 | then show ?thesis | |
| 321 | by (metis A_def L isLub_def isLub_lub r_def) | |
| 322 | qed | |
| 13115 | 323 | |
| 70194 | 324 | lemma lub_in_lattice: | 
| 325 | assumes "S \<subseteq> A" shows "lub S cl \<in> A" | |
| 326 | proof - | |
| 327 | obtain L where "isLub S cl L" | |
| 328 | using CL_imp_ex_isLub \<open>S \<subseteq> A\<close> by auto | |
| 329 | then show ?thesis | |
| 330 | by (metis A_def isLub_def isLub_lub) | |
| 331 | qed | |
| 13115 | 332 | |
| 64916 | 333 | lemma lubI: | 
| 70194 | 334 | assumes A: "S \<subseteq> A" "L \<in> A" and r: "\<forall>x \<in> S. (x, L) \<in> r" | 
| 335 | and clo: "\<And>z. \<lbrakk>z \<in> A; (\<forall>y \<in> S. (y, z) \<in> r)\<rbrakk> \<Longrightarrow> (L, z) \<in> r" | |
| 336 | shows "L = lub S cl" | |
| 337 | proof - | |
| 338 | obtain L where "isLub S cl L" | |
| 339 | using CL_imp_ex_isLub assms(1) by auto | |
| 340 | then show ?thesis | |
| 341 | by (simp add: antisymE A clo lub_in_lattice lub_least lub_upper r) | |
| 342 | qed | |
| 13115 | 343 | |
| 64916 | 344 | lemma lubIa: "\<lbrakk>S \<subseteq> A; isLub S cl L\<rbrakk> \<Longrightarrow> L = lub S cl" | 
| 70194 | 345 | by (meson isLub_lub lub_unique) | 
| 13115 | 346 | |
| 64916 | 347 | lemma isLub_in_lattice: "isLub S cl L \<Longrightarrow> L \<in> A" | 
| 64915 | 348 | by (simp add: isLub_def A_def) | 
| 13115 | 349 | |
| 64916 | 350 | lemma isLub_upper: "\<lbrakk>isLub S cl L; y \<in> S\<rbrakk> \<Longrightarrow> (y, L) \<in> r" | 
| 64915 | 351 | by (simp add: isLub_def r_def) | 
| 13115 | 352 | |
| 64916 | 353 | lemma isLub_least: "\<lbrakk>isLub S cl L; z \<in> A; \<forall>y \<in> S. (y, z) \<in> r\<rbrakk> \<Longrightarrow> (L, z) \<in> r" | 
| 64915 | 354 | by (simp add: isLub_def A_def r_def) | 
| 13115 | 355 | |
| 64916 | 356 | lemma isLubI: | 
| 67613 | 357 | "\<lbrakk>L \<in> A; \<forall>y \<in> S. (y, L) \<in> r; (\<forall>z \<in> A. (\<forall>y \<in> S. (y, z)\<in>r) \<longrightarrow> (L, z) \<in> r)\<rbrakk> \<Longrightarrow> isLub S cl L" | 
| 64915 | 358 | by (simp add: isLub_def A_def r_def) | 
| 13115 | 359 | |
| 64916 | 360 | end | 
| 361 | ||
| 13383 | 362 | |
| 61343 | 363 | subsection \<open>glb\<close> | 
| 13383 | 364 | |
| 64916 | 365 | context CL | 
| 366 | begin | |
| 367 | ||
| 368 | lemma glb_in_lattice: "S \<subseteq> A \<Longrightarrow> glb S cl \<in> A" | |
| 70194 | 369 | by (metis A_def CL.lub_in_lattice dualA_iff glb_dual_lub local.dual) | 
| 13115 | 370 | |
| 64916 | 371 | lemma glb_lower: "\<lbrakk>S \<subseteq> A; x \<in> S\<rbrakk> \<Longrightarrow> (glb S cl, x) \<in> r" | 
| 70194 | 372 | by (metis A_def CL.lub_upper dualA_iff dualr_iff glb_dual_lub local.dual r_def) | 
| 13115 | 373 | |
| 64916 | 374 | end | 
| 375 | ||
| 61343 | 376 | text \<open> | 
| 13383 | 377 | Reduce the sublattice property by using substructural properties; | 
| 61933 | 378 | abandoned see \<open>Tarski_4.ML\<close>. | 
| 61343 | 379 | \<close> | 
| 13115 | 380 | |
| 64916 | 381 | context CLF | 
| 382 | begin | |
| 383 | ||
| 384 | lemma [simp]: "f \<in> pset cl \<rightarrow> pset cl \<and> monotone f (pset cl) (order cl)" | |
| 64915 | 385 | using f_cl by (simp add: CLF_set_def) | 
| 13115 | 386 | |
| 64916 | 387 | declare f_cl [simp] | 
| 13115 | 388 | |
| 389 | ||
| 64916 | 390 | lemma f_in_funcset: "f \<in> A \<rightarrow> A" | 
| 64915 | 391 | by (simp add: A_def) | 
| 13115 | 392 | |
| 64916 | 393 | lemma monotone_f: "monotone f A r" | 
| 64915 | 394 | by (simp add: A_def r_def) | 
| 13115 | 395 | |
| 64916 | 396 | lemma CLF_dual: "(dual cl, f) \<in> CLF_set" | 
| 70194 | 397 | proof - | 
| 398 | have "Tarski.monotone f A (order (dual cl))" | |
| 399 | by (metis (no_types) A_def PO.monotone_dual PO_axioms dualA_iff monotone_f r_def) | |
| 400 | then show ?thesis | |
| 401 | by (simp add: A_def CLF_set_def CL_dualCL) | |
| 402 | qed | |
| 13115 | 403 | |
| 64916 | 404 | lemma dual: "CLF (dual cl) f" | 
| 64915 | 405 | by (rule CLF.intro) (rule CLF_dual) | 
| 27681 | 406 | |
| 64916 | 407 | end | 
| 408 | ||
| 13383 | 409 | |
| 61343 | 410 | subsection \<open>fixed points\<close> | 
| 13383 | 411 | |
| 17841 | 412 | lemma fix_subset: "fix f A \<subseteq> A" | 
| 64915 | 413 | by (auto simp: fix_def) | 
| 13115 | 414 | |
| 64915 | 415 | lemma fix_imp_eq: "x \<in> fix f A \<Longrightarrow> f x = x" | 
| 416 | by (simp add: fix_def) | |
| 13115 | 417 | |
| 64915 | 418 | lemma fixf_subset: "\<lbrakk>A \<subseteq> B; x \<in> fix (\<lambda>y \<in> A. f y) A\<rbrakk> \<Longrightarrow> x \<in> fix f B" | 
| 419 | by (auto simp: fix_def) | |
| 13115 | 420 | |
| 13383 | 421 | |
| 61343 | 422 | subsection \<open>lemmas for Tarski, lub\<close> | 
| 64915 | 423 | |
| 64916 | 424 | context CLF | 
| 425 | begin | |
| 426 | ||
| 70194 | 427 | lemma lubH_le_flubH: | 
| 428 |   assumes "H = {x \<in> A. (x, f x) \<in> r}"
 | |
| 429 | shows "(lub H cl, f (lub H cl)) \<in> r" | |
| 430 | proof (intro lub_least ballI) | |
| 431 | show "H \<subseteq> A" | |
| 432 | using assms | |
| 433 | by auto | |
| 434 | show "f (lub H cl) \<in> A" | |
| 435 | using \<open>H \<subseteq> A\<close> f_in_funcset lub_in_lattice by auto | |
| 436 | show "(x, f (lub H cl)) \<in> r" if "x \<in> H" for x | |
| 437 | proof - | |
| 438 | have "(f x, f (lub H cl)) \<in> r" | |
| 439 | by (meson \<open>H \<subseteq> A\<close> in_mono lub_in_lattice lub_upper monotoneE monotone_f that) | |
| 440 | moreover have "(x, f x) \<in> r" | |
| 441 | using assms that by blast | |
| 442 | ultimately show ?thesis | |
| 443 | using po.transE by blast | |
| 444 | qed | |
| 445 | qed | |
| 13115 | 446 | |
| 70202 | 447 | lemma lubH_is_fixp: | 
| 70194 | 448 |   assumes "H = {x \<in> A. (x, f x) \<in> r}"
 | 
| 70202 | 449 | shows "lub H cl \<in> fix f A" | 
| 70194 | 450 | proof - | 
| 70202 | 451 | have "(f (lub H cl), lub H cl) \<in> r" | 
| 452 | proof - | |
| 453 | have "(lub H cl, f (lub H cl)) \<in> r" | |
| 454 | using assms lubH_le_flubH by blast | |
| 455 | then have "(f (lub H cl), f (f (lub H cl))) \<in> r" | |
| 456 | by (meson PO_imp_refl_on monotoneE monotone_f refl_on_domain) | |
| 457 | then have "f (lub H cl) \<in> H" | |
| 458 | by (metis (no_types, lifting) PO_imp_refl_on assms mem_Collect_eq refl_on_domain) | |
| 459 | then show ?thesis | |
| 460 | by (simp add: assms lub_upper) | |
| 461 | qed | |
| 462 | with assms show ?thesis | |
| 463 | by (simp add: fix_def antisymE lubH_le_flubH lub_in_lattice) | |
| 70194 | 464 | qed | 
| 13115 | 465 | |
| 70202 | 466 | lemma fixf_le_lubH: | 
| 467 |   assumes "H = {x \<in> A. (x, f x) \<in> r}" "x \<in> fix f A"
 | |
| 468 | shows "(x, lub H cl) \<in> r" | |
| 469 | proof - | |
| 470 | have "x \<in> P \<Longrightarrow> x \<in> H" | |
| 471 | by (simp add: assms P_def fix_imp_eq [of _ f A] reflE fix_subset [of f A, THEN subsetD]) | |
| 472 | with assms show ?thesis | |
| 473 | by (metis (no_types, lifting) P_def lub_upper mem_Collect_eq subset_eq) | |
| 474 | qed | |
| 13115 | 475 | |
| 64916 | 476 | |
| 61343 | 477 | subsection \<open>Tarski fixpoint theorem 1, first part\<close> | 
| 64916 | 478 | |
| 70194 | 479 | lemma T_thm_1_lub: "lub P cl = lub {x \<in> A. (x, f x) \<in> r} cl"
 | 
| 480 | proof - | |
| 481 |   have "lub {x \<in> A. (x, f x) \<in> r} cl = lub (fix f A) cl"
 | |
| 482 | proof (rule antisymE) | |
| 483 |     show "(lub {x \<in> A. (x, f x) \<in> r} cl, lub (fix f A) cl) \<in> r"
 | |
| 484 | by (simp add: fix_subset lubH_is_fixp lub_upper) | |
| 485 | have "\<And>a. a \<in> fix f A \<Longrightarrow> a \<in> A" | |
| 486 | by (meson fix_subset subset_iff) | |
| 487 |     then show "(lub (fix f A) cl, lub {x \<in> A. (x, f x) \<in> r} cl) \<in> r"
 | |
| 488 | by (simp add: fix_subset fixf_le_lubH lubH_is_fixp lub_least) | |
| 489 | qed | |
| 490 | then show ?thesis | |
| 491 | using P_def by auto | |
| 492 | qed | |
| 13115 | 493 | |
| 70194 | 494 | lemma glbH_is_fixp: | 
| 495 |   assumes "H = {x \<in> A. (f x, x) \<in> r}" shows "glb H cl \<in> P"
 | |
| 61933 | 496 | \<comment> \<open>Tarski for glb\<close> | 
| 70194 | 497 | proof - | 
| 498 | have "glb H cl \<in> fix f (pset (dual cl))" | |
| 499 | using assms CLF.lubH_is_fixp [OF dual] PO.dualr_iff PO_axioms | |
| 500 | by (fastforce simp add: A_def r_def glb_dual_lub) | |
| 501 | then show ?thesis | |
| 502 | by (simp add: A_def P_def) | |
| 503 | qed | |
| 13115 | 504 | |
| 70194 | 505 | lemma T_thm_1_glb: "glb P cl = glb {x \<in> A. (f x, x) \<in> r} cl"
 | 
| 506 | unfolding glb_dual_lub P_def A_def r_def | |
| 507 | using CLF.T_thm_1_lub dualA_iff dualr_iff local.dual by force | |
| 64915 | 508 | |
| 13115 | 509 | |
| 61343 | 510 | subsection \<open>interval\<close> | 
| 13383 | 511 | |
| 64916 | 512 | lemma rel_imp_elem: "(x, y) \<in> r \<Longrightarrow> x \<in> A" | 
| 64915 | 513 | using CO_refl_on by (auto simp: refl_on_def) | 
| 514 | ||
| 64916 | 515 | lemma interval_subset: "\<lbrakk>a \<in> A; b \<in> A\<rbrakk> \<Longrightarrow> interval r a b \<subseteq> A" | 
| 64915 | 516 | by (simp add: interval_def) (blast intro: rel_imp_elem) | 
| 13115 | 517 | |
| 64916 | 518 | lemma intervalI: "\<lbrakk>(a, x) \<in> r; (x, b) \<in> r\<rbrakk> \<Longrightarrow> x \<in> interval r a b" | 
| 64915 | 519 | by (simp add: interval_def) | 
| 13115 | 520 | |
| 64916 | 521 | lemma interval_lemma1: "\<lbrakk>S \<subseteq> interval r a b; x \<in> S\<rbrakk> \<Longrightarrow> (a, x) \<in> r" | 
| 64915 | 522 | unfolding interval_def by fast | 
| 13115 | 523 | |
| 64916 | 524 | lemma interval_lemma2: "\<lbrakk>S \<subseteq> interval r a b; x \<in> S\<rbrakk> \<Longrightarrow> (x, b) \<in> r" | 
| 64915 | 525 | unfolding interval_def by fast | 
| 13115 | 526 | |
| 64916 | 527 | lemma a_less_lub: "\<lbrakk>S \<subseteq> A; S \<noteq> {}; \<forall>x \<in> S. (a,x) \<in> r; \<forall>y \<in> S. (y, L) \<in> r\<rbrakk> \<Longrightarrow> (a, L) \<in> r"
 | 
| 64915 | 528 | by (blast intro: transE) | 
| 13115 | 529 | |
| 64916 | 530 | lemma S_intv_cl: "\<lbrakk>a \<in> A; b \<in> A; S \<subseteq> interval r a b\<rbrakk> \<Longrightarrow> S \<subseteq> A" | 
| 64915 | 531 | by (simp add: subset_trans [OF _ interval_subset]) | 
| 13115 | 532 | |
| 64916 | 533 | lemma L_in_interval: | 
| 70194 | 534 |   assumes "b \<in> A" and S: "S \<subseteq> interval r a b" "isLub S cl L" "S \<noteq> {}"
 | 
| 535 | shows "L \<in> interval r a b" | |
| 536 | proof (rule intervalI) | |
| 537 | show "(a, L) \<in> r" | |
| 538 | by (meson PO_imp_trans all_not_in_conv S interval_lemma1 isLub_upper transD) | |
| 539 | show "(L, b) \<in> r" | |
| 540 | using \<open>b \<in> A\<close> assms interval_lemma2 isLub_least by auto | |
| 541 | qed | |
| 13115 | 542 | |
| 64916 | 543 | lemma G_in_interval: | 
| 70194 | 544 |   assumes "b \<in> A" and S: "S \<subseteq> interval r a b" "isGlb S cl G" "S \<noteq> {}"
 | 
| 545 | shows "G \<in> interval r a b" | |
| 546 | proof - | |
| 547 | have "a \<in> A" | |
| 548 |     using S(1) \<open>S \<noteq> {}\<close> interval_lemma1 rel_imp_elem by blast
 | |
| 549 | with assms show ?thesis | |
| 550 | by (metis (no_types) A_def CLF.L_in_interval dualA_iff interval_dual isGlb_dual_isLub local.dual) | |
| 551 | qed | |
| 13115 | 552 | |
| 64916 | 553 | lemma intervalPO: | 
| 64915 | 554 |   "\<lbrakk>a \<in> A; b \<in> A; interval r a b \<noteq> {}\<rbrakk>
 | 
| 555 | \<Longrightarrow> \<lparr>pset = interval r a b, order = induced (interval r a b) r\<rparr> \<in> PartialOrder" | |
| 556 | by (rule po_subset_po) (simp add: interval_subset) | |
| 13115 | 557 | |
| 64916 | 558 | lemma intv_CL_lub: | 
| 70194 | 559 |   assumes "a \<in> A" "b \<in> A" "interval r a b \<noteq> {}" and S: "S \<subseteq> interval r a b"
 | 
| 560 | shows "\<exists>L. isLub S \<lparr>pset = interval r a b, order = induced (interval r a b) r\<rparr> L" | |
| 561 | proof - | |
| 562 | obtain L where L: "isLub S cl L" | |
| 563 | by (meson CL_imp_ex_isLub S_intv_cl assms(1) assms(2) assms(4)) | |
| 564 | show ?thesis | |
| 565 | unfolding isLub_def potype.simps | |
| 566 | proof (intro exI impI conjI ballI) | |
| 567 |     let ?L = "(if S = {} then a else L)"
 | |
| 568 | show Lin: "?L \<in> interval r a b" | |
| 569 | using L L_in_interval assms left_in_interval by auto | |
| 570 | show "(y, ?L) \<in> induced (interval r a b) r" if "y \<in> S" for y | |
| 571 | proof - | |
| 572 |       have "S \<noteq> {}"
 | |
| 573 | using that by blast | |
| 574 | then show ?thesis | |
| 575 | using L Lin S indI isLub_upper that by auto | |
| 576 | qed | |
| 577 | show "(?L, z) \<in> induced (interval r a b) r" | |
| 578 | if "z \<in> interval r a b" and "\<forall>y\<in>S. (y, z) \<in> induced (interval r a b) r" for z | |
| 579 | using that L | |
| 580 | apply (simp add: isLub_def induced_def interval_imp_mem) | |
| 581 | by (metis (full_types) A_def Lin \<open>a \<in> A\<close> \<open>b \<in> A\<close> interval_subset r_def subset_eq) | |
| 582 | qed | |
| 583 | qed | |
| 13115 | 584 | |
| 64916 | 585 | lemmas intv_CL_glb = intv_CL_lub [THEN Rdual] | 
| 13115 | 586 | |
| 64916 | 587 | lemma interval_is_sublattice: "\<lbrakk>a \<in> A; b \<in> A; interval r a b \<noteq> {}\<rbrakk> \<Longrightarrow> interval r a b <<= cl"
 | 
| 64915 | 588 | apply (rule sublatticeI) | 
| 589 | apply (simp add: interval_subset) | |
| 70194 | 590 | by (simp add: CompleteLatticeI intervalPO intv_CL_glb intv_CL_lub) | 
| 13115 | 591 | |
| 64916 | 592 | lemmas interv_is_compl_latt = interval_is_sublattice [THEN sublattice_imp_CL] | 
| 13115 | 593 | |
| 13383 | 594 | |
| 61343 | 595 | subsection \<open>Top and Bottom\<close> | 
| 64915 | 596 | |
| 64916 | 597 | lemma Top_dual_Bot: "Top cl = Bot (dual cl)" | 
| 70194 | 598 | by (simp add: Top_def Bot_def least_def greatest_def) | 
| 13115 | 599 | |
| 64916 | 600 | lemma Bot_dual_Top: "Bot cl = Top (dual cl)" | 
| 70194 | 601 | by (simp add: Top_def Bot_def least_def greatest_def) | 
| 13115 | 602 | |
| 64916 | 603 | lemma Bot_in_lattice: "Bot cl \<in> A" | 
| 70194 | 604 | unfolding Bot_def least_def | 
| 64915 | 605 | apply (rule_tac a = "glb A cl" in someI2) | 
| 70194 | 606 | using glb_in_lattice glb_lower by (auto simp: A_def r_def) | 
| 13115 | 607 | |
| 64916 | 608 | lemma Top_in_lattice: "Top cl \<in> A" | 
| 70194 | 609 | using A_def CLF.Bot_in_lattice Top_dual_Bot local.dual by force | 
| 13115 | 610 | |
| 64916 | 611 | lemma Top_prop: "x \<in> A \<Longrightarrow> (x, Top cl) \<in> r" | 
| 70194 | 612 | unfolding Top_def greatest_def | 
| 64915 | 613 | apply (rule_tac a = "lub A cl" in someI2) | 
| 70194 | 614 | using lub_in_lattice lub_upper by (auto simp: A_def r_def) | 
| 13115 | 615 | |
| 64916 | 616 | lemma Bot_prop: "x \<in> A \<Longrightarrow> (Bot cl, x) \<in> r" | 
| 70194 | 617 | using A_def Bot_dual_Top CLF.Top_prop dualA_iff dualr_iff local.dual r_def by fastforce | 
| 13115 | 618 | |
| 64916 | 619 | lemma Top_intv_not_empty: "x \<in> A \<Longrightarrow> interval r x (Top cl) \<noteq> {}"
 | 
| 70194 | 620 | using Top_prop intervalI reflE by force | 
| 13115 | 621 | |
| 64916 | 622 | lemma Bot_intv_not_empty: "x \<in> A \<Longrightarrow> interval r (Bot cl) x \<noteq> {}"
 | 
| 70194 | 623 | using Bot_dual_Top Bot_prop intervalI reflE by fastforce | 
| 64915 | 624 | |
| 13115 | 625 | |
| 70202 | 626 | text \<open>the set of fixed points form a partial order\<close> | 
| 627 | proposition fixf_po: "\<lparr>pset = P, order = induced P r\<rparr> \<in> PartialOrder" | |
| 64915 | 628 | by (simp add: P_def fix_subset po_subset_po) | 
| 13115 | 629 | |
| 64916 | 630 | end | 
| 631 | ||
| 632 | context Tarski | |
| 633 | begin | |
| 634 | ||
| 635 | lemma Y_subset_A: "Y \<subseteq> A" | |
| 64915 | 636 | by (rule subset_trans [OF _ fix_subset]) (rule Y_ss [simplified P_def]) | 
| 13115 | 637 | |
| 64916 | 638 | lemma lubY_in_A: "lub Y cl \<in> A" | 
| 18750 | 639 | by (rule Y_subset_A [THEN lub_in_lattice]) | 
| 13115 | 640 | |
| 64916 | 641 | lemma lubY_le_flubY: "(lub Y cl, f (lub Y cl)) \<in> r" | 
| 70194 | 642 | proof (intro lub_least Y_subset_A ballI) | 
| 643 | show "f (lub Y cl) \<in> A" | |
| 644 | by (meson Tarski.monotone_def lubY_in_A monotone_f reflE rel_imp_elem) | |
| 645 | show "(x, f (lub Y cl)) \<in> r" if "x \<in> Y" for x | |
| 646 | proof | |
| 647 | have "\<And>A. Y \<subseteq> A \<Longrightarrow> x \<in> A" | |
| 648 | using that by blast | |
| 649 | moreover have "(x, lub Y cl) \<in> r" | |
| 650 | using that by (simp add: Y_subset_A lub_upper) | |
| 651 | ultimately show "(x, f (lub Y cl)) \<in> r" | |
| 652 | by (metis (no_types) Tarski.Y_ss Tarski_axioms Y_subset_A fix_imp_eq lubY_in_A monotoneE monotone_f) | |
| 653 | qed auto | |
| 654 | qed | |
| 13115 | 655 | |
| 64916 | 656 | lemma intY1_subset: "intY1 \<subseteq> A" | 
| 70194 | 657 | unfolding intY1_def using Top_in_lattice interval_subset lubY_in_A by auto | 
| 13115 | 658 | |
| 64916 | 659 | lemmas intY1_elem = intY1_subset [THEN subsetD] | 
| 13115 | 660 | |
| 70194 | 661 | lemma intY1_f_closed: | 
| 662 | assumes "x \<in> intY1" shows "f x \<in> intY1" | |
| 663 | proof (simp add: intY1_def interval_def, rule conjI) | |
| 664 | show "(lub Y cl, f x) \<in> r" | |
| 665 | using assms intY1_elem interval_imp_mem lubY_in_A unfolding intY1_def | |
| 666 | using lubY_le_flubY monotoneE monotone_f po.transE by blast | |
| 667 | then show "(f x, Top cl) \<in> r" | |
| 668 | by (meson PO_imp_refl_on Top_prop refl_onD2) | |
| 669 | qed | |
| 13115 | 670 | |
| 64916 | 671 | lemma intY1_mono: "monotone (\<lambda> x \<in> intY1. f x) intY1 (induced intY1 r)" | 
| 64915 | 672 | apply (auto simp add: monotone_def induced_def intY1_f_closed) | 
| 673 | apply (blast intro: intY1_elem monotone_f [THEN monotoneE]) | |
| 674 | done | |
| 13115 | 675 | |
| 64916 | 676 | lemma intY1_is_cl: "\<lparr>pset = intY1, order = induced intY1 r\<rparr> \<in> CompleteLattice" | 
| 70194 | 677 | unfolding intY1_def | 
| 678 | by (simp add: Top_in_lattice Top_intv_not_empty interv_is_compl_latt lubY_in_A) | |
| 13115 | 679 | |
| 64916 | 680 | lemma v_in_P: "v \<in> P" | 
| 70194 | 681 | proof - | 
| 682 | have "v \<in> fix (restrict f intY1) intY1" | |
| 683 | unfolding v_def | |
| 684 | apply (rule CLF.glbH_is_fixp | |
| 685 | [OF CLF.intro, unfolded CLF_set_def, of "\<lparr>pset = intY1, order = induced intY1 r\<rparr>", simplified]) | |
| 686 | using intY1_f_closed intY1_is_cl intY1_mono apply blast+ | |
| 687 | done | |
| 688 | then show ?thesis | |
| 689 | unfolding P_def | |
| 690 | by (meson fixf_subset intY1_subset) | |
| 691 | qed | |
| 13115 | 692 | |
| 64916 | 693 | lemma z_in_interval: "\<lbrakk>z \<in> P; \<forall>y\<in>Y. (y, z) \<in> induced P r\<rbrakk> \<Longrightarrow> z \<in> intY1" | 
| 70194 | 694 | unfolding intY1_def P_def | 
| 695 | by (meson Top_prop Y_subset_A fix_subset in_mono indE intervalI lub_least) | |
| 13115 | 696 | |
| 64916 | 697 | lemma tarski_full_lemma: "\<exists>L. isLub Y \<lparr>pset = P, order = induced P r\<rparr> L" | 
| 70194 | 698 | proof | 
| 699 | have "(y, v) \<in> induced P r" if "y \<in> Y" for y | |
| 700 | proof - | |
| 701 | have "(y, lub Y cl) \<in> r" | |
| 702 | by (simp add: Y_subset_A lub_upper that) | |
| 703 | moreover have "(lub Y cl, v) \<in> r" | |
| 704 | by (metis (no_types, lifting) CL.glb_in_lattice CL.intro intY1_def intY1_is_cl interval_imp_mem lub_dual_glb mem_Collect_eq select_convs(1) subsetI v_def) | |
| 705 | ultimately have "(y, v) \<in> r" | |
| 706 | using po.transE by blast | |
| 707 | then show ?thesis | |
| 708 | using Y_ss indI that v_in_P by auto | |
| 709 | qed | |
| 710 | moreover have "(v, z) \<in> induced P r" if "z \<in> P" "\<forall>y\<in>Y. (y, z) \<in> induced P r" for z | |
| 711 | proof (rule indI) | |
| 712 | have "((\<lambda>x \<in> intY1. f x) z, z) \<in> induced intY1 r" | |
| 713 | by (metis P_def fix_imp_eq in_mono indI intY1_subset reflE restrict_apply' that z_in_interval) | |
| 714 | then show "(v, z) \<in> r" | |
| 715 | by (metis (no_types, lifting) CL.glb_lower CL_def indE intY1_is_cl intY1_subset mem_Collect_eq select_convs(1,2) subsetI that v_def z_in_interval) | |
| 716 | qed (auto simp: that v_in_P) | |
| 717 | ultimately | |
| 718 | show "isLub Y \<lparr>pset = P, order = induced P r\<rparr> v" | |
| 719 | by (simp add: isLub_def v_in_P) | |
| 720 | qed | |
| 13115 | 721 | |
| 64916 | 722 | end | 
| 723 | ||
| 13115 | 724 | lemma CompleteLatticeI_simp: | 
| 70202 | 725 | "\<lbrakk>po \<in> PartialOrder; \<And>S. S \<subseteq> A \<Longrightarrow> \<exists>L. isLub S po L; A = pset po\<rbrakk> \<Longrightarrow> po \<in> CompleteLattice" | 
| 70194 | 726 | by (metis CompleteLatticeI Rdual) | 
| 727 | ||
| 64915 | 728 | theorem (in CLF) Tarski_full: "\<lparr>pset = P, order = induced P r\<rparr> \<in> CompleteLattice" | 
| 70194 | 729 | proof (intro CompleteLatticeI_simp allI impI) | 
| 730 | show "\<lparr>pset = P, order = induced P r\<rparr> \<in> PartialOrder" | |
| 731 | by (simp add: fixf_po) | |
| 732 | show "\<And>S. S \<subseteq> P \<Longrightarrow> \<exists>L. isLub S \<lparr>pset = P, order = induced P r\<rparr> L" | |
| 733 | unfolding P_def A_def r_def | |
| 734 | proof (rule Tarski.tarski_full_lemma [OF Tarski.intro [OF _ Tarski_axioms.intro]]) | |
| 735 | show "CLF cl f" .. | |
| 736 | qed | |
| 737 | qed auto | |
| 7112 | 738 | |
| 739 | end |