| author | wenzelm | 
| Tue, 04 Apr 2017 22:16:42 +0200 | |
| changeset 65382 | de848ac5e0d7 | 
| parent 62343 | 24106dc44def | 
| child 66453 | cc19f7ca2ed6 | 
| 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 | |
| 30662 | 8 | imports Option Main | 
| 26241 | 9 | begin | 
| 10 | ||
| 49190 | 11 | notation | 
| 12 |   bot ("\<bottom>") and
 | |
| 13 |   top ("\<top>") and
 | |
| 14 | inf (infixl "\<sqinter>" 70) and | |
| 15 | sup (infixl "\<squnion>" 65) and | |
| 16 |   Inf  ("\<Sqinter>_" [900] 900) and
 | |
| 17 |   Sup  ("\<Squnion>_" [900] 900)
 | |
| 18 | ||
| 61955 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
60679diff
changeset | 19 | syntax | 
| 49190 | 20 |   "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Sqinter>_./ _)" [0, 10] 10)
 | 
| 21 |   "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10)
 | |
| 22 |   "_SUP1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Squnion>_./ _)" [0, 10] 10)
 | |
| 23 |   "_SUP"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10)
 | |
| 24 | ||
| 25 | ||
| 30662 | 26 | instantiation option :: (preorder) preorder | 
| 26241 | 27 | begin | 
| 28 | ||
| 29 | definition less_eq_option where | |
| 37765 | 30 | "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 | 31 | |
| 32 | definition less_option where | |
| 37765 | 33 | "x < y \<longleftrightarrow> (case y of None \<Rightarrow> False | Some y \<Rightarrow> (case x of None \<Rightarrow> True | Some x \<Rightarrow> x < y))" | 
| 26241 | 34 | |
| 26258 | 35 | lemma less_eq_option_None [simp]: "None \<le> x" | 
| 26241 | 36 | by (simp add: less_eq_option_def) | 
| 37 | ||
| 26258 | 38 | lemma less_eq_option_None_code [code]: "None \<le> x \<longleftrightarrow> True" | 
| 26241 | 39 | by simp | 
| 40 | ||
| 26258 | 41 | lemma less_eq_option_None_is_None: "x \<le> None \<Longrightarrow> x = None" | 
| 26241 | 42 | by (cases x) (simp_all add: less_eq_option_def) | 
| 43 | ||
| 26258 | 44 | lemma less_eq_option_Some_None [simp, code]: "Some x \<le> None \<longleftrightarrow> False" | 
| 26241 | 45 | by (simp add: less_eq_option_def) | 
| 46 | ||
| 26258 | 47 | lemma less_eq_option_Some [simp, code]: "Some x \<le> Some y \<longleftrightarrow> x \<le> y" | 
| 26241 | 48 | by (simp add: less_eq_option_def) | 
| 49 | ||
| 26258 | 50 | lemma less_option_None [simp, code]: "x < None \<longleftrightarrow> False" | 
| 26241 | 51 | by (simp add: less_option_def) | 
| 52 | ||
| 26258 | 53 | lemma less_option_None_is_Some: "None < x \<Longrightarrow> \<exists>z. x = Some z" | 
| 26241 | 54 | by (cases x) (simp_all add: less_option_def) | 
| 55 | ||
| 26258 | 56 | lemma less_option_None_Some [simp]: "None < Some x" | 
| 26241 | 57 | by (simp add: less_option_def) | 
| 58 | ||
| 26258 | 59 | lemma less_option_None_Some_code [code]: "None < Some x \<longleftrightarrow> True" | 
| 26241 | 60 | by simp | 
| 61 | ||
| 26258 | 62 | lemma less_option_Some [simp, code]: "Some x < Some y \<longleftrightarrow> x < y" | 
| 26241 | 63 | by (simp add: less_option_def) | 
| 64 | ||
| 60679 | 65 | instance | 
| 66 | by standard | |
| 67 | (auto simp add: less_eq_option_def less_option_def less_le_not_le | |
| 68 | elim: order_trans split: option.splits) | |
| 26241 | 69 | |
| 60679 | 70 | end | 
| 30662 | 71 | |
| 60679 | 72 | instance option :: (order) order | 
| 73 | by standard (auto simp add: less_eq_option_def less_option_def split: option.splits) | |
| 74 | ||
| 75 | instance option :: (linorder) linorder | |
| 76 | by standard (auto simp add: less_eq_option_def less_option_def split: option.splits) | |
| 30662 | 77 | |
| 52729 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 haftmann parents: 
49190diff
changeset | 78 | instantiation option :: (order) order_bot | 
| 30662 | 79 | begin | 
| 80 | ||
| 60679 | 81 | definition bot_option where "\<bottom> = None" | 
| 30662 | 82 | |
| 60679 | 83 | instance | 
| 84 | by standard (simp add: bot_option_def) | |
| 30662 | 85 | |
| 86 | end | |
| 87 | ||
| 52729 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 haftmann parents: 
49190diff
changeset | 88 | instantiation option :: (order_top) order_top | 
| 30662 | 89 | begin | 
| 90 | ||
| 60679 | 91 | definition top_option where "\<top> = Some \<top>" | 
| 30662 | 92 | |
| 60679 | 93 | instance | 
| 94 | by standard (simp add: top_option_def less_eq_option_def split: option.split) | |
| 26241 | 95 | |
| 96 | end | |
| 30662 | 97 | |
| 60679 | 98 | instance option :: (wellorder) wellorder | 
| 99 | proof | |
| 100 | fix P :: "'a option \<Rightarrow> bool" | |
| 101 | fix z :: "'a option" | |
| 30662 | 102 | assume H: "\<And>x. (\<And>y. y < x \<Longrightarrow> P y) \<Longrightarrow> P x" | 
| 103 | have "P None" by (rule H) simp | |
| 60679 | 104 | then have P_Some [case_names Some]: "P z" if "\<And>x. z = Some x \<Longrightarrow> (P o Some) x" for z | 
| 105 | using \<open>P None\<close> that by (cases z) simp_all | |
| 106 | show "P z" | |
| 107 | proof (cases z rule: P_Some) | |
| 30662 | 108 | case (Some w) | 
| 60679 | 109 | show "(P o Some) w" | 
| 110 | proof (induct rule: less_induct) | |
| 30662 | 111 | case (less x) | 
| 60679 | 112 | have "P (Some x)" | 
| 113 | proof (rule H) | |
| 30662 | 114 | fix y :: "'a option" | 
| 115 | assume "y < Some x" | |
| 60679 | 116 | show "P y" | 
| 117 | proof (cases y rule: P_Some) | |
| 118 | case (Some v) | |
| 119 | with \<open>y < Some x\<close> have "v < x" by simp | |
| 30662 | 120 | with less show "(P o Some) v" . | 
| 121 | qed | |
| 122 | qed | |
| 123 | then show ?case by simp | |
| 124 | qed | |
| 125 | qed | |
| 126 | qed | |
| 127 | ||
| 49190 | 128 | instantiation option :: (inf) inf | 
| 129 | begin | |
| 130 | ||
| 131 | definition inf_option where | |
| 132 | "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)))" | |
| 133 | ||
| 60679 | 134 | lemma inf_None_1 [simp, code]: "None \<sqinter> y = None" | 
| 49190 | 135 | by (simp add: inf_option_def) | 
| 136 | ||
| 60679 | 137 | lemma inf_None_2 [simp, code]: "x \<sqinter> None = None" | 
| 49190 | 138 | by (cases x) (simp_all add: inf_option_def) | 
| 139 | ||
| 60679 | 140 | lemma inf_Some [simp, code]: "Some x \<sqinter> Some y = Some (x \<sqinter> y)" | 
| 49190 | 141 | by (simp add: inf_option_def) | 
| 142 | ||
| 143 | instance .. | |
| 144 | ||
| 30662 | 145 | end | 
| 49190 | 146 | |
| 147 | instantiation option :: (sup) sup | |
| 148 | begin | |
| 149 | ||
| 150 | definition sup_option where | |
| 151 | "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)))" | |
| 152 | ||
| 60679 | 153 | lemma sup_None_1 [simp, code]: "None \<squnion> y = y" | 
| 49190 | 154 | by (simp add: sup_option_def) | 
| 155 | ||
| 60679 | 156 | lemma sup_None_2 [simp, code]: "x \<squnion> None = x" | 
| 49190 | 157 | by (cases x) (simp_all add: sup_option_def) | 
| 158 | ||
| 60679 | 159 | lemma sup_Some [simp, code]: "Some x \<squnion> Some y = Some (x \<squnion> y)" | 
| 49190 | 160 | by (simp add: sup_option_def) | 
| 161 | ||
| 162 | instance .. | |
| 163 | ||
| 164 | end | |
| 165 | ||
| 60679 | 166 | instance option :: (semilattice_inf) semilattice_inf | 
| 167 | proof | |
| 49190 | 168 | fix x y z :: "'a option" | 
| 169 | show "x \<sqinter> y \<le> x" | |
| 60679 | 170 | by (cases x, simp_all, cases y, simp_all) | 
| 49190 | 171 | show "x \<sqinter> y \<le> y" | 
| 60679 | 172 | by (cases x, simp_all, cases y, simp_all) | 
| 49190 | 173 | show "x \<le> y \<Longrightarrow> x \<le> z \<Longrightarrow> x \<le> y \<sqinter> z" | 
| 60679 | 174 | by (cases x, simp_all, cases y, simp_all, cases z, simp_all) | 
| 49190 | 175 | qed | 
| 176 | ||
| 60679 | 177 | instance option :: (semilattice_sup) semilattice_sup | 
| 178 | proof | |
| 49190 | 179 | fix x y z :: "'a option" | 
| 180 | show "x \<le> x \<squnion> y" | |
| 60679 | 181 | by (cases x, simp_all, cases y, simp_all) | 
| 49190 | 182 | show "y \<le> x \<squnion> y" | 
| 60679 | 183 | by (cases x, simp_all, cases y, simp_all) | 
| 49190 | 184 | fix x y z :: "'a option" | 
| 185 | show "y \<le> x \<Longrightarrow> z \<le> x \<Longrightarrow> y \<squnion> z \<le> x" | |
| 60679 | 186 | by (cases y, simp_all, cases z, simp_all, cases x, simp_all) | 
| 49190 | 187 | qed | 
| 188 | ||
| 189 | instance option :: (lattice) lattice .. | |
| 190 | ||
| 191 | instance option :: (lattice) bounded_lattice_bot .. | |
| 192 | ||
| 193 | instance option :: (bounded_lattice_top) bounded_lattice_top .. | |
| 194 | ||
| 195 | instance option :: (bounded_lattice_top) bounded_lattice .. | |
| 196 | ||
| 197 | instance option :: (distrib_lattice) distrib_lattice | |
| 198 | proof | |
| 199 | fix x y z :: "'a option" | |
| 200 | show "x \<squnion> y \<sqinter> z = (x \<squnion> y) \<sqinter> (x \<squnion> z)" | |
| 60679 | 201 | by (cases x, simp_all, cases y, simp_all, cases z, simp_all add: sup_inf_distrib1 inf_commute) | 
| 202 | qed | |
| 49190 | 203 | |
| 204 | instantiation option :: (complete_lattice) complete_lattice | |
| 205 | begin | |
| 206 | ||
| 207 | definition Inf_option :: "'a option set \<Rightarrow> 'a option" where | |
| 208 | "\<Sqinter>A = (if None \<in> A then None else Some (\<Sqinter>Option.these A))" | |
| 209 | ||
| 60679 | 210 | lemma None_in_Inf [simp]: "None \<in> A \<Longrightarrow> \<Sqinter>A = None" | 
| 49190 | 211 | by (simp add: Inf_option_def) | 
| 212 | ||
| 213 | definition Sup_option :: "'a option set \<Rightarrow> 'a option" where | |
| 214 |   "\<Squnion>A = (if A = {} \<or> A = {None} then None else Some (\<Squnion>Option.these A))"
 | |
