diff -r 5c613b64bf0a -r ecf182ccc3ca src/HOLCF/FOCUS/Buffer.ML --- a/src/HOLCF/FOCUS/Buffer.ML Tue Sep 06 20:53:27 2005 +0200 +++ b/src/HOLCF/FOCUS/Buffer.ML Tue Sep 06 21:51:17 2005 +0200 @@ -1,13 +1,13 @@ -(* Title: HOLCF/FOCUS/Buffer.ML +(* Title: HOLCF/FOCUS/Buffer.ML ID: $Id$ - Author: David von Oheimb, TU Muenchen + Author: David von Oheimb, TU Muenchen *) -val set_cong = prove_goal Set.thy "!!X. A = B ==> (x:A) = (x:B)" (K [ - etac subst 1, rtac refl 1]); +val set_cong = prove_goal (theory "Set") "!!X. A = B ==> (x:A) = (x:B)" (K [ + etac subst 1, rtac refl 1]); -fun B_prover s tac eqs = prove_goal thy s (fn prems => [cut_facts_tac prems 1, - tac 1, auto_tac (claset(), simpset() addsimps eqs)]); +fun B_prover s tac eqs = prove_goal (the_context ()) s (fn prems => [cut_facts_tac prems 1, + tac 1, auto_tac (claset(), simpset() addsimps eqs)]); fun prove_forw s thm = B_prover s (dtac (thm RS iffD1)) []; fun prove_backw s thm eqs = B_prover s (rtac (thm RS iffD2)) eqs; @@ -15,74 +15,74 @@ (**** BufEq *******************************************************************) -val mono_BufEq_F = prove_goalw thy [mono_def, BufEq_F_def] - "mono BufEq_F" (K [Fast_tac 1]); +val mono_BufEq_F = prove_goalw (the_context ()) [mono_def, BufEq_F_def] + "mono BufEq_F" (K [Fast_tac 1]); val BufEq_fix = mono_BufEq_F RS (BufEq_def RS def_gfp_unfold); -val BufEq_unfold = prove_goal thy "(f:BufEq) = (!d. f\\(Md d\\<>) = <> & \ - \(!x. ? ff:BufEq. f\\(Md d\\\\\\x) = d\\(ff\\x)))" (K [ - stac (BufEq_fix RS set_cong) 1, - rewtac BufEq_F_def, - Asm_full_simp_tac 1]); +val BufEq_unfold = prove_goal (the_context ()) "(f:BufEq) = (!d. f\\(Md d\\<>) = <> & \ + \(!x. ? ff:BufEq. f\\(Md d\\\\\\x) = d\\(ff\\x)))" (K [ + stac (BufEq_fix RS set_cong) 1, + rewtac BufEq_F_def, + Asm_full_simp_tac 1]); val Buf_f_empty = prove_forw "f:BufEq \\ f\\<> = <>" BufEq_unfold; val Buf_f_d = prove_forw "f:BufEq \\ f\\(Md d\\<>) = <>" BufEq_unfold; -val Buf_f_d_req = prove_forw - "f:BufEq \\ \\ff. ff:BufEq \\ f\\(Md d\\\\\\x) = d\\ff\\x" BufEq_unfold; +val Buf_f_d_req = prove_forw + "f:BufEq \\ \\ff. ff:BufEq \\ f\\(Md d\\\\\\x) = d\\ff\\x" BufEq_unfold; (**** BufAC_Asm ***************************************************************) -val mono_BufAC_Asm_F = prove_goalw thy [mono_def, BufAC_Asm_F_def] - "mono BufAC_Asm_F" (K [Fast_tac 1]); +val mono_BufAC_Asm_F = prove_goalw (the_context ()) [mono_def, BufAC_Asm_F_def] + "mono BufAC_Asm_F" (K [Fast_tac 1]); val BufAC_Asm_fix = mono_BufAC_Asm_F RS (BufAC_Asm_def RS def_gfp_unfold); -val BufAC_Asm_unfold = prove_goal thy "(s:BufAC_Asm) = (s = <> | (? d x. \ +val BufAC_Asm_unfold = prove_goal (the_context ()) "(s:BufAC_Asm) = (s = <> | (? d x. \ \ s = Md d\\x & (x = <> | (ft\\x = Def \\ & (rt\\x):BufAC_Asm))))" (K [ - stac (BufAC_Asm_fix RS set_cong) 1, - rewtac BufAC_Asm_F_def, - Asm_full_simp_tac 1]); + stac (BufAC_Asm_fix RS set_cong) 1, + rewtac BufAC_Asm_F_def, + Asm_full_simp_tac 1]); val BufAC_Asm_empty = prove_backw "<> :BufAC_Asm" BufAC_Asm_unfold []; val BufAC_Asm_d = prove_backw "Md d\\<>:BufAC_Asm" BufAC_Asm_unfold []; val BufAC_Asm_d_req = prove_backw "x:BufAC_Asm ==> Md d\\\\\\x:BufAC_Asm" - BufAC_Asm_unfold []; + BufAC_Asm_unfold []; val BufAC_Asm_prefix2 = prove_forw "a\\b\\s:BufAC_Asm ==> s:BufAC_Asm" - BufAC_Asm_unfold; + BufAC_Asm_unfold; (**** BBufAC_Cmt **************************************************************) -val mono_BufAC_Cmt_F = prove_goalw thy [mono_def, BufAC_Cmt_F_def] - "mono BufAC_Cmt_F" (K [Fast_tac 1]); +val mono_BufAC_Cmt_F = prove_goalw (the_context ()) [mono_def, BufAC_Cmt_F_def] + "mono BufAC_Cmt_F" (K [Fast_tac 1]); val BufAC_Cmt_fix = mono_BufAC_Cmt_F RS (BufAC_Cmt_def RS def_gfp_unfold); -val BufAC_Cmt_unfold = prove_goal thy "((s,t):BufAC_Cmt) = (!d x. \ +val BufAC_Cmt_unfold = prove_goal (the_context ()) "((s,t):BufAC_Cmt) = (!d x. \ \(s = <> --> t = <>) & \ \(s = Md d\\<> --> t = <>) & \ \(s = Md d\\\\\\x --> ft\\t = Def d & (x, rt\\t):BufAC_Cmt))" (K [ - stac (BufAC_Cmt_fix RS set_cong) 1, - rewtac BufAC_Cmt_F_def, - Asm_full_simp_tac 1]); + stac (BufAC_Cmt_fix RS set_cong) 1, + rewtac BufAC_Cmt_F_def, + Asm_full_simp_tac 1]); val BufAC_Cmt_empty = prove_backw "f:BufEq ==> (<>, f\\<>):BufAC_Cmt" - BufAC_Cmt_unfold [Buf_f_empty]; + BufAC_Cmt_unfold [Buf_f_empty]; -val BufAC_Cmt_d = prove_backw "f:BufEq ==> (a\\\\, f\\(a\\\\)):BufAC_Cmt" - BufAC_Cmt_unfold [Buf_f_d]; +val BufAC_Cmt_d = prove_backw "f:BufEq ==> (a\\\\, f\\(a\\\\)):BufAC_Cmt" + BufAC_Cmt_unfold [Buf_f_d]; -val BufAC_Cmt_d2 = prove_forw +val BufAC_Cmt_d2 = prove_forw "(Md d\\\\, f\\(Md d\\\\)):BufAC_Cmt ==> f\\(Md d\\\\) = \\" BufAC_Cmt_unfold; -val BufAC_Cmt_d3 = prove_forw +val BufAC_Cmt_d3 = prove_forw "(Md d\\\\\\x, f\\(Md d\\\\\\x)):BufAC_Cmt ==> (x, rt\\(f\\(Md d\\\\\\x))):BufAC_Cmt" - BufAC_Cmt_unfold; + BufAC_Cmt_unfold; -val BufAC_Cmt_d32 = prove_forw +val BufAC_Cmt_d32 = prove_forw "(Md d\\\\\\x, f\\(Md d\\\\\\x)):BufAC_Cmt ==> ft\\(f\\(Md d\\\\\\x)) = Def d" - BufAC_Cmt_unfold; + BufAC_Cmt_unfold; (**** BufAC *******************************************************************) @@ -114,33 +114,33 @@ Goalw [BufAC_def] "f\\BufAC \\ \\ff\\BufAC. \\x. f\\(Md d\\\\\\x) = d\\ff\\x"; by (rtac (ex_elim_lemma RS iffD2) 1); by Safe_tac; -by (fast_tac (claset() addIs [BufAC_Cmt_d32 RS Def_maximal, - monofun_cfun_arg, BufAC_Asm_empty RS BufAC_Asm_d_req]) 1); +by (fast_tac (claset() addIs [BufAC_Cmt_d32 RS Def_maximal, + monofun_cfun_arg, BufAC_Asm_empty RS BufAC_Asm_d_req]) 1); by (auto_tac (claset() addIs [BufAC_Cmt_d3, BufAC_Asm_d_req],simpset())); qed "BufAC_f_d_req"; (**** BufSt *******************************************************************) -val mono_BufSt_F = prove_goalw thy [mono_def, BufSt_F_def] - "mono BufSt_F" (K [Fast_tac 1]); +val mono_BufSt_F = prove_goalw (the_context ()) [mono_def, BufSt_F_def] + "mono BufSt_F" (K [Fast_tac 1]); val BufSt_P_fix = mono_BufSt_F RS (BufSt_P_def RS def_gfp_unfold); -val BufSt_P_unfold = prove_goal thy "(h:BufSt_P) = (!s. h s\\<> = <> & \ - \ (!d x. h \\ \\(Md d\\x) = h (Sd d)\\x & \ +val BufSt_P_unfold = prove_goal (the_context ()) "(h:BufSt_P) = (!s. h s\\<> = <> & \ + \ (!d x. h \\ \\(Md d\\x) = h (Sd d)\\x & \ \ (? hh:BufSt_P. h (Sd d)\\(\\\\x) = d\\(hh \\ \\x))))" (K [ - stac (BufSt_P_fix RS set_cong) 1, - rewtac BufSt_F_def, - Asm_full_simp_tac 1]); + stac (BufSt_P_fix RS set_cong) 1, + rewtac BufSt_F_def, + Asm_full_simp_tac 1]); -val BufSt_P_empty = prove_forw "h:BufSt_P ==> h s \\ <> = <>" - BufSt_P_unfold; +val BufSt_P_empty = prove_forw "h:BufSt_P ==> h s \\ <> = <>" + BufSt_P_unfold; val BufSt_P_d = prove_forw "h:BufSt_P ==> h \\ \\(Md d\\x) = h (Sd d)\\x" - BufSt_P_unfold; + BufSt_P_unfold; val BufSt_P_d_req = prove_forw "h:BufSt_P ==> \\hh\\BufSt_P. \ - \ h (Sd d)\\(\\ \\x) = d\\(hh \\ \\x)" - BufSt_P_unfold; + \ h (Sd d)\\(\\ \\x) = d\\(hh \\ \\x)" + BufSt_P_unfold; (**** Buf_AC_imp_Eq ***********************************************************) @@ -210,11 +210,11 @@ (**** alternative (not strictly) stronger version of Buf_Eq *******************) -val Buf_Eq_alt_imp_Eq = prove_goalw thy [BufEq_def,BufEq_alt_def] - "BufEq_alt \\ BufEq" (K [ - rtac gfp_mono 1, - rewtac BufEq_F_def, - Fast_tac 1]); +val Buf_Eq_alt_imp_Eq = prove_goalw (the_context ()) [BufEq_def,BufEq_alt_def] + "BufEq_alt \\ BufEq" (K [ + rtac gfp_mono 1, + rewtac BufEq_F_def, + Fast_tac 1]); (* direct proof of "BufEq \\ BufEq_alt" seems impossible *) @@ -224,10 +224,10 @@ by (fast_tac (claset() addEs [BufAC_f_d, BufAC_f_d_req]) 1); qed "Buf_AC_imp_Eq_alt"; -bind_thm ("Buf_Eq_imp_Eq_alt", +bind_thm ("Buf_Eq_imp_Eq_alt", subset_trans OF [Buf_Eq_imp_AC, Buf_AC_imp_Eq_alt]); -bind_thm ("Buf_Eq_alt_eq", +bind_thm ("Buf_Eq_alt_eq", subset_antisym OF [Buf_Eq_alt_imp_Eq, Buf_Eq_imp_Eq_alt]); @@ -245,11 +245,11 @@ Goalw [BufSt_def, BufSt_P_def] "BufEq <= BufSt"; by Safe_tac; -by (EVERY'[rename_tac "f", res_inst_tac +by (EVERY'[rename_tac "f", res_inst_tac [("x","\\s. case s of Sd d => \\ x. f\\(Md d\\x)| \\ => f")] bexI] 1); by ( Simp_tac 1); by (etac weak_coinduct_image 1); -by (rewtac BufSt_F_def); +by (rewtac BufSt_F_def); by (Simp_tac 1); by Safe_tac; by ( EVERY'[rename_tac "s", induct_tac "s"] 1); @@ -262,6 +262,3 @@ qed "Buf_Eq_imp_St"; bind_thm ("Buf_Eq_eq_St", Buf_St_imp_Eq RS (Buf_Eq_imp_St RS subset_antisym)); - - -