# HG changeset patch # User Christian Urban # Date 1259383756 -3600 # Node ID d1aa26ecfecdaeaa6392b2bf9a37e67b856aa065 # Parent 3359033eff798b5789c767308d98147d762ca9bb# Parent 1c245f6911ddae9b97aa475c5d449983f24786c2 merged diff -r 1c245f6911dd -r d1aa26ecfecd FSet.thy --- a/FSet.thy Sat Nov 28 05:43:18 2009 +0100 +++ b/FSet.thy Sat Nov 28 05:49:16 2009 +0100 @@ -445,6 +445,7 @@ done ML {* fun r_mk_comb_tac_fset lthy = r_mk_comb_tac lthy rty [quot] [rel_refl] [trans2] rsp_thms *} +ML {* fun r_mk_comb_tac_fset' lthy = r_mk_comb_tac' lthy rty [quot] [rel_refl] [trans2] rsp_thms *} (* Construction site starts here *) lemma "P (x :: 'a list) (EMPTY :: 'c fset) \ (\e t. P x t \ P x (INSERT e t)) \ P x l" @@ -462,6 +463,8 @@ apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *}) apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *}) apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *}) +apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *}) +apply (tactic {* (r_mk_comb_tac_fset @{context}) 1 *}) apply (tactic {* (APPLY_RSP_TAC rty @{context}) 1 *}) apply (rule IDENTITY_QUOTIENT) apply (rule IDENTITY_QUOTIENT) diff -r 1c245f6911dd -r d1aa26ecfecd QuotMain.thy --- a/QuotMain.thy Sat Nov 28 05:43:18 2009 +0100 +++ b/QuotMain.thy Sat Nov 28 05:49:16 2009 +0100 @@ -827,6 +827,24 @@ | _ => no_tac) *} +ML {* +val ball_rsp_tac = + Subgoal.FOCUS (fn {concl, ...} => + case (term_of concl) of + (_ $ (_ $ (Const (@{const_name Ball}, _) $ _) $ (Const (@{const_name Ball}, _) $ _))) => + rtac @{thm FUN_REL_I} 1 + |_ => no_tac) +*} + +ML {* +val bex_rsp_tac = + Subgoal.FOCUS (fn {concl, context = ctxt, ...} => + case (term_of concl) of + (_ $ (_ $ (Const (@{const_name Bex}, _) $ _) $ (Const (@{const_name Bex}, _) $ _))) => + rtac @{thm FUN_REL_I} 1 + | _ => no_tac) +*} + ML {* (* Legacy *) fun needs_lift (rty as Type (rty_s, _)) ty = case ty of @@ -853,6 +871,8 @@ *} ML {* +<<<<<<< variant A +>>>>>>> variant B val ball_rsp_tac = Subgoal.FOCUS (fn {concl, context = ctxt, ...} => case (term_of concl) of @@ -881,6 +901,32 @@ *} ML {* +####### Ancestor +val ball_rsp_tac = + Subgoal.FOCUS (fn {concl, context = ctxt, ...} => + case (term_of concl) of + (_ $ (_ $ (Const (@{const_name Ball}, _) $ _) $ (Const (@{const_name Ball}, _) $ _))) => + ((simp_tac (HOL_ss addsimps @{thms FUN_REL.simps})) + THEN' rtac @{thm allI} THEN' rtac @{thm allI} THEN' rtac @{thm impI} + THEN' instantiate_tac @{thm RES_FORALL_RSP} ctxt THEN' + (simp_tac (HOL_ss addsimps @{thms FUN_REL.simps}))) 1 + |_ => no_tac) +*} + +ML {* +val bex_rsp_tac = + Subgoal.FOCUS (fn {concl, context = ctxt, ...} => + case (term_of concl) of + (_ $ (_ $ (Const (@{const_name Bex}, _) $ _) $ (Const (@{const_name Bex}, _) $ _))) => + ((simp_tac (HOL_ss addsimps @{thms FUN_REL.simps})) + THEN' rtac @{thm allI} THEN' rtac @{thm allI} THEN' rtac @{thm impI} + THEN' instantiate_tac @{thm RES_EXISTS_RSP} ctxt THEN' + (simp_tac (HOL_ss addsimps @{thms FUN_REL.simps}))) 1 + | _ => no_tac) +*} + +ML {* +======= end fun r_mk_comb_tac ctxt rty quot_thms rel_refl trans2 rsp_thms = (FIRST' [resolve_tac trans2, LAMBDA_RES_TAC ctxt, @@ -941,11 +987,16 @@ (* (R1 ===> R2) (\x\) (\\) \ \R1 x y\ \ R2 (\x) (\y) *) DT ctxt "2" (LAMBDA_RES_TAC ctxt), - (* R (Ball\) (Ball\) \ R (\) (\) *) + (* = (Ball\) (Ball\) \ = (\) (\) *) DT ctxt "3" (rtac @{thm RES_FORALL_RSP}), + (* (R1 ===> R2) (Ball\) (Ball\) \ \R1 x y\ \ R2 (Ball\x) (Ball\y) *) DT ctxt "4" (ball_rsp_tac ctxt), + + (* = (Bex\) (Bex\) \ = (\) (\) *) DT ctxt "5" (rtac @{thm RES_EXISTS_RSP}), + + (* (R1 ===> R2) (Bex\) (Bex\) \ \R1 x y\ \ R2 (Bex\x) (Bex\y) *) DT ctxt "6" (bex_rsp_tac ctxt), (* respectfulness of ops *)