| 215 | ||
| 60679 | 216 | lemma empty_Sup [simp]: "\<Squnion>{} = None"
 | 
| 49190 | 217 | by (simp add: Sup_option_def) | 
| 218 | ||
| 60679 | 219 | lemma singleton_None_Sup [simp]: "\<Squnion>{None} = None"
 | 
| 49190 | 220 | by (simp add: Sup_option_def) | 
| 221 | ||
| 60679 | 222 | instance | 
| 223 | proof | |
| 49190 | 224 | fix x :: "'a option" and A | 
| 225 | assume "x \<in> A" | |
| 226 | then show "\<Sqinter>A \<le> x" | |
| 227 | by (cases x) (auto simp add: Inf_option_def in_these_eq intro: Inf_lower) | |
| 228 | next | |
| 229 | fix z :: "'a option" and A | |
| 230 | assume *: "\<And>x. x \<in> A \<Longrightarrow> z \<le> x" | |
| 231 | show "z \<le> \<Sqinter>A" | |
| 232 | proof (cases z) | |
| 233 | case None then show ?thesis by simp | |
| 234 | next | |
| 235 | case (Some y) | |
| 236 | show ?thesis | |
| 237 | by (auto simp add: Inf_option_def in_these_eq Some intro!: Inf_greatest dest!: *) | |
| 238 | qed | |
| 239 | next | |
| 240 | fix x :: "'a option" and A | |
| 241 | assume "x \<in> A" | |
| 242 | then show "x \<le> \<Squnion>A" | |
| 243 | by (cases x) (auto simp add: Sup_option_def in_these_eq intro: Sup_upper) | |
| 244 | next | |
| 245 | fix z :: "'a option" and A | |
| 246 | assume *: "\<And>x. x \<in> A \<Longrightarrow> x \<le> z" | |
| 247 | show "\<Squnion>A \<le> z " | |
| 248 | proof (cases z) | |
| 249 | case None | |
| 250 | with * have "\<And>x. x \<in> A \<Longrightarrow> x = None" by (auto dest: less_eq_option_None_is_None) | |
| 251 |     then have "A = {} \<or> A = {None}" by blast
 | |
