Experiments with proving weird transp
authorCezary Kaliszyk <kaliszyk@in.tum.de>
Wed, 03 Mar 2010 17:49:34 +0100
changeset 1335 c3dfd4193b42
parent 1334 80441e27dfd6
child 1336 6ab8c46b3b4b
Experiments with proving weird transp
Nominal/Test.thy
--- a/Nominal/Test.thy	Wed Mar 03 17:47:29 2010 +0100
+++ b/Nominal/Test.thy	Wed Mar 03 17:49:34 2010 +0100
@@ -1,5 +1,5 @@
 theory Test
-imports "Parser"
+imports "Parser" "../Attic/Prove"
 begin
 
 text {* weirdo example from Peter Sewell's bestiary *}
@@ -14,6 +14,66 @@
 thm alpha_weird_raw.intros[no_vars]
 thm fv_weird_raw.simps[no_vars]
 
+thm eqvts
+
+local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding weird_inj}, []), (build_alpha_inj @{thms alpha_weird_raw.intros} @{thms weird_raw.distinct weird_raw.inject} @{thms alpha_weird_raw.cases} ctxt)) ctxt)) *}
+thm weird_inj
+
+local_setup {*
+(fn ctxt => snd (Local_Theory.note ((@{binding alpha_eqvt}, []),
+build_alpha_eqvts [@{term alpha_weird_raw}] [@{term "permute :: perm \<Rightarrow> weird_raw \<Rightarrow> weird_raw"}] @{thms permute_weird_raw.simps weird_inj} @{thm alpha_weird_raw.induct} ctxt) ctxt)) *}
+
+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
+  )
+*}
+
+
+prove {* (snd o snd) (build_alpha_refl_gl [@{term alpha_weird_raw}] ("x","y","z")) *}
+ML_prf {*
+fun transp_tac ctxt induct alpha_inj term_inj distinct cases eqvt =
+  ind_tac induct THEN_ALL_NEW
+  (TRY o rtac allI THEN' imp_elim_tac cases ctxt) THEN_ALL_NEW
+  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
+*}
+
+apply (tactic {* transp_tac @{context} @{thm alpha_weird_raw.induct} @{thms weird_inj}  @{thms weird_raw.inject} @{thms weird_raw.distinct} @{thms alpha_weird_raw.cases} @{thms alpha_eqvt} 1 *})
+apply (simp_all)
+apply (erule alpha_gen_compose_trans)
+apply assumption
+apply (simp add: alpha_eqvt)
+defer defer
+apply (erule alpha_gen_compose_trans)
+apply assumption
+apply (simp add: alpha_eqvt)
+apply (subgoal_tac "pia + pb + (pib + pa) = (pia + pib) + (pb + pa)")
+apply simp
+apply (erule alpha_gen_compose_trans)
+apply assumption
+apply (simp add: alpha_eqvt)
+sorry
+
+
 (*
 abbreviation "WBind \<equiv> WBind_raw"
 abbreviation "WP \<equiv> WP_raw"