diff -r 779af7c58743 -r 78b1d6c3ee9c src/HOL/MicroJava/J/JBasis.ML --- a/src/HOL/MicroJava/J/JBasis.ML Wed Dec 06 19:09:34 2000 +0100 +++ b/src/HOL/MicroJava/J/JBasis.ML Wed Dec 06 19:10:36 2000 +0100 @@ -4,92 +4,21 @@ Copyright 1999 TU Muenchen *) -val strip_tac1 = SELECT_GOAL (safe_tac (HOL_cs delrules [conjI, disjE, impCE])); - -Goalw [image_def] - "x \\ f``A ==> \\y. y \\ A \\ x = f y"; -by(Auto_tac); -qed "image_rev"; - -fun case_tac1 s i = EVERY [case_tac s i, rotate_tac ~1 i, rotate_tac ~1 (i+1)]; - -val select_split = prove_goalw Product_Type.thy [split_def] - "(SOME (x,y). P x y) = (SOME xy. P (fst xy) (snd xy))" (K [rtac refl 1]); - - -val split_beta = prove_goal Product_Type.thy "(\\(x,y). P x y) z = P (fst z) (snd z)" - (fn _ => [stac surjective_pairing 1, stac split 1, rtac refl 1]); -val split_beta2 = prove_goal Product_Type.thy "(\\(x,y). P x y) (w,z) = P w z" - (fn _ => [Auto_tac]); -val splitE2 = prove_goal Product_Type.thy "[|Q (split P z); !!x y. [|z = (x, y); Q (P x y)|] ==> R|] ==> R" (fn prems => [ - REPEAT (resolve_tac (prems@[surjective_pairing]) 1), - rtac (split_beta RS subst) 1, - rtac (hd prems) 1]); -val splitE2' = prove_goal Product_Type.thy "[|((\\(x,y). P x y) z) w; !!x y. [|z = (x, y); (P x y) w|] ==> R|] ==> R" (fn prems => [ - REPEAT (resolve_tac (prems@[surjective_pairing]) 1), - res_inst_tac [("P1","P")] (split_beta RS subst) 1, - rtac (hd prems) 1]); - - -fun pair_tac s = res_inst_tac [("p",s)] PairE THEN' hyp_subst_tac; - -val BallE = prove_goal thy "[|Ball A P; x \\ A ==> Q; P x ==> Q |] ==> Q" - (fn prems => [rtac ballE 1, resolve_tac prems 1, - eresolve_tac prems 1, eresolve_tac prems 1]); - -val set_cs2 = set_cs delrules [ballE] addSEs [BallE]; +(*###TO Product_Type*) +Goalw [split_def] "Eps (split P) = (SOME xy. P (fst xy) (snd xy))"; +by (rtac refl 1); +qed "select_split"; Addsimps [Let_def]; -(* To HOL.ML *) - -val ex1_some_eq_trivial = prove_goal HOL.thy "[| \\!x. P x; P y |] ==> Eps P = y" - (fn prems => [ - cut_facts_tac prems 1, - rtac some_equality 1, - atac 1, - etac ex1E 1, - etac all_dupE 1, - fast_tac HOL_cs 1]); - - -val ball_insert = prove_goalw equalities.thy [Ball_def] - "Ball (insert x A) P = (P x \\ Ball A P)" (fn _ => [ - fast_tac set_cs 1]); - -fun option_case_tac i = EVERY [ - etac option_caseE i, - rotate_tac ~2 (i+1), asm_full_simp_tac HOL_basic_ss (i+1), - rotate_tac ~2 i , asm_full_simp_tac HOL_basic_ss i]; - -val not_None_tac = EVERY' [dtac (not_None_eq RS iffD1),rotate_tac ~1,etac exE, - rotate_tac ~1,asm_full_simp_tac - (simpset() delsimps [split_paired_All,split_paired_Ex])]; - -Goal "{y. x = Some y} \\ {the x}"; -by Auto_tac; -qed "some_subset_the"; - -fun ex_ftac thm = EVERY' [forward_tac [thm], REPEAT o (etac exE), rotate_tac ~1, - asm_full_simp_tac (simpset() delsimps [split_paired_All,split_paired_Ex])]; - -val optionE = prove_goal thy - "[| opt = None ==> P; !!x. opt = Some x ==> P |] ==> P" - (fn prems => [ - case_tac "opt = None" 1, - eresolve_tac prems 1, - not_None_tac 1, - eresolve_tac prems 1]); - -fun option_case_tac2 s i = EVERY [ - case_tac s i, - rotate_tac ~1 (i+1), asm_full_simp_tac HOL_basic_ss (i+1), - rotate_tac ~1 i , asm_full_simp_tac HOL_basic_ss i]; - -val option_map_SomeD = prove_goalw thy [option_map_def] - "!!x. option_map f x = Some y ==> \\z. x = Some z \\ f z = y" (K [ - option_case_tac2 "x" 1, - Auto_tac]); +(* ### To HOL.ML *) +Goal "[| ?!x. P x; P y |] ==> Eps P = y"; +by (rtac some_equality 1); +by (atac 1); +by (etac ex1E 1); +by (etac all_dupE 1); +by (fast_tac HOL_cs 1); +qed "ex1_some_eq_trivial"; section "unique"; @@ -103,7 +32,7 @@ by (Simp_tac 1); qed "unique_Nil"; -Goalw [unique_def] "unique ((x,y)#l) = (unique l \\ (!y. (x,y) ~: set l))"; +Goalw [unique_def] "unique ((x,y)#l) = (unique l & (!y. (x,y) ~: set l))"; by (auto_tac (claset() addDs [fst_in_set_lemma],simpset())); qed "unique_Cons"; @@ -120,47 +49,14 @@ by (auto_tac (claset() addDs [fst_in_set_lemma],simpset()addsimps[inj_eq])); qed_spec_mp "unique_map_inj"; -Goal "!!l. unique l ==> unique (map (split (\\k. Pair (k, C))) l)"; -by(etac unique_map_inj 1); -by(rtac injI 1); -by Auto_tac; -qed "unique_map_Pair"; - -Goal "[|M = N; !!x. x\\N ==> f x = g x|] ==> f``M = g``N"; -by(rtac set_ext 1); -by(simp_tac (simpset() addsimps image_def::premises()) 1); -qed "image_cong"; - -val split_Pair_eq = prove_goal Product_Type.thy -"!!X. ((x, y), z) \\ split (\\x. Pair (x, Y)) `` A ==> y = Y" (K [ - etac imageE 1, - split_all_tac 1, - auto_tac(claset_of Product_Type.thy,simpset_of Product_Type.thy)]); - - (* More about Maps *) -val override_SomeD = prove_goalw thy [override_def] "(s ++ t) k = Some x ==> \ -\ t k = Some x | t k = None \\ s k = Some x" (fn prems => [ - cut_facts_tac prems 1, - case_tac "\\x. t k = Some x" 1, - etac exE 1, - rotate_tac ~1 1, - Asm_full_simp_tac 1, - asm_full_simp_tac (HOL_ss addsimps [not_None_eq RS sym]) 1, - rotate_tac ~1 1, - Asm_full_simp_tac 1]); +(*###Addsimps [fun_upd_same, fun_upd_other];*) -Addsimps [fun_upd_same, fun_upd_other]; - -Goal "unique xys --> (map_of xys x = Some y) = ((x,y) \\ set xys)"; -by(induct_tac "xys" 1); - by(Simp_tac 1); -by(force_tac (claset(), simpset() addsimps [unique_Cons]) 1); -qed_spec_mp "unique_map_of_Some_conv"; - -val in_set_get = unique_map_of_Some_conv RS iffD2; -val get_in_set = unique_map_of_Some_conv RS iffD1; +Goal "unique l --> (k, x) : set l --> map_of l k = Some x"; +by (induct_tac "l" 1); +by Auto_tac; +qed_spec_mp "map_of_SomeI"; Goal "(\\(x,y)\\set l. P x y) --> (\\x. \\y. map_of l x = Some y --> P x y)"; by(induct_tac "l" 1); @@ -169,50 +65,14 @@ by Auto_tac; bind_thm("Ball_set_table",result() RS mp); -val table_mono = prove_goal thy -"unique l' --> (\\xy. (xy)\\set l --> (xy)\\set l') -->\ -\ (\\k y. map_of l k = Some y --> map_of l' k = Some y)" (fn _ => [ - induct_tac "l" 1, - Auto_tac, - fast_tac (HOL_cs addSIs [in_set_get]) 1]) - RS mp RS mp RS spec RS spec RS mp; - -val table_map_Some' = prove_goal thy "map_of t k = Some (k', x) --> \ -\ map_of (map (\\u. ((fst u, fst (snd u)), snd (snd u))) t) (k, k') = Some x" (K [ - induct_tac "t" 1, - ALLGOALS Simp_tac, - case_tac1 "k = fst a" 1, - Auto_tac]) RS mp; -val table_map_Some = prove_goal thy -"map_of (map (\\((k,k'),x). (k,(k',x))) t) k = Some (k',x) --> \ -\map_of t (k, k') = Some x" (K [ - induct_tac "t" 1, - Auto_tac]) RS mp; - +Goal "map_of (map (\\((k,k'),x). (k,(k',x))) t) k = Some (k',x) --> \ +\map_of t (k, k') = Some x"; +by (induct_tac "t" 1); +by Auto_tac; +qed_spec_mp "table_of_remap_SomeD"; -val table_mapf_Some = prove_goal thy "!!f. \\x y. f x = f y --> x = y ==> \ -\ map_of (map (\\(k,x). (k,f x)) t) k = Some (f x) --> map_of t k = Some x" (K [ - induct_tac "t" 1, - Auto_tac]) RS mp; -val table_mapf_SomeD = prove_goal thy -"map_of (map (\\(k,x). (k, f x)) t) k = Some z --> (\\y. (k,y)\\set t \\ z = f y)"(K [ - induct_tac "t" 1, - Auto_tac]) RS mp; - -val table_mapf_Some2 = prove_goal thy -"!!k. map_of (map (\\(k,x). (k,C,x)) t) k = Some (D,x) ==> C = D \\ map_of t k = Some x" (K [ - forward_tac [table_mapf_SomeD] 1, - Auto_tac, - rtac table_mapf_Some 1, - atac 2, - Fast_tac 1]); - -val finite_map_of = rewrite_rule [dom_def] finite_dom_map_of; - -Goal -"map_of (map (\\(a,b). (a,f b)) xs) x = option_map f (map_of xs x)"; +(* ### To Map.ML *) +Goal "map_of (map (\\(a,b). (a,f b)) xs) x = option_map f (map_of xs x)"; by (induct_tac "xs" 1); -auto(); +by Auto_tac; qed "map_of_map"; - -