# HG changeset patch # User Cezary Kaliszyk # Date 1333107555 -7200 # Node ID 556fcd0e5fcb724046bd0c5ab6db41152e6be8fe # Parent 31bc3e2e80bfbf224162066cd2fc1413f6b8f0b6 Finish all subgoals about Aux. diff -r 31bc3e2e80bf -r 556fcd0e5fcb Nominal/Ex/AuxNoFCB.thy --- a/Nominal/Ex/AuxNoFCB.thy Fri Mar 30 09:11:30 2012 +0200 +++ b/Nominal/Ex/AuxNoFCB.thy Fri Mar 30 13:39:15 2012 +0200 @@ -268,9 +268,19 @@ apply simp apply simp apply simp - (* The last subgoal would be obvious if we'd define it with a recusor - and have eqvt for the function *) - sorry + apply clarify + apply (rule_tac s="(x \ xa) \ swapequal_sumC ((hla \ x) \ la, (hra \ x) \ ra, ta)" in trans) + apply (simp only: permute_pure) + apply (simp add: eqvt_at_def fresh_Pair_elim) + apply (simp add: flip_fresh_fresh) + apply (subgoal_tac "(x \ xa) \ (hla \ x) \ la = (hla \ xa) \ la") + apply (subgoal_tac "(x \ xa) \ (hra \ x) \ ra = (hra \ xa) \ ra") + apply simp + apply (subst permute_eqvt) + apply (simp add: flip_fresh_fresh flip_eqvt) + apply (subst permute_eqvt) + apply (simp add: flip_fresh_fresh flip_eqvt) + done termination (eqvt) by lexicographic_order @@ -341,7 +351,37 @@ lemma swapequal_reorder: " a \ x \ a \ y \ b \ x \ b \ y \ swapequal t s ((x, y) # (a, b) # xs) = swapequal t s ((a, b) # (x, y) # xs)" - sorry + apply (rule_tac x="(a, b, x, y, t, s, xs)" and ?'a="name" in obtain_fresh) + apply (rule_tac x="(a, b, x, y, t, s, xs, aa)" and ?'a="name" in obtain_fresh) + apply (rename_tac f g) + apply (simp add: fresh_Pair_elim fresh_at_base) + apply (subst swapequal.simps) + apply (auto simp add: fresh_Pair fresh_Cons fresh_at_base)[1] + apply (subgoal_tac "(x \ f) \ atom g \ t") + prefer 2 + apply (simp add: flip_at_base_simps fresh_at_base flip_def) + apply (subst swapequal.simps) + apply (simp add: fresh_Pair fresh_Cons fresh_permute_left) + apply rule apply assumption + apply (simp add: flip_at_base_simps fresh_at_base flip_def) + apply (subst swapequal.simps) + apply (simp add: fresh_Pair fresh_Cons fresh_at_base) + apply rule apply (rotate_tac 12) + apply assumption + apply (simp add: fresh_Pair fresh_Cons fresh_at_base) + apply (subst swapequal.simps) + apply (simp add: fresh_Pair fresh_Cons fresh_at_base fresh_permute_left) + apply (subgoal_tac "(a \ g) \ atom f \ t") + prefer 2 + apply (simp add: flip_at_base_simps fresh_at_base flip_def) + apply rule apply assumption + apply (simp add: flip_at_base_simps fresh_at_base flip_def) + apply (subgoal_tac "(a \ g) \ (x \ f) \ t = (x \ f) \ (a \ g) \ t") + apply (subgoal_tac "(b \ g) \ (y \ f) \ s = (y \ f) \ (b \ g) \ s") + apply simp + apply (smt flip_at_base_simps(3) flip_at_simps(1) flip_eqvt permute_eqvt) + apply (smt flip_at_base_simps(3) flip_at_simps(1) flip_eqvt permute_eqvt) + done lemma swapequal_lambda: assumes "distinct xs \ atom x \ xs \ atom y \ xs"