| author | Fabian Huch <huch@in.tum.de> | 
| Fri, 15 Dec 2023 10:48:05 +0100 | |
| changeset 79291 | e9a788a75775 | 
| parent 74334 | ead56ad40e15 | 
| child 81116 | 0fb1e2dd4122 | 
| permissions | -rw-r--r-- | 
| 26241 | 1 | (* Title: HOL/Library/Option_ord.thy | 
| 2 | Author: Florian Haftmann, TU Muenchen | |
| 3 | *) | |
| 4 | ||
| 60500 | 5 | section \<open>Canonical order on option type\<close> | 
| 26241 | 6 | |
| 7 | theory Option_ord | |
| 67006 | 8 | imports Main | 
| 26241 | 9 | begin | 
| 10 | ||
| 74334 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 wenzelm parents: 
69861diff
changeset | 11 | unbundle lattice_syntax | 
| 49190 | 12 | |
| 30662 | 13 | instantiation option :: (preorder) preorder | 
| 26241 | 14 | begin | 
| 15 | ||
| 16 | definition less_eq_option where | |
| 37765 | 17 | "x \<le> y \<longleftrightarrow> (case x of None \<Rightarrow> True | Some x \<Rightarrow> (case y of None \<Rightarrow> False | Some y \<Rightarrow> x \<le> y))" | 
| 26241 | 18 | |
| 19 | definition less_option where | |
| 37765 | 20 | "x < y \<longleftrightarrow> (case y of None \<Rightarrow> False | Some y \<Rightarrow> (case x of None \<Rightarrow> True | Some x \<Rightarrow> x < y))" | 
| 26241 | 21 | |
| 26258 | 22 | lemma less_eq_option_None [simp]: "None \<le> x" | 
| 26241 | 23 | by (simp add: less_eq_option_def) | 
| 24 | ||
| 26258 | 25 | lemma less_eq_option_None_code [code]: "None \<le> x \<longleftrightarrow> True" | 
| 26241 | 26 | by simp | 
| 27 | ||
| 26258 | 28 | lemma less_eq_option_None_is_None: "x \<le> None \<Longrightarrow> x = None" | 
| 26241 | 29 | by (cases x) (simp_all add: less_eq_option_def) | 
| 30 | ||
| 26258 | 31 | lemma less_eq_option_Some_None [simp, code]: "Some x \<le> None \<longleftrightarrow> False" | 
| 26241 | 32 | by (simp add: less_eq_option_def) | 
| 33 | ||
| 26258 | 34 | lemma less_eq_option_Some [simp, code]: "Some x \<le> Some y \<longleftrightarrow> x \<le> y" | 
| 26241 | 35 | by (simp add: less_eq_option_def) | 
| 36 | ||
| 26258 | 37 | lemma less_option_None [simp, code]: "x < None \<longleftrightarrow> False" | 
| 26241 | 38 | by (simp add: less_option_def) | 
| 39 | ||
| 26258 | 40 | lemma less_option_None_is_Some: "None < x \<Longrightarrow> \<exists>z. x = Some z" | 
| 26241 | 41 | by (cases x) (simp_all add: less_option_def) | 
| 42 | ||
| 26258 | 43 | lemma less_option_None_Some [simp]: "None < Some x" | 
| 26241 | 44 | by (simp add: less_option_def) | 
| 45 | ||
| 26258 | 46 | lemma less_option_None_Some_code [code]: "None < Some x \<longleftrightarrow> True" | 
| 26241 | 47 | by simp | 
| 48 | ||
| 26258 | 49 | lemma less_option_Some [simp, code]: "Some x < Some y \<longleftrightarrow> x < y" | 
| 26241 | 50 | by (simp add: less_option_def) | 
| 51 | ||
| 60679 | 52 | instance | 
| 53 | by standard | |
| 54 | (auto simp add: less_eq_option_def less_option_def less_le_not_le | |
| 55 | elim: order_trans split: option.splits) | |
| 26241 | 56 | |
| 60679 | 57 | end | 
| 30662 | 58 | |
| 60679 | 59 | instance option :: (order) order | 
| 60 | by standard (auto simp add: less_eq_option_def less_option_def split: option.splits) | |
| 61 | ||
| 62 | instance option :: (linorder) linorder | |
| 63 | by standard (auto simp add: less_eq_option_def less_option_def split: option.splits) | |
| 30662 | 64 | |
| 52729 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 haftmann parents: 
49190diff
changeset | 65 | instantiation option :: (order) order_bot | 
| 30662 | 66 | begin | 
| 67 | ||
| 60679 | 68 | definition bot_option where "\<bottom> = None" | 
| 30662 | 69 | |
| 60679 | 70 | instance | 
| 71 | by standard (simp add: bot_option_def) | |
| 30662 | 72 | |
| 73 | end | |
| 74 | ||
| 52729 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 haftmann parents: 
49190diff
changeset | 75 | instantiation option :: (order_top) order_top | 
| 30662 | 76 | begin | 
| 77 | ||
| 60679 | 78 | definition top_option where "\<top> = Some \<top>" | 
| 30662 | 79 | |
| 60679 | 80 | instance | 
| 81 | by standard (simp add: top_option_def less_eq_option_def split: option.split) | |
| 26241 | 82 | |
| 83 | end | |
| 30662 | 84 | |
| 60679 | 85 | instance option :: (wellorder) wellorder | 
| 86 | proof | |
| 87 | fix P :: "'a option \<Rightarrow> bool" | |
| 88 | fix z :: "'a option" | |
| 30662 | 89 | assume H: "\<And>x. (\<And>y. y < x \<Longrightarrow> P y) \<Longrightarrow> P x" | 
| 90 | have "P None" by (rule H) simp | |
| 67091 | 91 | then have P_Some [case_names Some]: "P z" if "\<And>x. z = Some x \<Longrightarrow> (P \<circ> Some) x" for z | 
| 60679 | 92 | using \<open>P None\<close> that by (cases z) simp_all | 
| 93 | show "P z" | |
| 94 | proof (cases z rule: P_Some) | |
| 30662 | 95 | case (Some w) | 
| 67091 | 96 | show "(P \<circ> Some) w" | 
| 60679 | 97 | proof (induct rule: less_induct) | 
| 30662 | 98 | case (less x) | 
| 60679 | 99 | have "P (Some x)" | 
| 100 | proof (rule H) | |
| 30662 | 101 | fix y :: "'a option" | 
| 102 | assume "y < Some x" | |
| 60679 | 103 | show "P y" | 
| 104 | proof (cases y rule: P_Some) | |
| 105 | case (Some v) | |
| 106 | with \<open>y < Some x\<close> have "v < x" by simp | |
| 67091 | 107 | with less show "(P \<circ> Some) v" . | 
| 30662 | 108 | qed | 
| 109 | qed | |
| 110 | then show ?case by simp | |
| 111 | qed | |
| 112 | qed | |
| 113 | qed | |
| 114 | ||
| 49190 | 115 | instantiation option :: (inf) inf | 
| 116 | begin | |
| 117 | ||
| 118 | definition inf_option where | |
| 119 | "x \<sqinter> y = (case x of None \<Rightarrow> None | Some x \<Rightarrow> (case y of None \<Rightarrow> None | Some y \<Rightarrow> Some (x \<sqinter> y)))" | |
| 120 | ||
| 60679 | 121 | lemma inf_None_1 [simp, code]: "None \<sqinter> y = None" | 
| 49190 | 122 | by (simp add: inf_option_def) | 
| 123 | ||
| 60679 | 124 | lemma inf_None_2 [simp, code]: "x \<sqinter> None = None" | 
| 49190 | 125 | by (cases x) (simp_all add: inf_option_def) | 
| 126 | ||
| 60679 | 127 | lemma inf_Some [simp, code]: "Some x \<sqinter> Some y = Some (x \<sqinter> y)" | 
| 49190 | 128 | by (simp add: inf_option_def) | 
| 129 | ||
| 130 | instance .. | |
| 131 | ||
| 30662 | 132 | end | 
| 49190 | 133 | |
| 134 | instantiation option :: (sup) sup | |
| 135 | begin | |
| 136 | ||
| 137 | definition sup_option where | |
| 138 | "x \<squnion> y = (case x of None \<Rightarrow> y | Some x' \<Rightarrow> (case y of None \<Rightarrow> x | Some y \<Rightarrow> Some (x' \<squnion> y)))" | |
| 139 | ||
| 60679 | 140 | lemma sup_None_1 [simp, code]: "None \<squnion> y = y" | 
| 49190 | 141 | by (simp add: sup_option_def) | 
| 142 | ||
| 60679 | 143 | lemma sup_None_2 [simp, code]: "x \<squnion> None = x" | 
| 49190 | 144 | by (cases x) (simp_all add: sup_option_def) | 
| 145 | ||
| 60679 | 146 | lemma sup_Some [simp, code]: "Some x \<squnion> Some y = Some (x \<squnion> y)" | 
| 49190 | 147 | by (simp add: sup_option_def) | 
| 148 | ||
| 149 | instance .. | |
| 150 | ||
| 151 | end | |
| 152 | ||
| 60679 | 153 | instance option :: (semilattice_inf) semilattice_inf | 
| 154 | proof | |
| 49190 | 155 | fix x y z :: "'a option" | 
| 156 | show "x \<sqinter> y \<le> x" | |
| 60679 | 157 | by (cases x, simp_all, cases y, simp_all) | 
| 49190 | 158 | show "x \<sqinter> y \<le> y" | 
| 60679 | 159 | by (cases x, simp_all, cases y, simp_all) | 
| 49190 | 160 | show "x \<le> y \<Longrightarrow> x \<le> z \<Longrightarrow> x \<le> y \<sqinter> z" | 
| 60679 | 161 | by (cases x, simp_all, cases y, simp_all, cases z, simp_all) | 
| 49190 | 162 | qed | 
| 163 | ||
| 60679 | 164 | instance option :: (semilattice_sup) semilattice_sup | 
| 165 | proof | |
| 49190 | 166 | fix x y z :: "'a option" | 
| 167 | show "x \<le> x \<squnion> y" | |
| 60679 | 168 | by (cases x, simp_all, cases y, simp_all) | 
| 49190 | 169 | show "y \<le> x \<squnion> y" | 
| 60679 | 170 | by (cases x, simp_all, cases y, simp_all) | 
| 49190 | 171 | fix x y z :: "'a option" | 
| 172 | show "y \<le> x \<Longrightarrow> z \<le> x \<Longrightarrow> y \<squnion> z \<le> x" | |
| 60679 | 173 | by (cases y, simp_all, cases z, simp_all, cases x, simp_all) | 
| 49190 | 174 | qed | 
| 175 | ||
| 176 | instance option :: (lattice) lattice .. | |
| 177 | ||
| 178 | instance option :: (lattice) bounded_lattice_bot .. | |
| 179 | ||
| 180 | instance option :: (bounded_lattice_top) bounded_lattice_top .. | |
| 181 | ||
| 182 | instance option :: (bounded_lattice_top) bounded_lattice .. | |
| 183 | ||
| 184 | instance option :: (distrib_lattice) distrib_lattice | |
| 185 | proof | |
| 186 | fix x y z :: "'a option" | |
| 187 | show "x \<squnion> y \<sqinter> z = (x \<squnion> y) \<sqinter> (x \<squnion> z)" | |
| 60679 | 188 | by (cases x, simp_all, cases y, simp_all, cases z, simp_all add: sup_inf_distrib1 inf_commute) | 
| 189 | qed | |
| 49190 | 190 | |
| 191 | instantiation option :: (complete_lattice) complete_lattice | |
| 192 | begin | |
| 193 | ||
| 194 | definition Inf_option :: "'a option set \<Rightarrow> 'a option" where | |
| 195 | "\<Sqinter>A = (if None \<in> A then None else Some (\<Sqinter>Option.these A))" | |
| 196 | ||
| 60679 | 197 | lemma None_in_Inf [simp]: "None \<in> A \<Longrightarrow> \<Sqinter>A = None" | 
| 49190 | 198 | by (simp add: Inf_option_def) | 
| 199 | ||
| 200 | definition Sup_option :: "'a option set \<Rightarrow> 'a option" where | |
| 201 |   "\<Squnion>A = (if A = {} \<or> A = {None} then None else Some (\<Squnion>Option.these A))"
 | |
