Nominal/Fv.thy
changeset 1352 cad5f3851569
parent 1339 5256f256edd8
child 1357 42b7abf779ec
--- a/Nominal/Fv.thy	Thu Mar 04 15:31:21 2010 +0100
+++ b/Nominal/Fv.thy	Thu Mar 04 15:31:34 2010 +0100
@@ -1,5 +1,5 @@
 theory Fv
-imports "Nominal2_Atoms" "Abs" "Perm" (* For testing *)
+imports "Nominal2_Atoms" "Abs" "Perm" "Rsp"
 begin
 
 (* Bindings are given as a list which has a length being equal
@@ -184,8 +184,8 @@
               val pi_supps = map ((curry op $) @{term "supp :: perm \<Rightarrow> atom set"}) rpis;
               val pi_supps_eq = HOLogic.mk_eq (mk_inter pi_supps, @{term "{} :: atom set"})
             in
-              if length pi_supps > 1 then
-                HOLogic.mk_conj (alpha_gen, pi_supps_eq) else alpha_gen
+              (*if length pi_supps > 1 then
+                HOLogic.mk_conj (alpha_gen, pi_supps_eq) else*) alpha_gen
             (* TODO Add some test that is makes sense *)
             end else @{term "True"}
         end
@@ -312,29 +312,34 @@
 *}
 
 ML {*
-fun reflp_tac induct inj =
+fun reflp_tac induct inj ctxt =
   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*)
-  (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}))
+  simp_tac ((mk_minimal_ss ctxt) addsimps inj) THEN_ALL_NEW
+  split_conjs THEN_ALL_NEW REPEAT o rtac @{thm exI[of _ "0 :: perm"]}
+  THEN_ALL_NEW split_conjs THEN_ALL_NEW asm_full_simp_tac (HOL_ss addsimps
+     @{thms alpha_gen fresh_star_def fresh_zero_perm permute_zero ball_triv
+       add_0_left supp_zero_perm Int_empty_left})
 *}
 
+
 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 impI THEN' etac induct) ORELSE' rtac induct) THEN_ALL_NEW
-  asm_full_simp_tac (HOL_ss addsimps inj) THEN_ALL_NEW
-  (REPEAT o etac conjE THEN' 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)))))
+fun symp_tac induct inj eqvt ctxt =
+  ind_tac induct THEN_ALL_NEW
+  simp_tac ((mk_minimal_ss ctxt) addsimps inj) THEN_ALL_NEW split_conjs
+  THEN_ALL_NEW
+  REPEAT o etac @{thm exi_neg}
+  THEN_ALL_NEW
+  split_conjs THEN_ALL_NEW
+  asm_full_simp_tac (HOL_ss addsimps @{thms supp_minus_perm minus_add[symmetric]}) THEN_ALL_NEW
+  (rtac @{thm alpha_gen_compose_sym} THEN' RANGE [
+    (asm_full_simp_tac (HOL_ss addsimps @{thms plus_perm_eq})),
+    (asm_full_simp_tac (HOL_ss addsimps (eqvt @ all_eqvts ctxt)))
+  ])
 *}
 
 ML {*
@@ -367,17 +372,38 @@
 by auto
 
 ML {*
+fun is_ex (Const ("Ex", _) $ Abs _) = true
+  | is_ex _ = false;
+*}
+
+ML {*
+fun eetac rule = Subgoal.FOCUS_PARAMS 
+  (fn (focus) =>
+     let
+       val concl = #concl focus
+       val prems = Logic.strip_imp_prems (term_of concl)
+       val exs = filter (fn x => is_ex (HOLogic.dest_Trueprop x)) prems
+       val cexs = map (SOME o (cterm_of (ProofContext.theory_of (#context focus)))) exs
+       val thins = map (fn cex => Drule.instantiate' [] [cex] Drule.thin_rl) cexs
+     in
+     (etac rule THEN' RANGE[
+        atac,
+        eresolve_tac thins
+     ]) 1
+     end
+  )
+*}
+
+ML {*
 fun transp_tac ctxt induct alpha_inj term_inj distinct cases eqvt =
-  ((rtac impI THEN' etac induct) ORELSE' rtac induct) THEN_ALL_NEW
+  ind_tac 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_ALL_NEW (REPEAT o etac conjE THEN' 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)))
-  )
+  asm_full_simp_tac ((mk_minimal_ss ctxt) addsimps alpha_inj) THEN_ALL_NEW
+  split_conjs THEN_ALL_NEW REPEAT o (eetac @{thm exi_sum} ctxt) THEN_ALL_NEW split_conjs
+  THEN_ALL_NEW (asm_full_simp_tac (HOL_ss addsimps (term_inj @ distinct)))
+  THEN_ALL_NEW split_conjs THEN_ALL_NEW
+  TRY o (etac @{thm alpha_gen_compose_trans} THEN' RANGE[atac]) THEN_ALL_NEW
+  (asm_full_simp_tac (HOL_ss addsimps (all_eqvts ctxt @ eqvt @ term_inj @ distinct)))
 *}
 
 lemma transp_aux:
@@ -400,8 +426,8 @@
 let
   val ([x, y, z], ctxt') = Variable.variant_fixes ["x","y","z"] ctxt;
   val (reflg, (symg, transg)) = build_alpha_refl_gl alphas (x, y, z)
-  fun reflp_tac' _ = reflp_tac term_induct alpha_inj 1;
-  fun symp_tac' _ = symp_tac alpha_induct alpha_inj eqvt 1;
+  fun reflp_tac' _ = reflp_tac term_induct alpha_inj ctxt 1;
+  fun symp_tac' _ = symp_tac alpha_induct alpha_inj eqvt ctxt 1;
   fun transp_tac' _ = transp_tac ctxt alpha_induct alpha_inj term_inj distinct cases eqvt 1;
   val reflt = Goal.prove ctxt' [] [] reflg reflp_tac';
   val symt = Goal.prove ctxt' [] [] symg symp_tac';