include set_simps and append_simps in fv_rsp
authorCezary Kaliszyk <kaliszyk@in.tum.de>
Wed, 12 May 2010 16:57:01 +0200
changeset 2116 ce228f7b2b72
parent 2115 9b109ef7bf47
child 2119 238062c4c9f2
include set_simps and append_simps in fv_rsp
Nominal/Ex/Classical.thy
Nominal/Ex/ExPS8.thy
Nominal/Rsp.thy
--- a/Nominal/Ex/Classical.thy	Wed May 12 16:39:10 2010 +0200
+++ b/Nominal/Ex/Classical.thy	Wed May 12 16:57:01 2010 +0200
@@ -13,7 +13,6 @@
 atom_decl coname
 
 ML {* val _ = cheat_equivp := true *}
-ML {* val _ = cheat_fv_rsp := true *}
 
 nominal_datatype trm =
    Ax "name" "coname"
--- a/Nominal/Ex/ExPS8.thy	Wed May 12 16:39:10 2010 +0200
+++ b/Nominal/Ex/ExPS8.thy	Wed May 12 16:57:01 2010 +0200
@@ -7,6 +7,7 @@
 atom_decl name
 
 ML {* val _ = cheat_fv_rsp := true *}
+ML {* val _ = cheat_equivp := true *}
 ML {* val _ = cheat_alpha_bn_rsp := true *}
 
 nominal_datatype exp =
--- a/Nominal/Rsp.thy	Wed May 12 16:39:10 2010 +0200
+++ b/Nominal/Rsp.thy	Wed May 12 16:57:01 2010 +0200
@@ -62,7 +62,7 @@
 fun fvbv_rsp_tac induct fvbv_simps ctxt =
   rtac induct THEN_ALL_NEW
   (TRY o rtac @{thm TrueI}) THEN_ALL_NEW
-  asm_full_simp_tac (HOL_ss addsimps (@{thms prod_fv.simps prod_rel.simps alphas} @ fvbv_simps)) THEN_ALL_NEW
+  asm_full_simp_tac (HOL_ss addsimps (@{thms prod_fv.simps prod_rel.simps set.simps append.simps alphas} @ fvbv_simps)) THEN_ALL_NEW
   REPEAT o eresolve_tac [conjE, exE] THEN_ALL_NEW
   asm_full_simp_tac (HOL_ss addsimps fvbv_simps) THEN_ALL_NEW
   TRY o blast_tac (claset_of ctxt)