| 202 | ||
| 60679 | 203 | lemma empty_Sup [simp]: "\<Squnion>{} = None"
 | 
| 49190 | 204 | by (simp add: Sup_option_def) | 
| 205 | ||
| 60679 | 206 | lemma singleton_None_Sup [simp]: "\<Squnion>{None} = None"
 | 
| 49190 | 207 | by (simp add: Sup_option_def) | 
| 208 | ||
| 60679 | 209 | instance | 
| 210 | proof | |
| 49190 | 211 | fix x :: "'a option" and A | 
| 212 | assume "x \<in> A" | |
| 213 | then show "\<Sqinter>A \<le> x" | |
| 214 | by (cases x) (auto simp add: Inf_option_def in_these_eq intro: Inf_lower) | |
| 215 | next | |
| 216 | fix z :: "'a option" and A | |
| 217 | assume *: "\<And>x. x \<in> A \<Longrightarrow> z \<le> x" | |
| 218 | show "z \<le> \<Sqinter>A" | |
| 219 | proof (cases z) | |
| 220 | case None then show ?thesis by simp | |
| 221 | next | |
| 222 | case (Some y) | |
| 223 | show ?thesis | |
| 224 | by (auto simp add: Inf_option_def in_these_eq Some intro!: Inf_greatest dest!: *) | |
| 225 | qed | |
| 226 | next | |
| 227 | fix x :: "'a option" and A | |
| 228 | assume "x \<in> A" | |
| 229 | then show "x \<le> \<Squnion>A" | |
| 230 | by (cases x) (auto simp add: Sup_option_def in_these_eq intro: Sup_upper) | |
| 231 | next | |
| 232 | fix z :: "'a option" and A | |
| 233 | assume *: "\<And>x. x \<in> A \<Longrightarrow> x \<le> z" | |
| 234 | show "\<Squnion>A \<le> z " | |
| 235 | proof (cases z) | |
| 236 | case None | |
| 237 | with * have "\<And>x. x \<in> A \<Longrightarrow> x = None" by (auto dest: less_eq_option_None_is_None) | |
| 238 |     then have "A = {} \<or> A = {None}" by blast
 | |
