src/HOL/Option.thy
author haftmann
Sun, 22 Jul 2012 09:56:34 +0200
changeset 49442 571cb1df0768
parent 47397 c4cf9d03c352
child 50204 3f85cd15a0cc
permissions -rw-r--r--
library theories for debugging and parallel computing using code generation towards Isabelle/ML
     1 (*  Title:      HOL/Option.thy
     2     Author:     Folklore
     3 *)
     4 
     5 header {* Datatype option *}
     6 
     7 theory Option
     8 imports Datatype
     9 begin
    10 
    11 datatype 'a option = None | Some 'a
    12 
    13 lemma not_None_eq [iff]: "(x ~= None) = (EX y. x = Some y)"
    14   by (induct x) auto
    15 
    16 lemma not_Some_eq [iff]: "(ALL y. x ~= Some y) = (x = None)"
    17   by (induct x) auto
    18 
    19 text{*Although it may appear that both of these equalities are helpful
    20 only when applied to assumptions, in practice it seems better to give
    21 them the uniform iff attribute. *}
    22 
    23 lemma inj_Some [simp]: "inj_on Some A"
    24 by (rule inj_onI) simp
    25 
    26 lemma option_caseE:
    27   assumes c: "(case x of None => P | Some y => Q y)"
    28   obtains
    29     (None) "x = None" and P
    30   | (Some) y where "x = Some y" and "Q y"
    31   using c by (cases x) simp_all
    32 
    33 lemma UNIV_option_conv: "UNIV = insert None (range Some)"
    34 by(auto intro: classical)
    35 
    36 
    37 subsubsection {* Operations *}
    38 
    39 primrec the :: "'a option => 'a" where
    40 "the (Some x) = x"
    41 
    42 primrec set :: "'a option => 'a set" where
    43 "set None = {}" |
    44 "set (Some x) = {x}"
    45 
    46 lemma ospec [dest]: "(ALL x:set A. P x) ==> A = Some x ==> P x"
    47   by simp
    48 
    49 declaration {* fn _ =>
    50   Classical.map_cs (fn cs => cs addSD2 ("ospec", @{thm ospec}))
    51 *}
    52 
    53 lemma elem_set [iff]: "(x : set xo) = (xo = Some x)"
    54   by (cases xo) auto
    55 
    56 lemma set_empty_eq [simp]: "(set xo = {}) = (xo = None)"
    57   by (cases xo) auto
    58 
    59 definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a option \<Rightarrow> 'b option" where
    60   "map = (%f y. case y of None => None | Some x => Some (f x))"
    61 
    62 lemma option_map_None [simp, code]: "map f None = None"
    63   by (simp add: map_def)
    64 
    65 lemma option_map_Some [simp, code]: "map f (Some x) = Some (f x)"
    66   by (simp add: map_def)
    67 
    68 lemma option_map_is_None [iff]:
    69     "(map f opt = None) = (opt = None)"
    70   by (simp add: map_def split add: option.split)
    71 
    72 lemma option_map_eq_Some [iff]:
    73     "(map f xo = Some y) = (EX z. xo = Some z & f z = y)"
    74   by (simp add: map_def split add: option.split)
    75 
    76 lemma option_map_comp:
    77     "map f (map g opt) = map (f o g) opt"
    78   by (simp add: map_def split add: option.split)
    79 
    80 lemma option_map_o_sum_case [simp]:
    81     "map f o sum_case g h = sum_case (map f o g) (map f o h)"
    82   by (rule ext) (simp split: sum.split)
    83 
    84 lemma map_cong: "x = y \<Longrightarrow> (\<And>a. y = Some a \<Longrightarrow> f a = g a) \<Longrightarrow> map f x = map g y"
    85 by (cases x) auto
    86 
    87 enriched_type map: Option.map proof -
    88   fix f g
    89   show "Option.map f \<circ> Option.map g = Option.map (f \<circ> g)"
    90   proof
    91     fix x
    92     show "(Option.map f \<circ> Option.map g) x= Option.map (f \<circ> g) x"
    93       by (cases x) simp_all
    94   qed
    95 next
    96   show "Option.map id = id"
    97   proof
    98     fix x
    99     show "Option.map id x = id x"
   100       by (cases x) simp_all
   101   qed
   102 qed
   103 
   104 primrec bind :: "'a option \<Rightarrow> ('a \<Rightarrow> 'b option) \<Rightarrow> 'b option" where
   105 bind_lzero: "bind None f = None" |
   106 bind_lunit: "bind (Some x) f = f x"
   107 
   108 lemma bind_runit[simp]: "bind x Some = x"
   109 by (cases x) auto
   110 
   111 lemma bind_assoc[simp]: "bind (bind x f) g = bind x (\<lambda>y. bind (f y) g)"
   112 by (cases x) auto
   113 
   114 lemma bind_rzero[simp]: "bind x (\<lambda>x. None) = None"
   115 by (cases x) auto
   116 
   117 lemma bind_cong: "x = y \<Longrightarrow> (\<And>a. y = Some a \<Longrightarrow> f a = g a) \<Longrightarrow> bind x f = bind y g"
   118 by (cases x) auto
   119 
   120 hide_const (open) set map bind
   121 hide_fact (open) map_cong bind_cong
   122 
   123 subsubsection {* Code generator setup *}
   124 
   125 definition is_none :: "'a option \<Rightarrow> bool" where
   126   [code_post]: "is_none x \<longleftrightarrow> x = None"
   127 
   128 lemma is_none_code [code]:
   129   shows "is_none None \<longleftrightarrow> True"
   130     and "is_none (Some x) \<longleftrightarrow> False"
   131   unfolding is_none_def by simp_all
   132 
   133 lemma [code_unfold]:
   134   "HOL.equal x None \<longleftrightarrow> is_none x"
   135   by (simp add: equal is_none_def)
   136 
   137 hide_const (open) is_none
   138 
   139 code_type option
   140   (SML "_ option")
   141   (OCaml "_ option")
   142   (Haskell "Maybe _")
   143   (Scala "!Option[(_)]")
   144 
   145 code_const None and Some
   146   (SML "NONE" and "SOME")
   147   (OCaml "None" and "Some _")
   148   (Haskell "Nothing" and "Just")
   149   (Scala "!None" and "Some")
   150 
   151 code_instance option :: equal
   152   (Haskell -)
   153 
   154 code_const "HOL.equal \<Colon> 'a option \<Rightarrow> 'a option \<Rightarrow> bool"
   155   (Haskell infix 4 "==")
   156 
   157 code_reserved SML
   158   option NONE SOME
   159 
   160 code_reserved OCaml
   161   option None Some
   162 
   163 code_reserved Scala
   164   Option None Some
   165 
   166 end