Nominal/Fv.thy
changeset 1301 c145c380e195
parent 1288 0203cd5cfd6c
child 1302 d3101a5b9c4d
--- a/Nominal/Fv.thy	Tue Mar 02 11:04:49 2010 +0100
+++ b/Nominal/Fv.thy	Tue Mar 02 12:28:07 2010 +0100
@@ -303,18 +303,26 @@
 fun reflp_tac induct inj =
   rtac induct THEN_ALL_NEW
   asm_full_simp_tac (HOL_ss addsimps inj) THEN_ALL_NEW
-  TRY o REPEAT_ALL_NEW (CHANGED o rtac conjI) THEN_ALL_NEW
+(*  TRY o REPEAT_ALL_NEW (CHANGED o rtac conjI) THEN_ALL_NEW*)
   (rtac @{thm exI[of _ "0 :: perm"]} THEN'
    asm_full_simp_tac (HOL_ss addsimps
      @{thms alpha_gen fresh_star_def fresh_zero_perm permute_zero ball_triv}))
 *}
 
+lemma exi_neg: "\<exists>(pi :: perm). P pi \<Longrightarrow> (\<And>(p :: perm). P p \<Longrightarrow> Q (- p)) \<Longrightarrow> \<exists>pi. Q pi"
+apply (erule exE)
+apply (rule_tac x="-pi" in exI)
+by auto
+
 ML {*
 fun symp_tac induct inj eqvt =
-  ((rtac @{thm impI} THEN' etac induct) ORELSE' rtac induct) THEN_ALL_NEW
+  (((rtac impI THEN' etac induct) ORELSE' rtac induct) THEN_ALL_NEW
   asm_full_simp_tac (HOL_ss addsimps inj) THEN_ALL_NEW
-  TRY o REPEAT_ALL_NEW (CHANGED o rtac conjI) THEN_ALL_NEW
-  (etac @{thm alpha_gen_compose_sym} THEN' eresolve_tac eqvt)
+  (etac @{thm exi_neg} THEN' REPEAT o etac conjE THEN'
+  (TRY o REPEAT_ALL_NEW (CHANGED o rtac conjI))) THEN_ALL_NEW
+  (asm_full_simp_tac HOL_ss) THEN_ALL_NEW
+  (etac @{thm alpha_gen_compose_sym} THEN' 
+    (asm_full_simp_tac (HOL_ss addsimps (@{thm atom_eqvt} :: eqvt)))))
 *}
 
 ML {*
@@ -340,14 +348,23 @@
   )
 *}
 
+
+lemma exi_sum: "\<exists>(pi :: perm). P pi \<Longrightarrow> \<exists>(pi :: perm). Q pi \<Longrightarrow> (\<And>(p :: perm) (pi :: perm). P p \<Longrightarrow> Q pi \<Longrightarrow> R (pi + p)) \<Longrightarrow> \<exists>pi. R pi"
+apply (erule exE)+
+apply (rule_tac x="pia + pi" in exI)
+by auto
+
 ML {*
 fun transp_tac ctxt induct alpha_inj term_inj distinct cases eqvt =
   ((rtac impI THEN' etac induct) ORELSE' rtac induct) THEN_ALL_NEW
   (TRY o rtac allI THEN' imp_elim_tac cases ctxt) THEN_ALL_NEW
   (
-    asm_full_simp_tac (HOL_ss addsimps alpha_inj @ term_inj @ distinct) THEN'
-    TRY o REPEAT_ALL_NEW (CHANGED o rtac conjI) THEN_ALL_NEW
-    (etac @{thm alpha_gen_compose_trans} THEN' RANGE [atac, eresolve_tac eqvt])
+    asm_full_simp_tac (HOL_ss addsimps alpha_inj @ term_inj @ distinct)
+    THEN_ALL_NEW (etac @{thm exi_sum} THEN' RANGE [atac]) THEN_ALL_NEW
+    (REPEAT o etac conjE THEN' (TRY o REPEAT_ALL_NEW (CHANGED o rtac conjI)))
+    THEN_ALL_NEW (asm_full_simp_tac HOL_ss) THEN_ALL_NEW
+    (etac @{thm alpha_gen_compose_trans} THEN' RANGE[atac]) THEN_ALL_NEW
+    (asm_full_simp_tac (HOL_ss addsimps (@{thm atom_eqvt} :: eqvt)))
   )
 *}