# HG changeset patch # User haftmann # Date 1346998818 -7200 # Node ID 3f85cd15a0ccc7c95f15c9b16dd54e1381f1649a # Parent 22f7e7b68f50b4e93888061f6cee31d44d51bb29 combinator Option.these diff -r 22f7e7b68f50 -r 3f85cd15a0cc NEWS --- a/NEWS Fri Sep 07 07:20:55 2012 +0200 +++ b/NEWS Fri Sep 07 08:20:18 2012 +0200 @@ -41,6 +41,8 @@ *** HOL *** +* New combinator "Option.these" with type "'a option set => 'a option". + * Renamed theory Library/List_Prefix to Library/Sublist. INCOMPATIBILITY. Related changes are: diff -r 22f7e7b68f50 -r 3f85cd15a0cc src/HOL/Option.thy --- a/src/HOL/Option.thy Fri Sep 07 07:20:55 2012 +0200 +++ b/src/HOL/Option.thy Fri Sep 07 08:20:18 2012 +0200 @@ -117,9 +117,57 @@ lemma bind_cong: "x = y \ (\a. y = Some a \ f a = g a) \ bind x f = bind y g" by (cases x) auto -hide_const (open) set map bind +definition these :: "'a option set \ 'a set" +where + "these A = the ` {x \ A. x \ None}" + +lemma these_empty [simp]: + "these {} = {}" + by (simp add: these_def) + +lemma these_insert_None [simp]: + "these (insert None A) = these A" + by (auto simp add: these_def) + +lemma these_insert_Some [simp]: + "these (insert (Some x) A) = insert x (these A)" +proof - + have "{y \ insert (Some x) A. y \ None} = insert (Some x) {y \ A. y \ None}" + by auto + then show ?thesis by (simp add: these_def) +qed + +lemma in_these_eq: + "x \ these A \ Some x \ A" +proof + assume "Some x \ A" + then obtain B where "A = insert (Some x) B" by auto + then show "x \ these A" by (auto simp add: these_def intro!: image_eqI) +next + assume "x \ these A" + then show "Some x \ A" by (auto simp add: these_def) +qed + +lemma these_image_Some_eq [simp]: + "these (Some ` A) = A" + by (auto simp add: these_def intro!: image_eqI) + +lemma Some_image_these_eq: + "Some ` these A = {x\A. x \ None}" + by (auto simp add: these_def image_image intro!: image_eqI) + +lemma these_empty_eq: + "these B = {} \ B = {} \ B = {None}" + by (auto simp add: these_def) + +lemma these_not_empty_eq: + "these B \ {} \ B \ {} \ B \ {None}" + by (auto simp add: these_empty_eq) + +hide_const (open) set map bind these hide_fact (open) map_cong bind_cong + subsubsection {* Code generator setup *} definition is_none :: "'a option \ bool" where @@ -164,3 +212,4 @@ Option None Some end +