| 252 | then show ?thesis by (simp add: Sup_option_def) | |
| 253 | next | |
| 254 | case (Some y) | |
| 255 | from * have "\<And>w. Some w \<in> A \<Longrightarrow> Some w \<le> z" . | |
| 256 | with Some have "\<And>w. w \<in> Option.these A \<Longrightarrow> w \<le> y" | |
| 257 | by (simp add: in_these_eq) | |
| 258 | then have "\<Squnion>Option.these A \<le> y" by (rule Sup_least) | |
| 259 | with Some show ?thesis by (simp add: Sup_option_def) | |
| 260 | qed | |
| 52729 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 haftmann parents: 
49190diff
changeset | 261 | next | 
| 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 haftmann parents: 
49190diff
changeset | 262 |   show "\<Squnion>{} = (\<bottom>::'a option)"
 | 
| 60679 | 263 | by (auto simp: bot_option_def) | 
| 52729 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 haftmann parents: 
49190diff
changeset | 264 |   show "\<Sqinter>{} = (\<top>::'a option)"
 | 
| 60679 | 265 | by (auto simp: top_option_def Inf_option_def) | 
| 49190 | 266 | qed | 
| 267 | ||
| 268 | end | |
| 269 | ||
| 270 | lemma Some_Inf: | |
| 271 | "Some (\<Sqinter>A) = \<Sqinter>(Some ` A)" | |
| 272 | by (auto simp add: Inf_option_def) | |
| 273 | ||
| 274 | lemma Some_Sup: | |
| 275 |   "A \<noteq> {} \<Longrightarrow> Some (\<Squnion>A) = \<Squnion>(Some ` A)"
 | |
