# HG changeset patch # User chaieb # Date 1236164087 0 # Node ID 478e5348ad3cad784e91bf9ed27aff8e37ef02cb # Parent 8c2649eb6a208966fe14c3a2a03c4b01b982e918# Parent 06b2d7f9f64b68fb02b3560eed29d96f48e64dd9 merged diff -r 06b2d7f9f64b -r 478e5348ad3c src/HOL/Library/Code_Index.thy --- a/src/HOL/Library/Code_Index.thy Wed Mar 04 10:33:14 2009 +0000 +++ b/src/HOL/Library/Code_Index.thy Wed Mar 04 10:54:47 2009 +0000 @@ -87,12 +87,14 @@ then show "P k" by simp qed simp_all -lemmas [code del] = index.recs index.cases - declare index_case [case_names nat, cases type: index] declare index.induct [case_names nat, induct type: index] -lemma [code]: +lemma index_decr [termination_simp]: + "k \ Code_Index.of_nat 0 \ Code_Index.nat_of k - Suc 0 < Code_Index.nat_of k" + by (cases k) simp + +lemma [simp, code]: "index_size = nat_of" proof (rule ext) fix k @@ -102,7 +104,7 @@ finally show "index_size k = nat_of k" . qed -lemma [code]: +lemma [simp, code]: "size = nat_of" proof (rule ext) fix k @@ -110,6 +112,8 @@ by (induct k) (simp_all del: zero_index_def Suc_index_def, simp_all) qed +lemmas [code del] = index.recs index.cases + lemma [code]: "eq_class.eq k l \ eq_class.eq (nat_of k) (nat_of l)" by (cases k, cases l) (simp add: eq) diff -r 06b2d7f9f64b -r 478e5348ad3c src/HOL/Option.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Option.thy Wed Mar 04 10:54:47 2009 +0000 @@ -0,0 +1,124 @@ +(* Title: HOL/Option.thy + Author: Folklore +*) + +header {* Datatype option *} + +theory Option +imports Datatype +begin + +datatype 'a option = None | Some 'a + +lemma not_None_eq [iff]: "(x ~= None) = (EX y. x = Some y)" + by (induct x) auto + +lemma not_Some_eq [iff]: "(ALL y. x ~= Some y) = (x = None)" + by (induct x) auto + +text{*Although it may appear that both of these equalities are helpful +only when applied to assumptions, in practice it seems better to give +them the uniform iff attribute. *} + +lemma option_caseE: + assumes c: "(case x of None => P | Some y => Q y)" + obtains + (None) "x = None" and P + | (Some) y where "x = Some y" and "Q y" + using c by (cases x) simp_all + +lemma insert_None_conv_UNIV: "insert None (range Some) = UNIV" + by (rule set_ext, case_tac x) auto + +lemma inj_Some [simp]: "inj_on Some A" + by (rule inj_onI) simp + + +subsubsection {* Operations *} + +primrec the :: "'a option => 'a" where +"the (Some x) = x" + +primrec set :: "'a option => 'a set" where +"set None = {}" | +"set (Some x) = {x}" + +lemma ospec [dest]: "(ALL x:set A. P x) ==> A = Some x ==> P x" + by simp + +declaration {* fn _ => + Classical.map_cs (fn cs => cs addSD2 ("ospec", thm "ospec")) +*} + +lemma elem_set [iff]: "(x : set xo) = (xo = Some x)" + by (cases xo) auto + +lemma set_empty_eq [simp]: "(set xo = {}) = (xo = None)" + by (cases xo) auto + +definition + map :: "('a \ 'b) \ 'a option \ 'b option" +where + [code del]: "map = (%f y. case y of None => None | Some x => Some (f x))" + +lemma option_map_None [simp, code]: "map f None = None" + by (simp add: map_def) + +lemma option_map_Some [simp, code]: "map f (Some x) = Some (f x)" + by (simp add: map_def) + +lemma option_map_is_None [iff]: + "(map f opt = None) = (opt = None)" + by (simp add: map_def split add: option.split) + +lemma option_map_eq_Some [iff]: + "(map f xo = Some y) = (EX z. xo = Some z & f z = y)" + by (simp add: map_def split add: option.split) + +lemma option_map_comp: + "map f (map g opt) = map (f o g) opt" + by (simp add: map_def split add: option.split) + +lemma option_map_o_sum_case [simp]: + "map f o sum_case g h = sum_case (map f o g) (map f o h)" + by (rule ext) (simp split: sum.split) + + +hide (open) const set map + +subsubsection {* Code generator setup *} + +definition + is_none :: "'a option \ bool" where + is_none_none [code post, symmetric, code inline]: "is_none x \ x = None" + +lemma is_none_code [code]: + shows "is_none None \ True" + and "is_none (Some x) \ False" + unfolding is_none_none [symmetric] by simp_all + +hide (open) const is_none + +code_type option + (SML "_ option") + (OCaml "_ option") + (Haskell "Maybe _") + +code_const None and Some + (SML "NONE" and "SOME") + (OCaml "None" and "Some _") + (Haskell "Nothing" and "Just") + +code_instance option :: eq + (Haskell -) + +code_const "eq_class.eq \ 'a\eq option \ 'a option \ bool" + (Haskell infixl 4 "==") + +code_reserved SML + option NONE SOME + +code_reserved OCaml + option None Some + +end diff -r 06b2d7f9f64b -r 478e5348ad3c src/Pure/axclass.ML --- a/src/Pure/axclass.ML Wed Mar 04 10:33:14 2009 +0000 +++ b/src/Pure/axclass.ML Wed Mar 04 10:54:47 2009 +0000 @@ -234,7 +234,10 @@ val map_inst_params = AxClassData.map o apsnd o apsnd; fun get_inst_param thy (c, tyco) = - (the o Symtab.lookup ((the o Symtab.lookup (fst (get_inst_params thy))) c)) tyco; + case Symtab.lookup ((the_default Symtab.empty o Symtab.lookup (fst (get_inst_params thy))) c) tyco + of SOME c' => c' + | NONE => error ("No instance parameter for constant " ^ quote c + ^ " on type constructor " ^ quote tyco); fun add_inst_param (c, tyco) inst = (map_inst_params o apfst o Symtab.map_default (c, Symtab.empty)) (Symtab.update_new (tyco, inst))