# HG changeset patch # User wenzelm # Date 1441042466 -7200 # Node ID 00a169fe5de4df65a7313c47752f79c17328e74e # Parent ca4ebc63d8aca4cb1d70b3495b8594c65894d663 misc tuning and modernization; diff -r ca4ebc63d8ac -r 00a169fe5de4 src/HOL/Option.thy --- a/src/HOL/Option.thy Mon Aug 31 19:04:24 2015 +0200 +++ b/src/HOL/Option.thy Mon Aug 31 19:34:26 2015 +0200 @@ -17,89 +17,86 @@ lemma [case_names None Some, cases type: option]: -- \for backward compatibility -- names of variables differ\ "(y = None \ P) \ (\a. y = Some a \ P) \ P" -by (rule option.exhaust) + by (rule option.exhaust) lemma [case_names None Some, induct type: option]: -- \for backward compatibility -- names of variables differ\ "P None \ (\option. P (Some option)) \ P option" -by (rule option.induct) + by (rule option.induct) text \Compatibility:\ - setup \Sign.mandatory_path "option"\ - lemmas inducts = option.induct lemmas cases = option.case - setup \Sign.parent_path\ -lemma not_None_eq [iff]: "(x ~= None) = (EX y. x = Some y)" +lemma not_None_eq [iff]: "x \ None \ (\y. x = Some y)" by (induct x) auto -lemma not_Some_eq [iff]: "(ALL y. x ~= Some y) = (x = None)" +lemma not_Some_eq [iff]: "(\y. x \ Some y) \ x = None" by (induct x) auto -text\Although it may appear that both of these equalities are helpful +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 inj_Some [simp]: "inj_on Some A" -by (rule inj_onI) simp + by (rule inj_onI) simp lemma case_optionE: - assumes c: "(case x of None => P | Some y => Q y)" + 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 split_option_all: "(\x. P x) \ P None \ (\x. P (Some x))" -by (auto intro: option.induct) + by (auto intro: option.induct) lemma split_option_ex: "(\x. P x) \ P None \ (\x. P (Some x))" -using split_option_all[of "\x. \P x"] by blast + using split_option_all[of "\x. \ P x"] by blast lemma UNIV_option_conv: "UNIV = insert None (range Some)" -by(auto intro: classical) + by (auto intro: classical) lemma rel_option_None1 [simp]: "rel_option P None x \ x = None" -by(cases x) simp_all + by (cases x) simp_all lemma rel_option_None2 [simp]: "rel_option P x None \ x = None" -by(cases x) simp_all + 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_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) + show "?rhs \ ?lhs" by (auto elim: option.rel_mono_strong) +qed lemma rel_option_reflI: "(\x. x \ set_option y \ P x x) \ rel_option P y y" -by(cases y) auto + by (cases y) auto subsubsection \Operations\ -lemma ospec [dest]: "(ALL x:set_option A. P x) ==> A = Some x ==> P x" +lemma ospec [dest]: "(\x\set_option A. P x) \ A = Some x \ P x" by simp setup \map_theory_claset (fn ctxt => ctxt addSD2 ("ospec", @{thm ospec}))\ -lemma elem_set [iff]: "(x : set_option xo) = (xo = Some x)" +lemma elem_set [iff]: "(x \ set_option xo) = (xo = Some x)" by (cases xo) auto lemma set_empty_eq [simp]: "(set_option xo = {}) = (xo = None)" by (cases xo) auto -lemma map_option_case: "map_option f y = (case y of None => None | Some x => Some (f x))" +lemma map_option_case: "map_option f y = (case y of None \ None | Some x \ Some (f x))" by (auto split: option.split) -lemma map_option_is_None [iff]: - "(map_option f opt = None) = (opt = None)" +lemma map_option_is_None [iff]: "(map_option f opt = None) = (opt = None)" by (simp add: map_option_case split add: option.split) -lemma map_option_eq_Some [iff]: - "(map_option f xo = Some y) = (EX z. xo = Some z & f z = y)" +lemma map_option_eq_Some [iff]: "(map_option f xo = Some y) = (\z. xo = Some z \ f z = y)" by (simp add: map_option_case split add: option.split) lemma map_option_o_case_sum [simp]: @@ -107,121 +104,114 @@ by (rule o_case_sum) lemma map_option_cong: "x = y \ (\a. y = Some a \ f a = g a) \ map_option f x = map_option g y" -by (cases x) auto + by (cases x) auto functor map_option: map_option -by(simp_all add: option.map_comp fun_eq_iff option.map_id) + by (simp_all add: option.map_comp fun_eq_iff option.map_id) -lemma case_map_option [simp]: - "case_option g h (map_option f x) = case_option g (h \ f) x" +lemma case_map_option [simp]: "case_option g h (map_option f x) = case_option g (h \ f) x" by (cases x) simp_all lemma rel_option_iff: "rel_option R x y = (case (x, y) of (None, None) \ True | (Some x, Some y) \ R x y | _ \ False)" -by (auto split: prod.split option.split) + by (auto split: prod.split option.split) definition is_none :: "'a option \ bool" -where [code_post]: "is_none x \ x = None" + 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) + by (simp_all add: is_none_def) lemma is_none_code [code]: "is_none None = True" "is_none (Some x) = False" -by simp_all + 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) + 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) + 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) + 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) + by (auto 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" +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 + by (cases f) simp_all lemma bind_runit[simp]: "bind x Some = x" -by (cases x) auto + by (cases x) auto lemma bind_assoc[simp]: "bind (bind x f) g = bind x (\y. bind (f y) g)" -by (cases x) auto + by (cases x) auto lemma bind_rzero[simp]: "bind x (\x. None) = None" -by (cases x) auto + by (cases x) auto lemma bind_cong: "x = y \ (\a. y = Some a \ f a = g a) \ bind x f = bind y g" -by (cases x) auto + by (cases x) auto -lemma bind_split: "P (bind m f) - \ (m = None \ P None) \ (\v. m=Some v \ P (f v))" - by (cases m) auto +lemma bind_split: "P (bind m f) \ (m = None \ P None) \ (\v. m = Some v \ P (f v))" + by (cases m) auto -lemma bind_split_asm: "P (bind m f) = (\( - m=None \ \P None - \ (\x. m=Some x \ \P (f x))))" +lemma bind_split_asm: "P (bind m f) \ \ (m = None \ \ P None \ (\x. m = Some x \ \ P (f x)))" by (cases m) auto 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 + 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 + 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 + 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) + unfolding simp_implies_def by (rule bind_option_cong) -lemma bind_option_cong_code: "x = y \ bind x f = bind y f" by simp +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}" + where "these A = the ` {x \ A. x \ None}" -lemma these_empty [simp]: - "these {} = {}" +lemma these_empty [simp]: "these {} = {}" by (simp add: these_def) -lemma these_insert_None [simp]: - "these (insert None A) = these A" +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)" +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" +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 @@ -231,20 +221,16 @@ then show "Some x \ A" by (auto simp add: these_def) qed -lemma these_image_Some_eq [simp]: - "these (Some ` A) = A" +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}" +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}" +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}" +lemma these_not_empty_eq: "these B \ {} \ B \ {} \ B \ {None}" by (auto simp add: these_empty_eq) hide_const (open) bind these @@ -255,6 +241,7 @@ context begin + interpretation lifting_syntax . lemma option_bind_transfer [transfer_rule]: @@ -264,7 +251,7 @@ lemma pred_option_parametric [transfer_rule]: "((A ===> op =) ===> rel_option A ===> op =) pred_option pred_option" -by(rule rel_funI)+(auto simp add: rel_option_unfold is_none_def dest: rel_funD) + by (rule rel_funI)+ (auto simp add: rel_option_unfold is_none_def dest: rel_funD) end @@ -276,7 +263,7 @@ by (auto simp add: UNIV_option_conv elim: finite_imageD intro: inj_Some) instance option :: (finite) finite - by default (simp add: UNIV_option_conv) + by standard (simp add: UNIV_option_conv) subsubsection \Code generator setup\