nipkow@30246: (* Title: HOL/Option.thy nipkow@30246: Author: Folklore nipkow@30246: *) nipkow@30246: nipkow@30246: header {* Datatype option *} nipkow@30246: nipkow@30246: theory Option haftmann@35719: imports Datatype nipkow@30246: begin nipkow@30246: nipkow@30246: datatype 'a option = None | Some 'a nipkow@30246: nipkow@30246: lemma not_None_eq [iff]: "(x ~= None) = (EX y. x = Some y)" nipkow@30246: by (induct x) auto nipkow@30246: nipkow@30246: lemma not_Some_eq [iff]: "(ALL y. x ~= Some y) = (x = None)" nipkow@30246: by (induct x) auto nipkow@30246: nipkow@30246: text{*Although it may appear that both of these equalities are helpful nipkow@30246: only when applied to assumptions, in practice it seems better to give nipkow@30246: them the uniform iff attribute. *} nipkow@30246: nipkow@31080: lemma inj_Some [simp]: "inj_on Some A" nipkow@31080: by (rule inj_onI) simp nipkow@31080: nipkow@30246: lemma option_caseE: nipkow@30246: assumes c: "(case x of None => P | Some y => Q y)" nipkow@30246: obtains nipkow@30246: (None) "x = None" and P nipkow@30246: | (Some) y where "x = Some y" and "Q y" nipkow@30246: using c by (cases x) simp_all nipkow@30246: kuncar@53010: lemma split_option_all: "(\x. P x) \ P None \ (\x. P (Some x))" kuncar@53010: by (auto intro: option.induct) kuncar@53010: kuncar@53010: lemma split_option_ex: "(\x. P x) \ P None \ (\x. P (Some x))" kuncar@53010: using split_option_all[of "\x. \P x"] by blast kuncar@53010: nipkow@31080: lemma UNIV_option_conv: "UNIV = insert None (range Some)" nipkow@31080: by(auto intro: classical) nipkow@31080: nipkow@30246: subsubsection {* Operations *} nipkow@30246: nipkow@30246: primrec the :: "'a option => 'a" where nipkow@30246: "the (Some x) = x" nipkow@30246: nipkow@30246: primrec set :: "'a option => 'a set" where nipkow@30246: "set None = {}" | nipkow@30246: "set (Some x) = {x}" nipkow@30246: nipkow@30246: lemma ospec [dest]: "(ALL x:set A. P x) ==> A = Some x ==> P x" nipkow@30246: by simp nipkow@30246: wenzelm@51703: setup {* map_theory_claset (fn ctxt => ctxt addSD2 ("ospec", @{thm ospec})) *} nipkow@30246: nipkow@30246: lemma elem_set [iff]: "(x : set xo) = (xo = Some x)" nipkow@30246: by (cases xo) auto nipkow@30246: nipkow@30246: lemma set_empty_eq [simp]: "(set xo = {}) = (xo = None)" nipkow@30246: by (cases xo) auto nipkow@30246: haftmann@31154: definition map :: "('a \ 'b) \ 'a option \ 'b option" where haftmann@31154: "map = (%f y. case y of None => None | Some x => Some (f x))" nipkow@30246: nipkow@30246: lemma option_map_None [simp, code]: "map f None = None" nipkow@30246: by (simp add: map_def) nipkow@30246: nipkow@30246: lemma option_map_Some [simp, code]: "map f (Some x) = Some (f x)" nipkow@30246: by (simp add: map_def) nipkow@30246: nipkow@30246: lemma option_map_is_None [iff]: nipkow@30246: "(map f opt = None) = (opt = None)" nipkow@30246: by (simp add: map_def split add: option.split) nipkow@30246: nipkow@30246: lemma option_map_eq_Some [iff]: nipkow@30246: "(map f xo = Some y) = (EX z. xo = Some z & f z = y)" nipkow@30246: by (simp add: map_def split add: option.split) nipkow@30246: nipkow@30246: lemma option_map_comp: nipkow@30246: "map f (map g opt) = map (f o g) opt" nipkow@30246: by (simp add: map_def split add: option.split) nipkow@30246: nipkow@30246: lemma option_map_o_sum_case [simp]: nipkow@30246: "map f o sum_case g h = sum_case (map f o g) (map f o h)" nipkow@30246: by (rule ext) (simp split: sum.split) nipkow@30246: krauss@46526: lemma map_cong: "x = y \ (\a. y = Some a \ f a = g a) \ map f x = map g y" krauss@46526: by (cases x) auto krauss@46526: haftmann@41505: enriched_type map: Option.map proof - haftmann@41372: fix f g haftmann@41372: show "Option.map f \ Option.map g = Option.map (f \ g)" haftmann@41372: proof haftmann@41372: fix x haftmann@41372: show "(Option.map f \ Option.map g) x= Option.map (f \ g) x" haftmann@41372: by (cases x) simp_all haftmann@41372: qed haftmann@40609: next haftmann@41372: show "Option.map id = id" haftmann@41372: proof haftmann@41372: fix x haftmann@41372: show "Option.map id x = id x" haftmann@41372: by (cases x) simp_all haftmann@41372: qed haftmann@40609: qed haftmann@40609: haftmann@51096: lemma option_case_map [simp]: haftmann@51096: "option_case g h (Option.map f x) = option_case g (h \ f) x" haftmann@51096: by (cases x) simp_all haftmann@51096: krauss@39149: primrec bind :: "'a option \ ('a \ 'b option) \ 'b option" where krauss@39149: bind_lzero: "bind None f = None" | krauss@39149: bind_lunit: "bind (Some x) f = f x" nipkow@30246: krauss@39149: lemma bind_runit[simp]: "bind x Some = x" krauss@39149: by (cases x) auto krauss@39149: krauss@39149: lemma bind_assoc[simp]: "bind (bind x f) g = bind x (\y. bind (f y) g)" krauss@39149: by (cases x) auto krauss@39149: krauss@39149: lemma bind_rzero[simp]: "bind x (\x. None) = None" krauss@39149: by (cases x) auto krauss@39149: krauss@46526: lemma bind_cong: "x = y \ (\a. y = Some a \ f a = g a) \ bind x f = bind y g" krauss@46526: by (cases x) auto krauss@46526: haftmann@49189: definition these :: "'a option set \ 'a set" haftmann@49189: where haftmann@49189: "these A = the ` {x \ A. x \ None}" haftmann@49189: haftmann@49189: lemma these_empty [simp]: haftmann@49189: "these {} = {}" haftmann@49189: by (simp add: these_def) haftmann@49189: haftmann@49189: lemma these_insert_None [simp]: haftmann@49189: "these (insert None A) = these A" haftmann@49189: by (auto simp add: these_def) haftmann@49189: haftmann@49189: lemma these_insert_Some [simp]: haftmann@49189: "these (insert (Some x) A) = insert x (these A)" haftmann@49189: proof - haftmann@49189: have "{y \ insert (Some x) A. y \ None} = insert (Some x) {y \ A. y \ None}" haftmann@49189: by auto haftmann@49189: then show ?thesis by (simp add: these_def) haftmann@49189: qed haftmann@49189: haftmann@49189: lemma in_these_eq: haftmann@49189: "x \ these A \ Some x \ A" haftmann@49189: proof haftmann@49189: assume "Some x \ A" haftmann@49189: then obtain B where "A = insert (Some x) B" by auto haftmann@49189: then show "x \ these A" by (auto simp add: these_def intro!: image_eqI) haftmann@49189: next haftmann@49189: assume "x \ these A" haftmann@49189: then show "Some x \ A" by (auto simp add: these_def) haftmann@49189: qed haftmann@49189: haftmann@49189: lemma these_image_Some_eq [simp]: haftmann@49189: "these (Some ` A) = A" haftmann@49189: by (auto simp add: these_def intro!: image_eqI) haftmann@49189: haftmann@49189: lemma Some_image_these_eq: haftmann@49189: "Some ` these A = {x\A. x \ None}" haftmann@49189: by (auto simp add: these_def image_image intro!: image_eqI) haftmann@49189: haftmann@49189: lemma these_empty_eq: haftmann@49189: "these B = {} \ B = {} \ B = {None}" haftmann@49189: by (auto simp add: these_def) haftmann@49189: haftmann@49189: lemma these_not_empty_eq: haftmann@49189: "these B \ {} \ B \ {} \ B \ {None}" haftmann@49189: by (auto simp add: these_empty_eq) haftmann@49189: haftmann@49189: hide_const (open) set map bind these krauss@46526: hide_fact (open) map_cong bind_cong nipkow@30246: haftmann@49189: nipkow@30246: subsubsection {* Code generator setup *} nipkow@30246: haftmann@31154: definition is_none :: "'a option \ bool" where haftmann@31998: [code_post]: "is_none x \ x = None" nipkow@30246: nipkow@30246: lemma is_none_code [code]: nipkow@30246: shows "is_none None \ True" nipkow@30246: and "is_none (Some x) \ False" haftmann@31154: unfolding is_none_def by simp_all haftmann@31154: haftmann@32069: lemma [code_unfold]: haftmann@38857: "HOL.equal x None \ is_none x" lammich@53940: "HOL.equal None = is_none" lammich@53940: by (auto simp add: equal is_none_def) nipkow@30246: wenzelm@36176: hide_const (open) is_none nipkow@30246: haftmann@52435: code_printing haftmann@52435: type_constructor option \ haftmann@52435: (SML) "_ option" haftmann@52435: and (OCaml) "_ option" haftmann@52435: and (Haskell) "Maybe _" haftmann@52435: and (Scala) "!Option[(_)]" haftmann@52435: | constant None \ haftmann@52435: (SML) "NONE" haftmann@52435: and (OCaml) "None" haftmann@52435: and (Haskell) "Nothing" haftmann@52435: and (Scala) "!None" haftmann@52435: | constant Some \ haftmann@52435: (SML) "SOME" haftmann@52435: and (OCaml) "Some _" haftmann@52435: and (Haskell) "Just" haftmann@52435: and (Scala) "Some" haftmann@52435: | class_instance option :: equal \ haftmann@52435: (Haskell) - haftmann@52435: | constant "HOL.equal :: 'a option \ 'a option \ bool" \ haftmann@52435: (Haskell) infix 4 "==" nipkow@30246: nipkow@30246: code_reserved SML nipkow@30246: option NONE SOME nipkow@30246: nipkow@30246: code_reserved OCaml nipkow@30246: option None Some nipkow@30246: haftmann@34886: code_reserved Scala haftmann@34886: Option None Some haftmann@34886: nipkow@30246: end haftmann@49189: