haftmann@26241: (* Title: HOL/Library/Option_ord.thy haftmann@26241: Author: Florian Haftmann, TU Muenchen haftmann@26241: *) haftmann@26241: haftmann@26263: header {* Canonical order on option type *} haftmann@26241: haftmann@26241: theory Option_ord haftmann@30662: imports Option Main haftmann@26241: begin haftmann@26241: haftmann@49190: notation haftmann@49190: bot ("\") and haftmann@49190: top ("\") and haftmann@49190: inf (infixl "\" 70) and haftmann@49190: sup (infixl "\" 65) and haftmann@49190: Inf ("\_" [900] 900) and haftmann@49190: Sup ("\_" [900] 900) haftmann@49190: haftmann@49190: syntax (xsymbols) haftmann@49190: "_INF1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) haftmann@49190: "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) haftmann@49190: "_SUP1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) haftmann@49190: "_SUP" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) haftmann@49190: haftmann@49190: haftmann@30662: instantiation option :: (preorder) preorder haftmann@26241: begin haftmann@26241: haftmann@26241: definition less_eq_option where haftmann@37765: "x \ y \ (case x of None \ True | Some x \ (case y of None \ False | Some y \ x \ y))" haftmann@26241: haftmann@26241: definition less_option where haftmann@37765: "x < y \ (case y of None \ False | Some y \ (case x of None \ True | Some x \ x < y))" haftmann@26241: haftmann@26258: lemma less_eq_option_None [simp]: "None \ x" haftmann@26241: by (simp add: less_eq_option_def) haftmann@26241: haftmann@26258: lemma less_eq_option_None_code [code]: "None \ x \ True" haftmann@26241: by simp haftmann@26241: haftmann@26258: lemma less_eq_option_None_is_None: "x \ None \ x = None" haftmann@26241: by (cases x) (simp_all add: less_eq_option_def) haftmann@26241: haftmann@26258: lemma less_eq_option_Some_None [simp, code]: "Some x \ None \ False" haftmann@26241: by (simp add: less_eq_option_def) haftmann@26241: haftmann@26258: lemma less_eq_option_Some [simp, code]: "Some x \ Some y \ x \ y" haftmann@26241: by (simp add: less_eq_option_def) haftmann@26241: haftmann@26258: lemma less_option_None [simp, code]: "x < None \ False" haftmann@26241: by (simp add: less_option_def) haftmann@26241: haftmann@26258: lemma less_option_None_is_Some: "None < x \ \z. x = Some z" haftmann@26241: by (cases x) (simp_all add: less_option_def) haftmann@26241: haftmann@26258: lemma less_option_None_Some [simp]: "None < Some x" haftmann@26241: by (simp add: less_option_def) haftmann@26241: haftmann@26258: lemma less_option_None_Some_code [code]: "None < Some x \ True" haftmann@26241: by simp haftmann@26241: haftmann@26258: lemma less_option_Some [simp, code]: "Some x < Some y \ x < y" haftmann@26241: by (simp add: less_option_def) haftmann@26241: haftmann@30662: instance proof haftmann@30662: qed (auto simp add: less_eq_option_def less_option_def less_le_not_le elim: order_trans split: option.splits) haftmann@26241: haftmann@26241: end haftmann@26241: haftmann@30662: instance option :: (order) order proof haftmann@30662: qed (auto simp add: less_eq_option_def less_option_def split: option.splits) haftmann@30662: haftmann@30662: instance option :: (linorder) linorder proof haftmann@30662: qed (auto simp add: less_eq_option_def less_option_def split: option.splits) haftmann@30662: haftmann@43815: instantiation option :: (order) bot haftmann@30662: begin haftmann@30662: haftmann@49190: definition bot_option where haftmann@49190: "\ = None" haftmann@30662: haftmann@30662: instance proof haftmann@30662: qed (simp add: bot_option_def) haftmann@30662: haftmann@30662: end haftmann@30662: haftmann@30662: instantiation option :: (top) top haftmann@30662: begin haftmann@30662: haftmann@49190: definition top_option where haftmann@49190: "\ = Some \" haftmann@30662: haftmann@30662: instance proof haftmann@30662: qed (simp add: top_option_def less_eq_option_def split: option.split) haftmann@26241: haftmann@26241: end haftmann@30662: haftmann@30662: instance option :: (wellorder) wellorder proof haftmann@30662: fix P :: "'a option \ bool" and z :: "'a option" haftmann@30662: assume H: "\x. (\y. y < x \ P y) \ P x" haftmann@30662: have "P None" by (rule H) simp haftmann@30662: then have P_Some [case_names Some]: haftmann@30662: "\z. (\x. z = Some x \ (P o Some) x) \ P z" haftmann@30662: proof - haftmann@30662: fix z haftmann@30662: assume "\x. z = Some x \ (P o Some) x" haftmann@30662: with `P None` show "P z" by (cases z) simp_all haftmann@30662: qed haftmann@30662: show "P z" proof (cases z rule: P_Some) haftmann@30662: case (Some w) haftmann@30662: show "(P o Some) w" proof (induct rule: less_induct) haftmann@30662: case (less x) haftmann@30662: have "P (Some x)" proof (rule H) haftmann@30662: fix y :: "'a option" haftmann@30662: assume "y < Some x" haftmann@30662: show "P y" proof (cases y rule: P_Some) haftmann@30662: case (Some v) with `y < Some x` have "v < x" by simp haftmann@30662: with less show "(P o Some) v" . haftmann@30662: qed haftmann@30662: qed haftmann@30662: then show ?case by simp haftmann@30662: qed haftmann@30662: qed haftmann@30662: qed haftmann@30662: haftmann@49190: instantiation option :: (inf) inf haftmann@49190: begin haftmann@49190: haftmann@49190: definition inf_option where haftmann@49190: "x \ y = (case x of None \ None | Some x \ (case y of None \ None | Some y \ Some (x \ y)))" haftmann@49190: haftmann@49190: lemma inf_None_1 [simp, code]: haftmann@49190: "None \ y = None" haftmann@49190: by (simp add: inf_option_def) haftmann@49190: haftmann@49190: lemma inf_None_2 [simp, code]: haftmann@49190: "x \ None = None" haftmann@49190: by (cases x) (simp_all add: inf_option_def) haftmann@49190: haftmann@49190: lemma inf_Some [simp, code]: haftmann@49190: "Some x \ Some y = Some (x \ y)" haftmann@49190: by (simp add: inf_option_def) haftmann@49190: haftmann@49190: instance .. haftmann@49190: haftmann@30662: end haftmann@49190: haftmann@49190: instantiation option :: (sup) sup haftmann@49190: begin haftmann@49190: haftmann@49190: definition sup_option where haftmann@49190: "x \ y = (case x of None \ y | Some x' \ (case y of None \ x | Some y \ Some (x' \ y)))" haftmann@49190: haftmann@49190: lemma sup_None_1 [simp, code]: haftmann@49190: "None \ y = y" haftmann@49190: by (simp add: sup_option_def) haftmann@49190: haftmann@49190: lemma sup_None_2 [simp, code]: haftmann@49190: "x \ None = x" haftmann@49190: by (cases x) (simp_all add: sup_option_def) haftmann@49190: haftmann@49190: lemma sup_Some [simp, code]: haftmann@49190: "Some x \ Some y = Some (x \ y)" haftmann@49190: by (simp add: sup_option_def) haftmann@49190: haftmann@49190: instance .. haftmann@49190: haftmann@49190: end haftmann@49190: haftmann@49190: instantiation option :: (semilattice_inf) semilattice_inf haftmann@49190: begin haftmann@49190: haftmann@49190: instance proof haftmann@49190: fix x y z :: "'a option" haftmann@49190: show "x \ y \ x" haftmann@49190: by - (cases x, simp_all, cases y, simp_all) haftmann@49190: show "x \ y \ y" haftmann@49190: by - (cases x, simp_all, cases y, simp_all) haftmann@49190: show "x \ y \ x \ z \ x \ y \ z" haftmann@49190: by - (cases x, simp_all, cases y, simp_all, cases z, simp_all) haftmann@49190: qed haftmann@49190: haftmann@49190: end haftmann@49190: haftmann@49190: instantiation option :: (semilattice_sup) semilattice_sup haftmann@49190: begin haftmann@49190: haftmann@49190: instance proof haftmann@49190: fix x y z :: "'a option" haftmann@49190: show "x \ x \ y" haftmann@49190: by - (cases x, simp_all, cases y, simp_all) haftmann@49190: show "y \ x \ y" haftmann@49190: by - (cases x, simp_all, cases y, simp_all) haftmann@49190: fix x y z :: "'a option" haftmann@49190: show "y \ x \ z \ x \ y \ z \ x" haftmann@49190: by - (cases y, simp_all, cases z, simp_all, cases x, simp_all) haftmann@49190: qed haftmann@49190: haftmann@49190: end haftmann@49190: haftmann@49190: instance option :: (lattice) lattice .. haftmann@49190: haftmann@49190: instance option :: (lattice) bounded_lattice_bot .. haftmann@49190: haftmann@49190: instance option :: (bounded_lattice_top) bounded_lattice_top .. haftmann@49190: haftmann@49190: instance option :: (bounded_lattice_top) bounded_lattice .. haftmann@49190: haftmann@49190: instance option :: (distrib_lattice) distrib_lattice haftmann@49190: proof haftmann@49190: fix x y z :: "'a option" haftmann@49190: show "x \ y \ z = (x \ y) \ (x \ z)" haftmann@49190: by - (cases x, simp_all, cases y, simp_all, cases z, simp_all add: sup_inf_distrib1 inf_commute) haftmann@49190: qed haftmann@49190: haftmann@49190: instantiation option :: (complete_lattice) complete_lattice haftmann@49190: begin haftmann@49190: haftmann@49190: definition Inf_option :: "'a option set \ 'a option" where haftmann@49190: "\A = (if None \ A then None else Some (\Option.these A))" haftmann@49190: haftmann@49190: lemma None_in_Inf [simp]: haftmann@49190: "None \ A \ \A = None" haftmann@49190: by (simp add: Inf_option_def) haftmann@49190: haftmann@49190: definition Sup_option :: "'a option set \ 'a option" where haftmann@49190: "\A = (if A = {} \ A = {None} then None else Some (\Option.these A))" haftmann@49190: haftmann@49190: lemma empty_Sup [simp]: haftmann@49190: "\{} = None" haftmann@49190: by (simp add: Sup_option_def) haftmann@49190: haftmann@49190: lemma singleton_None_Sup [simp]: haftmann@49190: "\{None} = None" haftmann@49190: by (simp add: Sup_option_def) haftmann@49190: haftmann@49190: instance proof haftmann@49190: fix x :: "'a option" and A haftmann@49190: assume "x \ A" haftmann@49190: then show "\A \ x" haftmann@49190: by (cases x) (auto simp add: Inf_option_def in_these_eq intro: Inf_lower) haftmann@49190: next haftmann@49190: fix z :: "'a option" and A haftmann@49190: assume *: "\x. x \ A \ z \ x" haftmann@49190: show "z \ \A" haftmann@49190: proof (cases z) haftmann@49190: case None then show ?thesis by simp haftmann@49190: next haftmann@49190: case (Some y) haftmann@49190: show ?thesis haftmann@49190: by (auto simp add: Inf_option_def in_these_eq Some intro!: Inf_greatest dest!: *) haftmann@49190: qed haftmann@49190: next haftmann@49190: fix x :: "'a option" and A haftmann@49190: assume "x \ A" haftmann@49190: then show "x \ \A" haftmann@49190: by (cases x) (auto simp add: Sup_option_def in_these_eq intro: Sup_upper) haftmann@49190: next haftmann@49190: fix z :: "'a option" and A haftmann@49190: assume *: "\x. x \ A \ x \ z" haftmann@49190: show "\A \ z " haftmann@49190: proof (cases z) haftmann@49190: case None haftmann@49190: with * have "\x. x \ A \ x = None" by (auto dest: less_eq_option_None_is_None) haftmann@49190: then have "A = {} \ A = {None}" by blast haftmann@49190: then show ?thesis by (simp add: Sup_option_def) haftmann@49190: next haftmann@49190: case (Some y) haftmann@49190: from * have "\w. Some w \ A \ Some w \ z" . haftmann@49190: with Some have "\w. w \ Option.these A \ w \ y" haftmann@49190: by (simp add: in_these_eq) haftmann@49190: then have "\Option.these A \ y" by (rule Sup_least) haftmann@49190: with Some show ?thesis by (simp add: Sup_option_def) haftmann@49190: qed haftmann@49190: qed haftmann@49190: haftmann@49190: end haftmann@49190: haftmann@49190: lemma Some_Inf: haftmann@49190: "Some (\A) = \(Some ` A)" haftmann@49190: by (auto simp add: Inf_option_def) haftmann@49190: haftmann@49190: lemma Some_Sup: haftmann@49190: "A \ {} \ Some (\A) = \(Some ` A)" haftmann@49190: by (auto simp add: Sup_option_def) haftmann@49190: haftmann@49190: lemma Some_INF: haftmann@49190: "Some (\x\A. f x) = (\x\A. Some (f x))" haftmann@49190: by (simp add: INF_def Some_Inf image_image) haftmann@49190: haftmann@49190: lemma Some_SUP: haftmann@49190: "A \ {} \ Some (\x\A. f x) = (\x\A. Some (f x))" haftmann@49190: by (simp add: SUP_def Some_Sup image_image) haftmann@49190: haftmann@49190: instantiation option :: (complete_distrib_lattice) complete_distrib_lattice haftmann@49190: begin haftmann@49190: haftmann@49190: instance proof haftmann@49190: fix a :: "'a option" and B haftmann@49190: show "a \ \B = (\b\B. a \ b)" haftmann@49190: proof (cases a) haftmann@49190: case None haftmann@49190: then show ?thesis by (simp add: INF_def) haftmann@49190: next haftmann@49190: case (Some c) haftmann@49190: show ?thesis haftmann@49190: proof (cases "None \ B") haftmann@49190: case True haftmann@49190: then have "Some c = (\b\B. Some c \ b)" haftmann@49190: by (auto intro!: antisym INF_lower2 INF_greatest) haftmann@49190: with True Some show ?thesis by simp haftmann@49190: next haftmann@49190: case False then have B: "{x \ B. \y. x = Some y} = B" by auto (metis not_Some_eq) haftmann@49190: from sup_Inf have "Some c \ Some (\Option.these B) = Some (\b\Option.these B. c \ b)" by simp haftmann@49190: then have "Some c \ \(Some ` Option.these B) = (\x\Some ` Option.these B. Some c \ x)" haftmann@49190: by (simp add: Some_INF Some_Inf) haftmann@49190: with Some B show ?thesis by (simp add: Some_image_these_eq) haftmann@49190: qed haftmann@49190: qed haftmann@49190: show "a \ \B = (\b\B. a \ b)" haftmann@49190: proof (cases a) haftmann@49190: case None haftmann@49190: then show ?thesis by (simp add: SUP_def image_constant_conv bot_option_def) haftmann@49190: next haftmann@49190: case (Some c) haftmann@49190: show ?thesis haftmann@49190: proof (cases "B = {} \ B = {None}") haftmann@49190: case True haftmann@49190: then show ?thesis by (auto simp add: SUP_def) haftmann@49190: next haftmann@49190: have B: "B = {x \ B. \y. x = Some y} \ {x \ B. x = None}" haftmann@49190: by auto haftmann@49190: then have Sup_B: "\B = \({x \ B. \y. x = Some y} \ {x \ B. x = None})" haftmann@49190: and SUP_B: "\f. (\x \ B. f x) = (\x \ {x \ B. \y. x = Some y} \ {x \ B. x = None}. f x)" haftmann@49190: by simp_all haftmann@49190: have Sup_None: "\{x. x = None \ x \ B} = None" haftmann@49190: by (simp add: bot_option_def [symmetric]) haftmann@49190: have SUP_None: "(\x\{x. x = None \ x \ B}. Some c \ x) = None" haftmann@49190: by (simp add: bot_option_def [symmetric]) haftmann@49190: case False then have "Option.these B \ {}" by (simp add: these_not_empty_eq) haftmann@49190: moreover from inf_Sup have "Some c \ Some (\Option.these B) = Some (\b\Option.these B. c \ b)" haftmann@49190: by simp haftmann@49190: ultimately have "Some c \ \(Some ` Option.these B) = (\x\Some ` Option.these B. Some c \ x)" haftmann@49190: by (simp add: Some_SUP Some_Sup) haftmann@49190: with Some show ?thesis haftmann@49190: by (simp add: Some_image_these_eq Sup_B SUP_B Sup_None SUP_None SUP_union Sup_union_distrib) haftmann@49190: qed haftmann@49190: qed haftmann@49190: qed haftmann@49190: haftmann@49190: end haftmann@49190: haftmann@49190: instantiation option :: (complete_linorder) complete_linorder haftmann@49190: begin haftmann@49190: haftmann@49190: instance .. haftmann@49190: haftmann@49190: end haftmann@49190: haftmann@49190: haftmann@49190: no_notation haftmann@49190: bot ("\") and haftmann@49190: top ("\") and haftmann@49190: inf (infixl "\" 70) and haftmann@49190: sup (infixl "\" 65) and haftmann@49190: Inf ("\_" [900] 900) and haftmann@49190: Sup ("\_" [900] 900) haftmann@49190: haftmann@49190: no_syntax (xsymbols) haftmann@49190: "_INF1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) haftmann@49190: "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) haftmann@49190: "_SUP1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) haftmann@49190: "_SUP" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) haftmann@49190: haftmann@49190: end haftmann@49190: