# HG changeset patch # User paulson # Date 844677298 -7200 # Node ID ff04984186e997d10380f36d2b29d9f34f40bb86 # Parent 4d7a4b25a11f73ff0ec8c5fcf312c93290c37291 Tidied up some proofs diff -r 4d7a4b25a11f -r ff04984186e9 src/HOL/MiniML/I.ML --- a/src/HOL/MiniML/I.ML Mon Oct 07 10:31:50 1996 +0200 +++ b/src/HOL/MiniML/I.ML Mon Oct 07 10:34:58 1996 +0200 @@ -9,7 +9,6 @@ (* case Var n *) by (simp_tac (!simpset addsimps [app_subst_list] setloop (split_inside_tac [expand_if])) 1); - by (fast_tac (HOL_cs addss !simpset) 1); (* case Abs e *) by (asm_full_simp_tac (!simpset setloop (split_inside_tac [expand_bind])) 1); @@ -21,14 +20,14 @@ by (fast_tac (HOL_cs addss (!simpset addsimps [new_tv_subst])) 2); by (fast_tac (HOL_cs addIs [new_tv_Suc_list RS mp,new_tv_subst_le, less_imp_le,lessI]) 1); - +(** LEVEL 10 **) by (strip_tac 1); by (REPEAT (etac allE 1)); by (etac impE 1); by (fast_tac (HOL_cs addss (!simpset addsimps [new_tv_subst])) 2); by (fast_tac (HOL_cs addIs [new_tv_Suc_list RS mp,new_tv_subst_le, less_imp_le,lessI]) 1); - +(** LEVEL 15 **) (* case App e1 e2 *) by (simp_tac (!simpset setloop (split_inside_tac [expand_bind])) 1); by (strip_tac 1); @@ -49,6 +48,7 @@ by (eres_inst_tac [("x","s2'")] allE 1); by (eres_inst_tac [("x","t2")] allE 1); by (eres_inst_tac [("x","n2")] allE 1); +(** LEVEL 34 **) by (rtac conjI 1); by (strip_tac 1); by (mp_tac 1); @@ -61,6 +61,7 @@ by (fast_tac (HOL_cs addDs [sym RS W_var_geD,new_tv_subst_le,new_tv_list_le] addss !simpset) 1); by (fast_tac (HOL_cs addss (!simpset addsimps [subst_comp_tel])) 1); +(** LEVEL 45 **) by (strip_tac 1); by (rename_tac "s2 t2' n2'" 1); by (rtac conjI 1); @@ -77,6 +78,7 @@ by (fast_tac (HOL_cs addss (!simpset addsimps [subst_comp_tel,subst_comp_te])) 1); by (strip_tac 1); by (mp_tac 1); +(** LEVEL 60 **) by (mp_tac 1); by (etac exE 1); by (etac conjE 1); @@ -90,25 +92,26 @@ by (REPEAT(EVERY1 [asm_full_simp_tac (!simpset addsimps [subst_comp_tel,subst_comp_te]), REPEAT o etac conjE, hyp_subst_tac])); -by (rtac exI 1); +(** LEVEL 70 **) by (safe_tac HOL_cs); - by (fast_tac HOL_cs 1); by (simp_tac (!simpset addsimps [o_def,subst_comp_te]) 2); by (subgoal_tac "new_tv n2 s & new_tv n2 r & new_tv n2 ra" 1); by (asm_full_simp_tac (!simpset addsimps [new_tv_subst]) 1); by ((forward_tac [new_tv_subst_tel] 1) THEN (atac 1)); by ((dres_inst_tac [("a","$ s a")] new_tv_W 1) THEN (atac 1)); by (safe_tac HOL_cs); - by (fast_tac (HOL_cs addDs[sym RS W_var_geD,new_tv_subst_le,new_tv_list_le] + by (best_tac (HOL_cs addDs[sym RS W_var_geD,new_tv_subst_le,new_tv_list_le] addss !simpset) 1); by (fast_tac (HOL_cs addDs [sym RS W_var_geD,new_tv_subst_le,new_tv_list_le] addss !simpset) 1); +(** LEVEL 79 **) by (dres_inst_tac [("e","expr1")] (sym RS W_var_geD) 1); by ((dtac new_tv_subst_tel 1) THEN (atac 1)); by ((dres_inst_tac [("ts","$ s a")] new_tv_list_le 1) THEN (atac 1)); by ((dtac new_tv_subst_tel 1) THEN (atac 1)); -by (fast_tac (HOL_cs addDs [new_tv_W] addss - (!simpset addsimps [subst_comp_tel])) 1); +by (best_tac (HOL_cs addDs [new_tv_W] + addss (!simpset addsimps [subst_comp_tel])) 1); +(** LEVEL 84 **) qed_spec_mp "I_correct_wrt_W"; (*** @@ -133,22 +136,23 @@ goal I.thy "!a m s. \ \ new_tv m a & new_tv m s --> I e a m s = Fail --> W e ($s a) m = Fail"; by (expr.induct_tac "e" 1); - by(simp_tac (!simpset addsimps [app_subst_list] + by (simp_tac (!simpset addsimps [app_subst_list] setloop (split_tac [expand_if])) 1); - by(Simp_tac 1); - by(strip_tac 1); - br conjI 1; - by(strip_tac 1); - by(subgoal_tac "TVar m # $ s a = $s(TVar m # a)" 1); - by(asm_simp_tac (HOL_ss addsimps + by (Simp_tac 1); + by (strip_tac 1); + by (rtac conjI 1); + by (strip_tac 1); + by (subgoal_tac "TVar m # $ s a = $s(TVar m # a)" 1); + by (asm_simp_tac (HOL_ss addsimps [new_tv_Suc_list, lessI RS less_imp_le RS new_tv_subst_le]) 1); - be conjE 1; - bd (new_tv_not_free_tv RS not_free_impl_id) 1; - by(Asm_simp_tac 1); - by(strip_tac 1); - be exE 1; - by(split_all_tac 1); - by(Full_simp_tac 1); + by (etac conjE 1); + by (dtac (new_tv_not_free_tv RS not_free_impl_id) 1); + by (Asm_simp_tac 1); + by (strip_tac 1); + by (etac exE 1); + by (split_all_tac 1); + by (Full_simp_tac 1); +(** LEVEL 15 **) by (Asm_simp_tac 1); by (strip_tac 1); by (etac exE 1); @@ -156,43 +160,41 @@ by (split_all_tac 1); by (Full_simp_tac 1); by (dtac lemma 1); - by(fast_tac HOL_cs 1); + by (fast_tac HOL_cs 1); +(** LEVEL 23 **) by (etac exE 1); by (etac conjE 1); by (hyp_subst_tac 1); by (Asm_simp_tac 1); -by (rtac exI 1); -by (rtac conjI 1); - br refl 1; -by (Simp_tac 1); by (etac disjE 1); - br disjI1 1; - by(full_simp_tac (!simpset addsimps [o_def,subst_comp_tel]) 1); - by(EVERY[etac allE 1, etac allE 1, etac allE 1, + by (rtac disjI1 1); +(** LEVEL 29 **) + by (full_simp_tac (!simpset addsimps [o_def,subst_comp_tel]) 1); + by (EVERY[etac allE 1, etac allE 1, etac allE 1, etac impE 1, etac impE 2, atac 2, atac 2]); - br conjI 1; - by(fast_tac (HOL_cs addIs [W_var_ge RS new_tv_list_le]) 1); - br new_tv_subst_comp_2 1; - by(fast_tac (HOL_cs addIs [W_var_ge RS new_tv_subst_le]) 1); - by(fast_tac (HOL_cs addSIs [new_tv_subst_tel]addIs[new_tv_W RS conjunct1])1); + by (rtac conjI 1); + by (fast_tac (HOL_cs addIs [W_var_ge RS new_tv_list_le]) 1); + by (rtac new_tv_subst_comp_2 1); + by (fast_tac (HOL_cs addIs [W_var_ge RS new_tv_subst_le]) 1); + by (fast_tac (HOL_cs addSIs [new_tv_subst_tel]addIs[new_tv_W RS conjunct1])1); by (rtac disjI2 1); by (etac exE 1); by (split_all_tac 1); by (etac conjE 1); +(** LEVEL 40 **) by (Full_simp_tac 1); by (dtac lemma 1); - br conjI 1; - by(fast_tac (HOL_cs addIs [W_var_ge RS new_tv_list_le]) 1); - br new_tv_subst_comp_1 1; - by(fast_tac (HOL_cs addIs [W_var_ge RS new_tv_subst_le]) 1); - by(fast_tac (HOL_cs addSIs [new_tv_subst_tel]addIs[new_tv_W RS conjunct1])1); + by (rtac conjI 1); + by (fast_tac (HOL_cs addIs [W_var_ge RS new_tv_list_le]) 1); + by (rtac new_tv_subst_comp_1 1); + by (fast_tac (HOL_cs addIs [W_var_ge RS new_tv_subst_le]) 1); + by (fast_tac (HOL_cs addSIs [new_tv_subst_tel]addIs[new_tv_W RS conjunct1])1); by (etac exE 1); by (etac conjE 1); by (hyp_subst_tac 1); +(** LEVEL 50 **) by (asm_full_simp_tac (!simpset addsimps [o_def,subst_comp_te RS sym,subst_comp_tel RS sym]) 1); -by (asm_simp_tac (!simpset addcongs [conj_cong] addsimps - [eq_sym_conv,subst_comp_te RS sym,subst_comp_tel RS sym]) 1); qed_spec_mp "I_complete_wrt_W"; (*** diff -r 4d7a4b25a11f -r ff04984186e9 src/HOL/MiniML/Maybe.ML --- a/src/HOL/MiniML/Maybe.ML Mon Oct 07 10:31:50 1996 +0200 +++ b/src/HOL/MiniML/Maybe.ML Mon Oct 07 10:34:58 1996 +0200 @@ -22,7 +22,6 @@ goal Maybe.thy "((m bind f) = Fail) = ((m=Fail) | (? p. m = Ok p & f p = Fail))"; by (simp_tac (!simpset setloop (split_tac [expand_bind])) 1); -by (fast_tac HOL_cs 1); qed "bind_eq_Fail"; Addsimps [bind_eq_Fail]; diff -r 4d7a4b25a11f -r ff04984186e9 src/HOL/MiniML/W.ML --- a/src/HOL/MiniML/W.ML Mon Oct 07 10:31:50 1996 +0200 +++ b/src/HOL/MiniML/W.ML Mon Oct 07 10:34:58 1996 +0200 @@ -235,6 +235,7 @@ by (res_inst_tac [("x","s'")] exI 1); by (Asm_simp_tac 1); +(** LEVEL 10 **) (* case Abs e *) by (strip_tac 1); by (eresolve_tac has_type_casesE 1); @@ -245,6 +246,7 @@ by (fast_tac (HOL_cs addss (!simpset addcongs [conj_cong] setloop (split_tac [expand_bind]))) 1); +(** LEVEL 17 **) (* case App e1 e2 *) by (strip_tac 1); by (eresolve_tac has_type_casesE 1); @@ -266,6 +268,7 @@ by (fast_tac (HOL_cs addIs [sym RS W_var_geD,new_tv_W RS conjunct1,new_tv_list_le,new_tv_subst_tel]) 1); +(** LEVEL 35 **) by (subgoal_tac "$ (%x.if x=ma then t' else (if x:(free_tv t - free_tv sa) then r x \ \ else ra x)) ($ sa t) = \ @@ -280,6 +283,7 @@ by (subgoal_tac "na ~=ma" 2); by (fast_tac (HOL_cs addDs [new_tv_W,sym RS W_var_geD, new_tv_not_free_tv,new_tv_le]) 3); +(** LEVEL 42 **) by (case_tac "na:free_tv sa" 2); (* na ~: free_tv sa *) by (asm_simp_tac (!simpset addsimps [not_free_impl_id] @@ -293,6 +297,7 @@ (* na ~: dom sa *) by (asm_full_simp_tac (!simpset addsimps [dom_def] setloop (split_tac [expand_if])) 3); +(** LEVEL 50 **) (* na : dom sa *) by (rtac eq_free_eq_subst_te 2); by (strip_tac 2); @@ -307,6 +312,7 @@ setloop (split_tac [expand_if]))) 2); by (Simp_tac 2); +(** LEVEL 60 **) by (rtac eq_free_eq_subst_te 2); by (strip_tac 2 ); by (subgoal_tac "na ~= ma" 2); @@ -315,6 +321,7 @@ by (dtac (sym RS W_var_geD) 3); by (fast_tac (HOL_cs addDs [new_tv_list_le,new_tv_subst_tel,new_tv_W,new_tv_not_free_tv]) 3); by (case_tac "na: free_tv t - free_tv sa" 2); +(** LEVEL 68 **) (* case na ~: free_tv t - free_tv sa *) by ( asm_full_simp_tac (!simpset setloop (split_tac [expand_if])) 3); (* case na : free_tv t - free_tv sa *)