# HG changeset patch # User Cezary Kaliszyk # Date 1269638561 -3600 # Node ID aa999d263b1057281523e01f62520b09ec49bc33 # Parent d1293d30c65753a155d947ba5c137a0de2b5cb47 Modified abs_gen_sym and abs_gen_trans so it becomes usable in the proofs. diff -r d1293d30c657 -r aa999d263b10 Nominal/Abs.thy --- a/Nominal/Abs.thy Fri Mar 26 17:22:17 2010 +0100 +++ b/Nominal/Abs.thy Fri Mar 26 22:22:41 2010 +0100 @@ -587,24 +587,36 @@ by (simp_all add: fresh_zero_perm) lemma alpha_gen_sym: - assumes a: "R (p \ x) y \ R (- p \ y) x" + assumes a: "R (p \ x) y \ R y (p \ x)" + and b: "\x y p. R x y \ R (p \ x) (p \ y)" shows "(bs, x) \gen R f p (cs, y) \ (cs, y) \gen R f (- p) (bs, x)" and "(bs, x) \res R f p (cs, y) \ (cs, y) \res R f (- p) (bs, x)" and "(ds, x) \lst R f p (es, y) \ (es, y) \lst R f (- p) (ds, x)" + unfolding alphas fresh_star_def + apply (erule_tac [1] conjE)+ + apply (erule_tac [2] conjE)+ + apply (erule_tac [3] conjE)+ using a - unfolding alphas - unfolding fresh_star_def - by (auto simp add: fresh_minus_perm) - + apply (simp_all add: fresh_minus_perm) + apply (rotate_tac [!] -1) + apply (drule_tac [!] p="-p" in b) + apply (simp_all) + apply (metis permute_minus_cancel(2)) + apply (metis permute_minus_cancel(2)) + done lemma alpha_gen_trans: - assumes a: "\R (p \ x) y; R (q \ y) z\ \ R ((q + p) \ x) z" + assumes a: "(\c. R y c \ R (p \ x) c)" + and b: "R (p \ x) y" + and c: "R (q \ y) z" + and d: "\x y p. R x y \ R (p \ x) (p \ y)" shows "\(bs, x) \gen R f p (cs, y); (cs, y) \gen R f q (ds, z)\ \ (bs, x) \gen R f (q + p) (ds, z)" and "\(bs, x) \res R f p (cs, y); (cs, y) \res R f q (ds, z)\ \ (bs, x) \res R f (q + p) (ds, z)" and "\(es, x) \lst R f p (gs, y); (gs, y) \lst R f q (hs, z)\ \ (es, x) \lst R f (q + p) (hs, z)" - using a unfolding alphas unfolding fresh_star_def - by (simp_all add: fresh_plus_perm) + apply (simp_all add: fresh_plus_perm) + apply (metis a c d permute_minus_cancel(2))+ + done lemma alpha_gen_eqvt: assumes a: "R (q \ x) y \ R (p \ (q \ x)) (p \ y)" diff -r d1293d30c657 -r aa999d263b10 Nominal/Manual/Term5.thy --- a/Nominal/Manual/Term5.thy Fri Mar 26 17:22:17 2010 +0100 +++ b/Nominal/Manual/Term5.thy Fri Mar 26 22:22:41 2010 +0100 @@ -1,5 +1,5 @@ theory Term5 -imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp" "Abs" "Perm" "Fv" "Rsp" "../Attic/Prove" +imports "../Nominal2_Atoms" "../Nominal2_Eqvt" "../Nominal2_Supp" "../Abs" "../Perm" "../Fv" "../Rsp" begin atom_decl name @@ -23,7 +23,7 @@ print_theorems local_setup {* snd o define_fv_alpha (Datatype.the_info @{theory} "Term5.rtrm5") - [[[], [], [(SOME (@{term rbv5}, true), 0, 1)]], [[], []]] [(@{term rbv5}, 1, [[], [0, 2]])] *} + [[[], [], [(SOME (@{term rbv5}, true), 0, 1)]], [[], []]] [(@{term rbv5}, 1, [[], [(0,NONE), (2,SOME @{term rbv5})]])] *} print_theorems notation @@ -31,7 +31,7 @@ alpha_rlts ("_ \l _" [100, 100] 100) thm alpha_rtrm5_alpha_rlts_alpha_rbv5.intros -local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha5_inj}, []), (build_alpha_inj @{thms alpha_rtrm5_alpha_rlts_alpha_rbv5.intros} @{thms rtrm5.distinct rtrm5.inject rlts.distinct rlts.inject} @{thms alpha_rtrm5.cases alpha_rlts.cases alpha_rbv5.cases} ctxt)) ctxt)) *} +local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha5_inj}, []), (build_rel_inj @{thms alpha_rtrm5_alpha_rlts_alpha_rbv5.intros} @{thms rtrm5.distinct rtrm5.inject rlts.distinct rlts.inject} @{thms alpha_rtrm5.cases alpha_rlts.cases alpha_rbv5.cases} ctxt)) ctxt)) *} thm alpha5_inj lemma rbv5_eqvt[eqvt]: @@ -71,7 +71,17 @@ "(a \5 b \ b \5 a) \ (x \l y \ y \l x) \ (alpha_rbv5 x y \ alpha_rbv5 y x)" -apply (tactic {* symp_tac @{thm alpha_rtrm5_alpha_rlts_alpha_rbv5.induct} @{thms alpha5_inj} @{thms alpha5_eqvt} @{context} 1 *}) +apply (rule alpha_rtrm5_alpha_rlts_alpha_rbv5.induct) +apply (simp_all add: alpha5_inj) +apply (erule exE) +apply (rule_tac x="-pi" in exI) +apply (rule alpha_gen_sym) +apply (simp_all add: alphas) +apply (case_tac x) +apply (case_tac y) +apply (simp add: alpha5_eqvt) +apply clarify +apply simp done lemma alpha5_symp1: diff -r d1293d30c657 -r aa999d263b10 Nominal/Manual/Term5n.thy --- a/Nominal/Manual/Term5n.thy Fri Mar 26 17:22:17 2010 +0100 +++ b/Nominal/Manual/Term5n.thy Fri Mar 26 22:22:41 2010 +0100 @@ -1,5 +1,5 @@ theory Term5 -imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp" "Abs" "Perm" "Fv" "Rsp" "../Attic/Prove" +imports "../Nominal2_Atoms" "../Nominal2_Eqvt" "../Nominal2_Supp" "../Abs" "../Perm" "../Fv" "../Rsp" begin atom_decl name @@ -23,7 +23,7 @@ print_theorems local_setup {* snd o define_fv_alpha (Datatype.the_info @{theory} "Term5.rtrm5") - [[[], [], [(SOME (@{term rbv5}, false), 0, 1)]], [[], []]] [(@{term rbv5}, 1, [[], [0, 2]])] *} + [[[], [], [(SOME (@{term rbv5}, false), 0, 1)]], [[], []]] [(@{term rbv5}, 1, [[], [(0,NONE), (2,SOME @{term rbv5})]])] *} print_theorems notation @@ -31,7 +31,7 @@ alpha_rlts ("_ \l _" [100, 100] 100) thm alpha_rtrm5_alpha_rlts_alpha_rbv5.intros -local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha5_inj}, []), (build_alpha_inj @{thms alpha_rtrm5_alpha_rlts_alpha_rbv5.intros} @{thms rtrm5.distinct rtrm5.inject rlts.distinct rlts.inject} @{thms alpha_rtrm5.cases alpha_rlts.cases alpha_rbv5.cases} ctxt)) ctxt)) *} +local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha5_inj}, []), (build_rel_inj @{thms alpha_rtrm5_alpha_rlts_alpha_rbv5.intros} @{thms rtrm5.distinct rtrm5.inject rlts.distinct rlts.inject} @{thms alpha_rtrm5.cases alpha_rlts.cases alpha_rbv5.cases} ctxt)) ctxt)) *} thm alpha5_inj lemma rbv5_eqvt[eqvt]: @@ -65,13 +65,45 @@ "(a \5 b \ b \5 a) \ (x \l y \ y \l x) \ (alpha_rbv5 x y \ alpha_rbv5 y x)" -sorry +apply (rule alpha_rtrm5_alpha_rlts_alpha_rbv5.induct) +apply (simp_all add: alpha5_inj) +apply (erule conjE)+ +apply (erule exE) +apply (rule_tac x="-pi" in exI) +apply (rule alpha_gen_sym) +apply (simp add: alphas) +apply (simp add: alpha5_eqvt) +apply (simp add: alphas) +apply clarify +apply simp +done lemma alpha5_transp: "(a \5 b \ (\c. b \5 c \ a \5 c)) \ (x \l y \ (\z. y \l z \ x \l z)) \ (alpha_rbv5 k l \ (\m. alpha_rbv5 l m \ alpha_rbv5 k m))" -sorry +apply (rule alpha_rtrm5_alpha_rlts_alpha_rbv5.induct) +apply (rule_tac [!] allI) +apply (simp_all add: alpha5_inj) +apply (tactic {* (imp_elim_tac @{thms alpha_rtrm5.cases alpha_rlts.cases alpha_rbv5.cases} @{context}) 1 *}) +apply (simp_all add: alpha5_inj) +defer +apply (tactic {* (imp_elim_tac @{thms alpha_rtrm5.cases alpha_rlts.cases alpha_rbv5.cases} @{context}) 1 *}) +apply (simp_all add: alpha5_inj) +apply (tactic {* (imp_elim_tac @{thms alpha_rtrm5.cases alpha_rlts.cases alpha_rbv5.cases} @{context}) 1 *}) +apply (simp_all add: alpha5_inj) +apply (tactic {* (imp_elim_tac @{thms alpha_rtrm5.cases alpha_rlts.cases alpha_rbv5.cases} @{context}) 1 *}) +apply (simp_all add: alpha5_inj) +apply (erule conjE)+ +apply (erule exE)+ +apply (rule_tac x="pi + pia" in exI) +apply (rule alpha_gen_trans) +prefer 6 +apply assumption +apply (simp_all add: alphas alpha5_eqvt) +apply (clarify) +apply simp +done lemma alpha5_equivp: "equivp alpha_rtrm5"