| 239 | then show ?thesis by (simp add: Sup_option_def) | |
| 240 | next | |
| 241 | case (Some y) | |
| 242 | from * have "\<And>w. Some w \<in> A \<Longrightarrow> Some w \<le> z" . | |
| 243 | with Some have "\<And>w. w \<in> Option.these A \<Longrightarrow> w \<le> y" | |
| 244 | by (simp add: in_these_eq) | |
| 245 | then have "\<Squnion>Option.these A \<le> y" by (rule Sup_least) | |
| 246 | with Some show ?thesis by (simp add: Sup_option_def) | |
| 247 | qed | |
| 52729 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 haftmann parents: 
49190diff
changeset | 248 | next | 
| 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 haftmann parents: 
49190diff
changeset | 249 |   show "\<Squnion>{} = (\<bottom>::'a option)"
 | 
| 60679 | 250 | by (auto simp: bot_option_def) | 
| 52729 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 haftmann parents: 
49190diff
changeset | 251 |   show "\<Sqinter>{} = (\<top>::'a option)"
 | 
| 60679 | 252 | by (auto simp: top_option_def Inf_option_def) | 
| 49190 | 253 | qed | 
| 254 | ||
| 255 | end | |
| 256 | ||
| 257 | lemma Some_Inf: | |
| 258 | "Some (\<Sqinter>A) = \<Sqinter>(Some ` A)" | |
| 259 | by (auto simp add: Inf_option_def) | |
| 260 | ||
| 261 | lemma Some_Sup: | |
| 262 |   "A \<noteq> {} \<Longrightarrow> Some (\<Squnion>A) = \<Squnion>(Some ` A)"
 | |
