diff -r b4404b13f775 -r 37c0d7953cba Nominal/Ex/Classical.thy --- a/Nominal/Ex/Classical.thy Wed Jun 29 17:01:09 2011 +0100 +++ b/Nominal/Ex/Classical.thy Wed Jun 29 19:21:26 2011 +0100 @@ -2,6 +2,7 @@ imports "../Nominal2" begin + (* example from Urban's PhD *) atom_decl name @@ -177,7 +178,9 @@ then have "(q \ (as \ supp x)) \* f (q \ (as \ supp x)) (q \ x) c" apply(simp add: fresh_star_eqvt set_eqvt) sorry (* perm? *) - then have "r \ (bs \ supp y) \* f (r \ (bs \ supp y)) (r \ y) c" using qq2 apply (simp add: inter_eqvt) + then have "r \ (bs \ supp y) \* f (r \ (bs \ supp y)) (r \ y) c" using qq2 + apply (simp add: inter_eqvt) + sorry (* rest similar reversing it other way around... *) show ?thesis sorry qed @@ -476,6 +479,8 @@ apply(blast) done + + end