| 276 | by (auto simp add: Sup_option_def) | |
| 277 | ||
| 278 | lemma Some_INF: | |
| 279 | "Some (\<Sqinter>x\<in>A. f x) = (\<Sqinter>x\<in>A. Some (f x))" | |
| 56166 | 280 | using Some_Inf [of "f ` A"] by (simp add: comp_def) | 
| 49190 | 281 | |
| 282 | lemma Some_SUP: | |
| 283 |   "A \<noteq> {} \<Longrightarrow> Some (\<Squnion>x\<in>A. f x) = (\<Squnion>x\<in>A. Some (f x))"
 | |
| 56166 | 284 | using Some_Sup [of "f ` A"] by (simp add: comp_def) | 
| 49190 | 285 | |
| 60679 | 286 | instance option :: (complete_distrib_lattice) complete_distrib_lattice | 
| 287 | proof | |
| 49190 | 288 | fix a :: "'a option" and B | 
| 289 | show "a \<squnion> \<Sqinter>B = (\<Sqinter>b\<in>B. a \<squnion> b)" | |
| 290 | proof (cases a) | |
| 291 | case None | |
| 62343 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61955diff
changeset | 292 | then show ?thesis by simp | 
| 49190 | 293 | next | 
| 294 | case (Some c) | |
| 295 | show ?thesis | |
| 296 | proof (cases "None \<in> B") | |
| 297 | case True | |
| 298 | then have "Some c = (\<Sqinter>b\<in>B. Some c \<squnion> b)" | |
| 299 | by (auto intro!: antisym INF_lower2 INF_greatest) | |
| 300 | with True Some show ?thesis by simp | |
| 301 | next | |
| 302 |       case False then have B: "{x \<in> B. \<exists>y. x = Some y} = B" by auto (metis not_Some_eq)
 | |