| 263 | by (auto simp add: Sup_option_def) | |
| 264 | ||
| 265 | lemma Some_INF: | |
| 266 | "Some (\<Sqinter>x\<in>A. f x) = (\<Sqinter>x\<in>A. Some (f x))" | |
| 69861 
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
 haftmann parents: 
69661diff
changeset | 267 | by (simp add: Some_Inf image_comp) | 
| 49190 | 268 | |
| 269 | lemma Some_SUP: | |
| 270 |   "A \<noteq> {} \<Longrightarrow> Some (\<Squnion>x\<in>A. f x) = (\<Squnion>x\<in>A. Some (f x))"
 | |
| 69861 
62e47f06d22c
avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
 haftmann parents: 
69661diff
changeset | 271 | by (simp add: Some_Sup image_comp) | 
| 49190 | 272 | |
| 69313 | 273 | lemma option_Inf_Sup: "\<Sqinter>(Sup ` A) \<le> \<Squnion>(Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y})"
 | 
| 274 |   for A :: "('a::complete_distrib_lattice option) set set"
 | |
| 67829 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 275 | proof (cases "{} \<in> A")
 | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 276 | case True | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 277 | then show ?thesis | 
| 67951 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 278 | by (rule INF_lower2, simp_all) | 
| 67829 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 279 | next | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 280 | case False | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 281 |   from this have X: "{} \<notin> A"
 | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 282 | by simp | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 283 | then show ?thesis | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 284 |   proof (cases "{None} \<in> A")
 | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 285 | case True | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 286 | then show ?thesis | 
