(* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) theory LemmaBucket imports HaskellLemmaBucket SpecValid_R SubMonadLib begin lemma corres_underlying_trivial: "\ nf' \ no_fail P' f \ \ corres_underlying Id nf nf' (=) \ P' f f" by (auto simp add: corres_underlying_def Id_def no_fail_def) lemma hoare_spec_gen_asm: "\ F \ s \ \P\ f \Q\ \ \ s \ \P and K F\ f \Q\" "\ F \ s \ \P\ f' \Q\,\E\ \ \ s \ \P and K F\ f' \Q\,\E\" unfolding spec_valid_def spec_validE_def validE_def apply (clarsimp simp only: pred_conj_def conj_assoc[symmetric] intro!: hoare_gen_asm[unfolded pred_conj_def])+ done lemma spec_validE_fail: "s \ \P\ fail \Q\,\E\" by wp+ lemma mresults_fail: "mresults fail = {}" by (simp add: mresults_def fail_def) lemma gets_symb_exec_l: "corres_underlying sr nf nf' dc P P' (gets f) (return x)" by (simp add: corres_underlying_def return_def simpler_gets_def split_def) lemmas mapM_x_wp_inv = mapM_x_wp[where S=UNIV, simplified] lemma mapM_wp_inv: "(\x. \P\ f x \\rv. P\) \ \P\ mapM f xs \\rv. P\" apply (rule valid_return_unit) apply (fold mapM_x_mapM) apply (erule mapM_x_wp_inv) done lemmas mapM_x_wp' = mapM_x_wp [OF _ subset_refl] lemma corres_underlying_similar: "\ a = a'; b = b'; nf' \ no_fail \ (f a b) \ \ corres_underlying Id nf nf' dc \ \ (f a b) (f a' b')" by (simp add: corres_underlying_def no_fail_def, blast) lemma corres_underlying_gets_pre_lhs: "(\x. corres_underlying S nf nf' r (P x) P' (g x) g') \ corres_underlying S nf nf' r (\s. P (f s) s) P' (gets f >>= (\x. g x)) g'" apply (simp add: simpler_gets_def bind_def split_def corres_underlying_def) apply force done lemma mapM_x_inv_wp: assumes x: "\s. I s \ Q s" assumes y: "\x. x \ set xs \ \P\ m x \\rv. I\" assumes z: "\s. I s \ P s" shows "\I\ mapM_x m xs \\rv. Q\" apply (rule hoare_post_imp) apply (erule x) apply (rule mapM_x_wp) apply (rule hoare_pre_imp [OF _ y]) apply (erule z) apply assumption apply simp done lemma mapM_x_accumulate_checks': assumes P: "\x. x \ set xs' \ \\\ f x \\rv. P x\" assumes P': "\x y. \ x \ set xs'; y \ set xs' \ \ \P y\ f x \\rv. P y\" shows "set xs \ set xs' \ \\\ mapM_x f xs \\rv s. \x \ set xs. P x s\" apply (induct xs) apply (simp add: mapM_x_Nil) apply (simp add: mapM_x_Cons) apply (rule hoare_pre) apply (wp mapM_x_wp'[OF P']) apply blast apply simp apply assumption apply simp apply (rule P) apply simp apply simp done lemmas mapM_x_accumulate_checks = mapM_x_accumulate_checks'[OF _ _ subset_refl] (* Other *) lemma isRight_rel_sum_comb2: "\ (f \ r) v v'; isRight v' \ \ isRight v \ r (theRight v) (theRight v')" by (cases v; clarsimp) lemma isRight_case_sum: "isRight x \ case_sum f g x = g (theRight x)" by (cases x; clarsimp) lemma enumerate_append:"enumerate i (xs @ ys) = enumerate i xs @ enumerate (i + length xs) ys" apply (induct xs arbitrary:ys i) apply clarsimp+ done lemma enumerate_bound:"(a, b) \ set (enumerate n xs) \ a < n + length xs" by (metis add.commute in_set_enumerate_eq prod.sel(1)) lemma enumerate_exceed:"(n + length xs, b) \ set (enumerate n xs)" by (metis enumerate_bound less_not_refl) lemma all_pair_unwrap:"(\a. P (fst a) (snd a)) = (\a b. P a b)" by force lemma if_fold[simp]:"(if P then Q else if P then R else S) = (if P then Q else S)" by presburger lemma disjoint_subset_both:"\A' \ A; B' \ B; A \ B = {}\ \ A' \ B' = {}" by blast lemma union_split: "\A \ C = {}; B \ C = {}\ \ (A \ B) \ C = {}" by (simp add: inf_sup_distrib2) lemma dom_expand: "dom (\x. if P x then Some y else None) = {x. P x}" using if_option_Some by fastforce lemma range_translate: "(range f = range g) = ((\x. \y. f x = g y) \ (\x. \y. f y = g x))" by (rule iffI, rule conjI, clarsimp, blast, clarsimp, metis f_inv_into_f range_eqI, clarsimp, subst set_eq_subset, rule conjI, clarsimp, rename_tac arg, erule_tac x=arg and P="\x. (\y. f x = g y)" in allE, clarsimp, clarsimp, rename_tac arg, erule_tac x=arg and P="\x. (\y. f y = g x)" in allE, clarsimp, metis range_eqI) lemma ran_expand: "\x. P x \ ran (\x. if P x then Some y else None) = {y}" by (rule subset_antisym, (clarsimp simp:ran_def)+) lemma map_upd_expand: "f(x \ y) = f ++ (\z. if z = x then Some y else None)" by (rule ext, rename_tac w, case_tac "w = x", simp, simp add:map_add_def) lemma map_upd_subI: "\f \\<^sub>m g; f x = None\ \ f \\<^sub>m g(x \ y)" by (rule_tac f="\i. if i = x then Some y else None" in map_add_le_mapE, simp add:map_le_def, rule ballI, rename_tac a, rule conjI, erule_tac x=x in ballE, clarsimp, erule disjE, clarsimp, clarsimp simp:map_add_def, clarsimp, erule disjE, clarsimp, clarsimp simp:map_add_def, clarsimp simp:map_add_def, erule_tac x=a in ballE, erule disjE, (case_tac "g a"; simp_all), clarsimp+) lemma all_ext: "\x. f x = g x \ f = g" by presburger lemma conjI2: "\B; B \ A\ \ A \ B" by auto (* Trivial lemmas for dealing with messy CNode obligations. *) lemma Least2: "\\P 0; \P 1; P (2::nat)\ \ Least P = 2" by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3)) lemma Least3: "\\P 0; \P 1; \P 2; P (3::nat)\ \ Least P = 3" by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3)) lemma Least4: "\\P 0; \P 1; \P 2; \P 3; P (4::nat)\ \ Least P = 4" by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3)) lemma Least5: "\\P 0; \P 1; \P 2; \P 3; \P 4; P (5::nat)\ \ Least P = 5" by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3)) lemma Least6: "\\P 0; \P 1; \P 2; \P 3; \P 4; \P 5; P (6::nat)\ \ Least P = 6" by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3)) lemma Least7: "\\P 0; \P 1; \P 2; \P 3; \P 4; \P 5; \P 6; P (7::nat)\ \ Least P = 7" by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3)) lemma Least8: "\\P 0; \P 1; \P 2; \P 3; \P 4; \P 5; \P 6; \P 7; P (8::nat)\ \ Least P = 8" by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3)) lemma Least9: "\\P 0; \P 1; \P 2; \P 3; \P 4; \P 5; \P 6; \P 7; \P 8; P (9::nat)\ \ Least P = 9" by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3)) lemma Least10: "\\P 0; \P 1; \P 2; \P 3; \P 4; \P 5; \P 6; \P 7; \P 8; \P 9; P (10::nat)\ \ Least P = 10" by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3)) lemma Least11: "\\P 0; \P 1; \P 2; \P 3; \P 4; \P 5; \P 6; \P 7; \P 8; \P 9; \P 10; P (11::nat)\ \ Least P = 11" by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3)) lemma Least12: "\\P 0; \P 1; \P 2; \P 3; \P 4; \P 5; \P 6; \P 7; \P 8; \P 9; \P 10; \P 11; P (12::nat)\ \ Least P = 12" by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3)) lemma Least13: "\\P 0; \P 1; \P 2; \P 3; \P 4; \P 5; \P 6; \P 7; \P 8; \P 9; \P 10; \P 11; \P 12; P (13::nat)\ \ Least P = 13" by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3)) lemma Least14: "\\P 0; \P 1; \P 2; \P 3; \P 4; \P 5; \P 6; \P 7; \P 8; \P 9; \P 10; \P 11; \P 12; \P 13; P (14::nat)\ \ Least P = 14" by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3)) lemma Least15: "\\P 0; \P 1; \P 2; \P 3; \P 4; \P 5; \P 6; \P 7; \P 8; \P 9; \P 10; \P 11; \P 12; \P 13; \P 14; P (15::nat)\ \ Least P = 15" by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3)) lemma Least16: "\\P 0; \P 1; \P 2; \P 3; \P 4; \P 5; \P 6; \P 7; \P 8; \P 9; \P 10; \P 11; \P 12; \P 13; \P 14; \P 15; P (16::nat)\ \ Least P = 16" by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3)) lemma Least17: "\\P 0; \P 1; \P 2; \P 3; \P 4; \P 5; \P 6; \P 7; \P 8; \P 9; \P 10; \P 11; \P 12; \P 13; \P 14; \P 15; \P 16; P (17::nat)\ \ Least P = 17" by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3)) lemma Least18: "\\P 0; \P 1; \P 2; \P 3; \P 4; \P 5; \P 6; \P 7; \P 8; \P 9; \P 10; \P 11; \P 12; \P 13; \P 14; \P 15; \P 16; \P 17; P (18::nat)\ \ Least P = 18" by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3)) lemma Least19: "\\P 0; \P 1; \P 2; \P 3; \P 4; \P 5; \P 6; \P 7; \P 8; \P 9; \P 10; \P 11; \P 12; \P 13; \P 14; \P 15; \P 16; \P 17; \P 18; P (19::nat)\ \ Least P = 19" by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3)) lemma Least20: "\\P 0; \P 1; \P 2; \P 3; \P 4; \P 5; \P 6; \P 7; \P 8; \P 9; \P 10; \P 11; \P 12; \P 13; \P 14; \P 15; \P 16; \P 17; \P 18; \P 19; P (20::nat)\ \ Least P = 20" by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3)) lemma Least21: "\\P 0; \P 1; \P 2; \P 3; \P 4; \P 5; \P 6; \P 7; \P 8; \P 9; \P 10; \P 11; \P 12; \P 13; \P 14; \P 15; \P 16; \P 17; \P 18; \P 19; \P 20; P (21::nat)\ \ Least P = 21" by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3)) lemma Least22: "\\P 0; \P 1; \P 2; \P 3; \P 4; \P 5; \P 6; \P 7; \P 8; \P 9; \P 10; \P 11; \P 12; \P 13; \P 14; \P 15; \P 16; \P 17; \P 18; \P 19; \P 20; \P 21; P (22::nat)\ \ Least P = 22" by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3)) lemma Least23: "\\P 0; \P 1; \P 2; \P 3; \P 4; \P 5; \P 6; \P 7; \P 8; \P 9; \P 10; \P 11; \P 12; \P 13; \P 14; \P 15; \P 16; \P 17; \P 18; \P 19; \P 20; \P 21; \P 22; P (23::nat)\ \ Least P = 23" by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3)) lemma Least24: "\\P 0; \P 1; \P 2; \P 3; \P 4; \P 5; \P 6; \P 7; \P 8; \P 9; \P 10; \P 11; \P 12; \P 13; \P 14; \P 15; \P 16; \P 17; \P 18; \P 19; \P 20; \P 21; \P 22; \P 23; P (24::nat)\ \ Least P = 24" by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3)) lemma Least25: "\\P 0; \P 1; \P 2; \P 3; \P 4; \P 5; \P 6; \P 7; \P 8; \P 9; \P 10; \P 11; \P 12; \P 13; \P 14; \P 15; \P 16; \P 17; \P 18; \P 19; \P 20; \P 21; \P 22; \P 23; \P 24; P (25::nat)\ \ Least P = 25" by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3)) lemma Least26: "\\P 0; \P 1; \P 2; \P 3; \P 4; \P 5; \P 6; \P 7; \P 8; \P 9; \P 10; \P 11; \P 12; \P 13; \P 14; \P 15; \P 16; \P 17; \P 18; \P 19; \P 20; \P 21; \P 22; \P 23; \P 24; \P 25; P (26::nat)\ \ Least P = 26" by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3)) lemma Least27: "\\P 0; \P 1; \P 2; \P 3; \P 4; \P 5; \P 6; \P 7; \P 8; \P 9; \P 10; \P 11; \P 12; \P 13; \P 14; \P 15; \P 16; \P 17; \P 18; \P 19; \P 20; \P 21; \P 22; \P 23; \P 24; \P 25; \P 26; P (27::nat)\ \ Least P = 27" by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3)) lemma Least28: "\\P 0; \P 1; \P 2; \P 3; \P 4; \P 5; \P 6; \P 7; \P 8; \P 9; \P 10; \P 11; \P 12; \P 13; \P 14; \P 15; \P 16; \P 17; \P 18; \P 19; \P 20; \P 21; \P 22; \P 23; \P 24; \P 25; \P 26; \P 27; P (28::nat)\ \ Least P = 28" by (simp add: Least_Suc eval_nat_numeral(2) eval_nat_numeral(3)) lemma map_add_discard: "\ cond x \ (f ++ (\x. if cond x then (g x) else None)) x = f x" by (simp add: map_add_def) lemma dom_split:"\\x \ S. \y. f x = Some y; \x. x \ S \ f x = None\ \ dom f = S" by (auto simp:dom_def) lemma map_set_in: "x \ f ` S = (\y\S. f y = x)" by blast lemma map_length_split: "map (length \ (\(a, b). P a b # map (f a b) (Q a b))) xs = map (\(a, b). 1 + length (Q a b)) xs" by clarsimp lemma sum_suc: "(\x \ xs. Suc (f x)) = length xs + (\x \ xs. f x)" apply (induct xs) by clarsimp+ lemma sum_suc_pair: "(\(a, b) \ xs. Suc (f a b)) = length xs + (\(a, b) \ xs. f a b)" apply (induct xs) by clarsimp+ lemma fold_add_sum: "fold (+) ((map (\(a, b). f a b) xs)::nat list) 0 = (\(a, b) \ xs. f a b)" apply (subst fold_plus_sum_list_rev) apply (subst sum_list_rev) by clarsimp lemma set_of_enumerate:"card (set (enumerate n xs)) = length xs" by (metis distinct_card distinct_enumerate length_enumerate) lemma collapse_fst: "fst ` (\x. (f x, g x)) ` s = f ` s" by force lemma collapse_fst2: "fst ` (\(x, y). (f x, g y)) ` s = (\x. f (fst x)) ` s" by force lemma collapse_fst3: "(\x. f (fst x)) ` set (enumerate n xs) = f ` set [n.. 'b option" assumes "finite (UNIV::'a set)" shows "card (dom f) \ CARD('a)" by (simp add: assms card_mono) lemma third_in: "(a, b, c) \ S \ c \ (snd \ snd) ` S" by (metis (erased, opaque_lifting) map_set_in image_comp snd_conv) lemma third_in2: "(a \ (snd \ snd) ` (set (enumerate i xs))) = (a \ snd ` (set xs))" by (metis map_map map_snd_enumerate set_map) lemma map_of_enum: "map_of (enumerate n xs) x = Some y \ y \ set xs" apply (clarsimp) by (metis enumerate_eq_zip in_set_zipE) lemma map_of_append: "(map_of xs ++ map_of ys) x = (case map_of ys x of None \ map_of xs x | Some x' \ Some x')" by (simp add: map_add_def) lemma map_of_append2: "(map_of xs ++ map_of ys ++ map_of zs) x = (case map_of zs x of None \ (case map_of ys x of None \ map_of xs x | Some x' \ Some x') | Some x' \ Some x')" by (simp add: map_add_def) lemma map_of_in_set_map: "map_of (map (\(n, y). (f n, y)) xs) x = Some z \ z \ snd ` set xs" proof - assume "map_of (map (\(n, y). (f n, y)) xs) x = Some z" hence "(x, z) \ (\(uu, y). (f uu, y)) ` set xs" using map_of_SomeD by fastforce thus "z \ snd ` set xs" using map_set_in by fastforce qed lemma pair_in_enum: "(a, b) \ set (enumerate x ys) \ b \ set ys" by (metis enumerate_eq_zip in_set_zip2) lemma distinct_inj: "inj f \ distinct xs = distinct (map f xs)" apply (induct xs) apply simp apply (simp add: inj_image_mem_iff) done lemma distinct_map_via_ran: "distinct (map fst xs) \ ran (map_of xs) = set (map snd xs)" apply (cut_tac xs="map fst xs" and ys="map snd xs" in ran_map_of_zip[symmetric]) apply clarsimp+ by (simp add: ran_distinct) lemma in_ran_in_set: "x \ ran (map_of xs) \ x \ set (map snd xs)" by (metis (mono_tags) map_set_in map_of_SomeD ranE set_map snd_conv) lemma in_ran_map_app: "x \ ran (xs ++ ys ++ zs) \ x \ ran xs \ x \ ran ys \ x \ ran zs" proof - assume a1: "x \ ran (xs ++ ys ++ zs)" obtain bb :: "'a \ ('b \ 'a option) \ 'b" where "\x0 x1. (\v2. x1 v2 = Some x0) = (x1 (bb x0 x1) = Some x0)" by moura hence f2: "\f a. (\ (\b. f b = Some a) \ f (bb a f) = Some a) \ ((\b. f b = Some a) \ (\b. f b \ Some a))" by blast have "\b. (xs ++ ys ++ zs) b = Some x" using a1 by (simp add: ran_def) hence f3: "(xs ++ ys ++ zs) (bb x (xs ++ ys ++ zs)) = Some x" using f2 by meson { assume "ys (bb x (xs ++ ys ++ zs)) \ None \ xs (bb x (xs ++ ys ++ zs)) \ Some x" { assume "ys (bb x (xs ++ ys ++ zs)) \ Some x \ (ys (bb x (xs ++ ys ++ zs)) \ None \ xs (bb x (xs ++ ys ++ zs)) \ Some x)" hence "\b. zs b = Some x" using f3 by auto hence ?thesis by (simp add: ran_def) } hence ?thesis using ran_def by fastforce } thus ?thesis using ran_def by fastforce qed lemma none_some_map: "None \ S \ Some x \ S = (x \ the ` S)" apply (rule iffI) apply force apply (subst in_these_eq[symmetric]) apply (clarsimp simp:Option.these_def) apply (case_tac "\y. xa = Some y") by clarsimp+ lemma none_some_map2: "the ` Set.filter (\s. \ Option.is_none s) (range f) = ran f" apply (rule subset_antisym) apply clarsimp apply (case_tac "f x", simp_all) apply (simp add: ranI) apply clarsimp apply (subst none_some_map[symmetric]) apply clarsimp+ apply (erule ranE) by (metis range_eqI) lemma prop_map_of_prop:"\\z \ set xs. P (g z); map_of (map (\x. (f x, g x)) xs) y = Some a\ \ P a" using map_of_SomeD by fastforce lemma range_subsetI2: "\y\A. \x. f x = y \ A \ range f" by fast lemma insert_strip: "x \ y \ (x \ insert y S) = (x \ S)" by simp lemma dom_map_add: "dom ys = A \ dom (xs ++ ys) = A \ dom xs" by simp lemma set_compre_unwrap: "({x. P x} \ S) = (\x. P x \ x \ S)" by blast lemma map_add_same: "\xs = ys; zs = ws\ \ xs ++ zs = ys ++ ws" by simp lemma map_add_find_left: "n k = None \ (m ++ n) k = m k" by (simp add:map_add_def) lemma map_length_split_triple: "map (length \ (\(a, b, c). P a b c # map (f a b c) (Q a b c))) xs = map (\(a, b, c). 1 + length (Q a b c)) xs" by fastforce lemma sum_suc_triple: "(\(a, b, c)\xs. Suc (f a b c)) = length xs + (\(a, b, c)\xs. f a b c)" by (induct xs; clarsimp) lemma sum_enumerate: "(\(a, b)\enumerate n xs. P b) = (\b\xs. P b)" by (induct xs arbitrary:n; clarsimp) lemma dom_map_fold:"dom (fold (++) (map (\x. [f x \ g x]) xs) ms) = dom ms \ set (map f xs)" by (induct xs arbitrary:f g ms; clarsimp) lemma list_ran_prop:"map_of (map (\x. (f x, g x)) xs) i = Some t \ \x \ set xs. g x = t" by (induct xs arbitrary:f g t i; clarsimp split:if_split_asm) lemma in_set_enumerate_eq2:"(a, b) \ set (enumerate n xs) \ (b = xs ! (a - n))" by (simp add: in_set_enumerate_eq) lemma subset_eq_notI: "\a\ B;a\ C\ \ \ B \ C" by auto lemma nat_divide_less_eq: fixes b :: nat shows "0 < c \ (b div c < a) = (b < a * c)" using td_gal_lt by blast lemma strengthen_imp_same_first_conj: "(b \ (a \ c) \ (a' \ c')) \ ((a \ b \ c) \ (a' \ b \ c'))" by blast lemma conj_impD: "a \ b \ a \ b" by blast lemma set_list_mem_nonempty: "x \ set xs \ xs \ []" by auto lemma strenghten_False_imp: "\P \ P \ Q" by blast lemma None_in_range_Some [simp]: "(None \ range Some) = False" by auto lemma foldl_fun_or_alt: "foldl (\x y. x \ f y) b ls = foldl (\) b (map f ls)" apply (induct ls) apply clarsimp apply clarsimp by (simp add: foldl_map) lemma sorted_imp_sorted_filter: "sorted xs \ sorted (filter P xs)" by (metis filter_sort sorted_sort sorted_sort_id) lemma sorted_list_of_set_already_sorted: "\ distinct xs; sorted xs \ \ sorted_list_of_set (set xs) = xs" by (simp add: sorted_list_of_set_sort_remdups distinct_remdups_id sorted_sort_id) lemma inj_on_domD: "\inj_on f (dom f); f x = Some z; f y = Some z\ \ x = y" by (erule inj_onD) clarsimp+ lemma not_emptyI: "\x A B. \x\A; x\B\ \ A \ B \ {}" by auto lemma add_mask_lower_bits2: "\is_aligned (x :: 'a :: len word) n; p && ~~ mask n = 0\ \ x + p && ~~ mask n = x" apply (subst word_plus_and_or_coroll) apply (simp add: aligned_mask_disjoint and_mask_0_iff_le_mask) apply (clarsimp simp: bit.conj_disj_distrib2) done (* FIXME: move to GenericLib *) lemma if3_fold2: "(if P then x else if Q then x else y) = (if P \ Q then x else y)" by (rule z3_rule) lemma inter_UNIV_minus[simp]: "x \ (UNIV - y) = x-y" by blast lemma imp_and_strg: "Q \ C \ (A \ Q \ C) \ C" by blast lemma cases_conj_strg: "A \ B \ (P \ A) \ (\ P \ B)" by simp lemma and_not_not_or_imp: "(~ A & ~ B | C) = ((A | B) \ C)" by blast end