# HG changeset patch # User Andreas Lochbihler # Date 1423662310 -3600 # Node ID 0c5c5ad5d2e009f3cd3e2fef62062e31034c4d09 # Parent ef8ac8d2315e13d70b87cfdb5d6c7b720101b181 add lemmas about functions on option diff -r ef8ac8d2315e -r 0c5c5ad5d2e0 src/HOL/Option.thy --- a/src/HOL/Option.thy Wed Feb 11 14:19:46 2015 +0100 +++ b/src/HOL/Option.thy Wed Feb 11 14:45:10 2015 +0100 @@ -62,6 +62,21 @@ lemma UNIV_option_conv: "UNIV = insert None (range Some)" by(auto intro: classical) +lemma rel_option_None1 [simp]: "rel_option P None x \ x = None" +by(cases x) simp_all + +lemma rel_option_None2 [simp]: "rel_option P x None \ x = None" +by(cases x) simp_all + +lemma rel_option_inf: "inf (rel_option A) (rel_option B) = rel_option (inf A B)" (is "?lhs = ?rhs") +proof(rule antisym) + show "?lhs \ ?rhs" by(auto elim!: option.rel_cases) +qed(auto elim: option.rel_mono_strong) + +lemma rel_option_reflI: + "(\x. x \ set_option y \ P x x) \ rel_option P y y" +by(cases y) auto + subsubsection {* Operations *} lemma ospec [dest]: "(ALL x:set_option A. P x) ==> A = Some x ==> P x" @@ -106,10 +121,43 @@ | _ \ False)" by (auto split: prod.split option.split) +definition is_none :: "'a option \ bool" +where [code_post]: "is_none x \ x = None" + +lemma is_none_simps [simp]: + "is_none None" + "\ is_none (Some x)" +by(simp_all add: is_none_def) + +lemma is_none_code [code]: + "is_none None = True" + "is_none (Some x) = False" +by simp_all + +lemma rel_option_unfold: + "rel_option R x y \ + (is_none x \ is_none y) \ (\ is_none x \ \ is_none y \ R (the x) (the y))" +by(simp add: rel_option_iff split: option.split) + +lemma rel_optionI: + "\ is_none x \ is_none y; \ \ is_none x; \ is_none y \ \ P (the x) (the y) \ + \ rel_option P x y" +by(simp add: rel_option_unfold) + +lemma is_none_map_option [simp]: "is_none (map_option f x) \ is_none x" +by(simp add: is_none_def) + +lemma the_map_option: "\ is_none x \ the (map_option f x) = f (the x)" +by(clarsimp simp add: is_none_def) + + primrec bind :: "'a option \ ('a \ 'b option) \ 'b option" where bind_lzero: "bind None f = None" | bind_lunit: "bind (Some x) f = f x" +lemma is_none_bind: "is_none (bind f g) \ is_none f \ is_none (g (the f))" +by(cases f) simp_all + lemma bind_runit[simp]: "bind x Some = x" by (cases x) auto @@ -133,6 +181,24 @@ lemmas bind_splits = bind_split bind_split_asm +lemma bind_eq_Some_conv: "bind f g = Some x \ (\y. f = Some y \ g y = Some x)" +by(cases f) simp_all + +lemma map_option_bind: "map_option f (bind x g) = bind x (map_option f \ g)" +by(cases x) simp_all + +lemma bind_option_cong: + "\ x = y; \z. z \ set_option y \ f z = g z \ \ bind x f = bind y g" +by(cases y) simp_all + +lemma bind_option_cong_simp: + "\ x = y; \z. z \ set_option y =simp=> f z = g z \ \ bind x f = bind y g" +unfolding simp_implies_def by(rule bind_option_cong) + +lemma bind_option_cong_code: "x = y \ bind x f = bind y f" by simp +setup \Code_Simp.map_ss (Simplifier.add_cong @{thm bind_option_cong_code})\ + + definition these :: "'a option set \ 'a set" where "these A = the ` {x \ A. x \ None}" @@ -210,15 +276,7 @@ subsubsection {* Code generator setup *} -definition is_none :: "'a option \ bool" where - [code_post]: "is_none x \ x = None" - -lemma is_none_code [code]: - shows "is_none None \ True" - and "is_none (Some x) \ False" - unfolding is_none_def by simp_all - -lemma [code_unfold]: +lemma equal_None_code_unfold [code_unfold]: "HOL.equal x None \ is_none x" "HOL.equal None = is_none" by (auto simp add: equal is_none_def)