| 67951 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 287 | by (rule INF_lower2, simp_all) | 
| 49190 | 288 | next | 
| 67829 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 289 | case False | 
| 67951 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 290 | |
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 291 |     {fix y
 | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 292 | assume A: "y \<in> A" | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 293 |       have "Sup (y - {None}) = Sup y"
 | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 294 | by (metis (no_types, lifting) Sup_option_def insert_Diff_single these_insert_None these_not_empty_eq) | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 295 |       from A and this have "(\<exists>z. y - {None} = z - {None} \<and> z \<in> A) \<and> \<Squnion>y = \<Squnion>(y - {None})"
 | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 296 | by auto | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 297 | } | 
| 67829 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 298 |     from this have A: "Sup ` A = (Sup ` {y - {None} | y. y\<in>A})"
 | 
| 67951 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 299 | by (auto simp add: image_def) | 
| 67829 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 300 | |
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 301 |     have [simp]: "\<And>y. y \<in> A \<Longrightarrow> \<exists>ya. {ya. \<exists>x. x \<in> y \<and> (\<exists>y. x = Some y) \<and> ya = the x} 
 | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 302 |           = {y. \<exists>x\<in>ya - {None}. y = the x} \<and> ya \<in> A"
 | 
| 67951 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 303 | by (rule exI, auto) | 
| 67829 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 304 | |
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 305 | have [simp]: "\<And>y. y \<in> A \<Longrightarrow> | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 306 |          (\<exists>ya. y - {None} = ya - {None} \<and> ya \<in> A) \<and> \<Squnion>{ya. \<exists>x\<in>y - {None}. ya = the x} 
 | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 307 |           = \<Squnion>{ya. \<exists>x. x \<in> y \<and> (\<exists>y. x = Some y) \<and> ya = the x}"
 | 
| 67951 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 308 | apply (safe, blast) | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 309 | by (rule arg_cong [of _ _ Sup], auto) | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 310 |     {fix y
 | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 311 | assume [simp]: "y \<in> A" | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 312 |       have "\<exists>x. (\<exists>y. x = {ya. \<exists>x\<in>y - {None}. ya = the x} \<and> y \<in> A) \<and> \<Squnion>{ya. \<exists>x. x \<in> y \<and> (\<exists>y. x = Some y) \<and> ya = the x} = \<Squnion>x"
 | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 313 |       and "\<exists>x. (\<exists>y. x = y - {None} \<and> y \<in> A) \<and> \<Squnion>{ya. \<exists>x\<in>y - {None}. ya = the x} = \<Squnion>{y. \<exists>xa. xa \<in> x \<and> (\<exists>y. xa = Some y) \<and> y = the xa}"
 | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 314 |          apply (rule exI [of _ "{ya. \<exists>x. x \<in> y \<and> (\<exists>y. x = Some y) \<and> ya = the x}"], simp)
 | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 315 |         by (rule exI [of _ "y - {None}"], simp)
 | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 316 | } | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 317 |     from this have C: "(\<lambda>x.  (\<Squnion>Option.these x)) ` {y - {None} |y. y \<in> A} =  (Sup ` {the ` (y - {None}) |y. y \<in> A})"
 | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 318 | by (simp add: image_def Option.these_def, safe, simp_all) | 