| 303 | from sup_Inf have "Some c \<squnion> Some (\<Sqinter>Option.these B) = Some (\<Sqinter>b\<in>Option.these B. c \<squnion> b)" by simp | |
| 304 | then have "Some c \<squnion> \<Sqinter>(Some ` Option.these B) = (\<Sqinter>x\<in>Some ` Option.these B. Some c \<squnion> x)" | |
| 56166 | 305 | by (simp add: Some_INF Some_Inf comp_def) | 
| 62343 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61955diff
changeset | 306 | with Some B show ?thesis by (simp add: Some_image_these_eq cong del: strong_INF_cong) | 
| 49190 | 307 | qed | 
| 308 | qed | |
| 309 | show "a \<sqinter> \<Squnion>B = (\<Squnion>b\<in>B. a \<sqinter> b)" | |
| 310 | proof (cases a) | |
| 311 | case None | |
| 62343 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61955diff
changeset | 312 | then show ?thesis by (simp add: image_constant_conv bot_option_def cong del: strong_SUP_cong) | 
| 49190 | 313 | next | 
| 314 | case (Some c) | |
| 315 | show ?thesis | |
| 316 |     proof (cases "B = {} \<or> B = {None}")
 | |
| 317 | case True | |
| 56166 | 318 | then show ?thesis by auto | 
| 49190 | 319 | next | 
| 320 |       have B: "B = {x \<in> B. \<exists>y. x = Some y} \<union> {x \<in> B. x = None}"
 | |
| 321 | by auto | |
| 322 |       then have Sup_B: "\<Squnion>B = \<Squnion>({x \<in> B. \<exists>y. x = Some y} \<union> {x \<in> B. x = None})"
 | |
| 323 |         and SUP_B: "\<And>f. (\<Squnion>x \<in> B. f x) = (\<Squnion>x \<in> {x \<in> B. \<exists>y. x = Some y} \<union> {x \<in> B. x = None}. f x)"
 | |
| 324 | by simp_all | |
| 325 |       have Sup_None: "\<Squnion>{x. x = None \<and> x \<in> B} = None"
 | |
| 326 | by (simp add: bot_option_def [symmetric]) | |
| 327 |       have SUP_None: "(\<Squnion>x\<in>{x. x = None \<and> x \<in> B}. Some c \<sqinter> x) = None"
 | |
| 328 | by (simp add: bot_option_def [symmetric]) | |
| 329 |       case False then have "Option.these B \<noteq> {}" by (simp add: these_not_empty_eq)
 | |
| 330 | moreover from inf_Sup have "Some c \<sqinter> Some (\<Squnion>Option.these B) = Some (\<Squnion>b\<in>Option.these B. c \<sqinter> b)" | |
| 331 | by simp | |
| 332 | ultimately have "Some c \<sqinter> \<Squnion>(Some ` Option.these B) = (\<Squnion>x\<in>Some ` Option.these B. Some c \<sqinter> x)" | |
| 56166 | 333 | by (simp add: Some_SUP Some_Sup comp_def) | 
| 49190 | 334 | with Some show ?thesis | 
| 62343 
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
 haftmann parents: 
61955diff
changeset | 335 | by (simp add: Some_image_these_eq Sup_B SUP_B Sup_None SUP_None SUP_union Sup_union_distrib cong del: strong_SUP_cong) | 
| 49190 | 336 | qed | 
| 337 | qed | |
| 338 | qed | |
| 339 | ||
| 60679 | 340 | instance option :: (complete_linorder) complete_linorder .. | 
| 49190 | 341 | |
| 342 | ||
| 343 | no_notation | |
| 344 |   bot ("\<bottom>") and
 | |
| 345 |   top ("\<top>") and
 | |
| 346 | inf (infixl "\<sqinter>" 70) and | |
| 347 | sup (infixl "\<squnion>" 65) and | |
| 348 |   Inf  ("\<Sqinter>_" [900] 900) and
 | |
| 349 |   Sup  ("\<Squnion>_" [900] 900)
 | |
| 350 | ||
| 61955 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
60679diff
changeset | 351 | no_syntax | 
| 49190 | 352 |   "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Sqinter>_./ _)" [0, 10] 10)
 | 
| 353 |   "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10)
 | |
| 354 |   "_SUP1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Squnion>_./ _)" [0, 10] 10)
 | |
| 355 |   "_SUP"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10)
 | |
| 356 | ||
| 357 | end |