| 67829 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 319 | |
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 320 | have D: "\<forall> f . \<exists>Y\<in>A. f Y \<notin> Y \<Longrightarrow> False" | 
| 67951 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 321 | by (drule spec [of _ "\<lambda> Y . SOME x . x \<in> Y"], simp add: X some_in_eq) | 
| 67829 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 322 | |
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 323 |     define F where "F = (\<lambda> Y . SOME x::'a option . x \<in> (Y - {None}))"
 | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 324 | |
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 325 |     have G: "\<And> Y . Y \<in> A \<Longrightarrow> \<exists> x . x \<in> Y - {None}"
 | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 326 | by (metis False X all_not_in_conv insert_Diff_single these_insert_None these_not_empty_eq) | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 327 | |
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 328 |     have F: "\<And> Y . Y \<in> A \<Longrightarrow> F Y \<in> (Y - {None})"
 | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 329 | by (metis F_def G empty_iff some_in_eq) | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 330 | |
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 331 | have "Some \<bottom> \<le> Inf (F ` A)" | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 332 | by (metis (no_types, lifting) Diff_iff F Inf_option_def bot.extremum image_iff | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 333 | less_eq_option_Some singletonI) | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 334 | |
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 335 | from this have "Inf (F ` A) \<noteq> None" | 
| 67951 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 336 | by (cases "\<Sqinter>x\<in>A. F x", simp_all) | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 337 | |
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 338 |     from this have "Inf (F ` A) \<noteq> None \<and> Inf (F ` A) \<in> Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y}"
 | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 339 | using F by auto | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 340 | |
| 67829 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 341 |     from this have "\<exists> x . x \<noteq> None \<and> x \<in> Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y}"
 | 
| 67951 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 342 | by blast | 
| 67829 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 343 | |
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 344 |     from this have E:" Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y} = {None} \<Longrightarrow> False"
 | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 345 | by blast | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 346 | |
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 347 |     have [simp]: "((\<Squnion>x\<in>{f ` A |f. \<forall>Y\<in>A. f Y \<in> Y}. \<Sqinter>x) = None) = False"
 | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 348 |       by (metis (no_types, lifting) E Sup_option_def \<open>\<exists>x. x \<noteq> None \<and> x \<in> Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y}\<close> 
 | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 349 | ex_in_conv option.simps(3)) | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 350 | |
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 351 |     have B: "Option.these ((\<lambda>x. Some (\<Squnion>Option.these x)) ` {y - {None} |y. y \<in> A}) 
 | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 352 |         = ((\<lambda>x. (\<Squnion> Option.these x)) ` {y - {None} |y. y \<in> A})"
 | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 353 | by (metis image_image these_image_Some_eq) | 
| 67951 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 354 |     {
 | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 355 | fix f | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 356 |       assume A: "\<And> Y . (\<exists>y. Y = the ` (y - {None}) \<and> y \<in> A) \<Longrightarrow> f Y \<in> Y"
 | 
| 67829 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 357 | |
| 67951 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 358 |       have "\<And>xa. xa \<in> A \<Longrightarrow> f {y. \<exists>a\<in>xa - {None}. y = the a} = f (the ` (xa - {None}))"
 | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 359 | by (simp add: image_def) | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 360 |       from this have [simp]: "\<And>xa. xa \<in> A \<Longrightarrow> \<exists>x\<in>A. f {y. \<exists>a\<in>xa - {None}. y = the a} = f (the ` (x - {None}))"
 | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 361 | by blast | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 362 |       have "\<And>xa. xa \<in> A \<Longrightarrow> f (the ` (xa - {None})) = f {y. \<exists>a \<in> xa - {None}. y = the a} \<and> xa \<in> A"
 | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 363 | by (simp add: image_def) | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 364 |       from this have [simp]: "\<And>xa. xa \<in> A \<Longrightarrow> \<exists>x. f (the ` (xa - {None})) = f {y. \<exists>a\<in>x - {None}. y = the a} \<and> x \<in> A"
 | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 365 | by blast | 
| 67829 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 366 | |
| 67951 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 367 |       {
 | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 368 | fix Y | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 369 |         have "Y \<in> A \<Longrightarrow> Some (f (the ` (Y - {None}))) \<in> Y"
 | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 370 |           using A [of "the ` (Y - {None})"] apply (simp add: image_def)
 | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 371 | using option.collapse by fastforce | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 372 | } | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 373 |       from this have [simp]: "\<And> Y . Y \<in> A \<Longrightarrow> Some (f (the ` (Y - {None}))) \<in> Y"
 | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 374 | by blast | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 375 |       have [simp]: "(\<Sqinter>x\<in>A. Some (f {y. \<exists>x\<in>x - {None}. y = the x})) = \<Sqinter>{Some (f {y. \<exists>a\<in>x - {None}. y = the a}) |x. x \<in> A}"
 | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 376 | by (simp add: Setcompr_eq_image) | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 377 | |
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 378 |       have [simp]: "\<exists>x. (\<exists>f. x = {y. \<exists>x\<in>A. y = f x} \<and> (\<forall>Y\<in>A. f Y \<in> Y)) \<and> \<Sqinter>{Some (f {y. \<exists>a\<in>x - {None}. y = the a}) |x. x \<in> A} = \<Sqinter>x"
 | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 379 |         apply (rule exI [of _ "{Some (f {y. \<exists>a\<in>x - {None}. y = the a}) | x . x\<in> A}"], safe)
 | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 380 |         by (rule exI [of _ "(\<lambda> Y . Some (f (the ` (Y - {None})))) "], safe, simp_all)
 | 
| 67829 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 381 | |
| 67951 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 382 |       {
 | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 383 | fix xb | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 384 |         have "xb \<in> A \<Longrightarrow> (\<Sqinter>x\<in>{{ya. \<exists>x\<in>y - {None}. ya = the x} |y. y \<in> A}. f x) \<le> f {y. \<exists>x\<in>xb - {None}. y = the x}"
 | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 385 |           apply (rule INF_lower2 [of "{y. \<exists>x\<in>xb - {None}. y = the x}"])
 | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 386 | by blast+ | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 387 | } | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 388 |       from this have [simp]: "(\<Sqinter>x\<in>{the ` (y - {None}) |y. y \<in> A}. f x) \<le> the (\<Sqinter>Y\<in>A. Some (f (the ` (Y - {None}))))"
 | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 389 | apply (simp add: Inf_option_def image_def Option.these_def) | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 390 | by (rule Inf_greatest, clarsimp) | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 391 |       have [simp]: "the (\<Sqinter>Y\<in>A. Some (f (the ` (Y - {None})))) \<in> Option.these (Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y})"
 | 
| 69661 | 392 | apply (auto simp add: Option.these_def) | 
| 393 | apply (rule imageI) | |
| 394 | apply auto | |
| 395 |         using \<open>\<And>Y. Y \<in> A \<Longrightarrow> Some (f (the ` (Y - {None}))) \<in> Y\<close> apply blast
 | |
| 396 | apply (auto simp add: Some_INF [symmetric]) | |
| 397 | done | |
| 67951 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 398 |       have "(\<Sqinter>x\<in>{the ` (y - {None}) |y. y \<in> A}. f x) \<le> \<Squnion>Option.these (Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y})"
 | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 399 |         by (rule Sup_upper2 [of "the (Inf ((\<lambda> Y . Some (f (the ` (Y - {None})) )) ` A))"], simp_all)
 | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 400 | } | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 401 |     from this have X: "\<And> f . \<forall>Y. (\<exists>y. Y = the ` (y - {None}) \<and> y \<in> A) \<longrightarrow> f Y \<in> Y \<Longrightarrow>
 | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 402 |       (\<Sqinter>x\<in>{the ` (y - {None}) |y. y \<in> A}. f x) \<le> \<Squnion>Option.these (Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y})"
 | 
| 67829 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 403 | by blast | 
| 67951 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 404 | |
| 67829 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 405 | |
| 67951 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 406 |     have [simp]: "\<And> x . x\<in>{y - {None} |y. y \<in> A} \<Longrightarrow>  x \<noteq> {} \<and> x \<noteq> {None}"
 | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 407 | using F by fastforce | 
| 67829 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 408 | |
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 409 |     have "(Inf (Sup `A)) = (Inf (Sup ` {y - {None} | y. y\<in>A}))"
 | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 410 | by (subst A, simp) | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 411 | |
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 412 |     also have "... = (\<Sqinter>x\<in>{y - {None} |y. y \<in> A}. if x = {} \<or> x = {None} then None else Some (\<Squnion>Option.these x))"
 | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 413 | by (simp add: Sup_option_def) | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 414 | |
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 415 |     also have "... = (\<Sqinter>x\<in>{y - {None} |y. y \<in> A}. Some (\<Squnion>Option.these x))"
 | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 416 | using G by fastforce | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 417 | |
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 418 |     also have "... = Some (\<Sqinter>Option.these ((\<lambda>x. Some (\<Squnion>Option.these x)) ` {y - {None} |y. y \<in> A}))"
 | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 419 | by (simp add: Inf_option_def, safe) | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 420 | |
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 421 |     also have "... =  Some (\<Sqinter> ((\<lambda>x.  (\<Squnion>Option.these x)) ` {y - {None} |y. y \<in> A}))"
 | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 422 | by (simp add: B) | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 423 | |
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 424 |     also have "... = Some (Inf (Sup ` {the ` (y - {None}) |y. y \<in> A}))"
 | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 425 | by (unfold C, simp) | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 426 | thm Inf_Sup | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 427 |     also have "... = Some (\<Squnion>x\<in>{f ` {the ` (y - {None}) |y. y \<in> A} |f. \<forall>Y. (\<exists>y. Y = the ` (y - {None}) \<and> y \<in> A) \<longrightarrow> f Y \<in> Y}. \<Sqinter>x) "
 | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 428 | by (simp add: Inf_Sup) | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 429 | |
| 69313 | 430 |     also have "... \<le> \<Squnion> (Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y})"
 | 
| 431 |     proof (cases "\<Squnion> (Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y})")
 | |
| 67951 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 432 | case None | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 433 | then show ?thesis by (simp add: less_eq_option_def) | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 434 | next | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 435 | case (Some a) | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 436 | then show ?thesis | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 437 | apply simp | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 438 | apply (rule Sup_least, safe) | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 439 | apply (simp add: Sup_option_def) | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 440 |         apply (cases "(\<forall>f. \<exists>Y\<in>A. f Y \<notin> Y) \<or> Inf ` {f ` A |f. \<forall>Y\<in>A. f Y \<in> Y} = {None}", simp_all)
 | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 441 | by (drule X, simp) | 
| 
655aa11359dc
Removed some uses of deprecated _tac methods. (Patch from Viorel Preoteasa)
 Manuel Eberl <eberlm@in.tum.de> parents: 
67829diff
changeset | 442 | qed | 
| 67829 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 443 | finally show ?thesis by simp | 
| 49190 | 444 | qed | 
| 445 | qed | |
| 446 | ||
| 67829 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 447 | instance option :: (complete_distrib_lattice) complete_distrib_lattice | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 448 | by (standard, simp add: option_Inf_Sup) | 
| 
2a6ef5ba4822
Changes to complete distributive lattices due to Viorel Preoteasa
 Manuel Eberl <eberlm@in.tum.de> parents: 
67091diff
changeset | 449 | |
| 60679 | 450 | instance option :: (complete_linorder) complete_linorder .. | 
| 49190 | 451 | |
| 74334 
ead56ad40e15
bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
 wenzelm parents: 
69861diff
changeset | 452 | unbundle no_lattice_syntax | 
| 49190 | 453 | |
| 454 | end |