Quot/Nominal/Abs.thy
changeset 1024 b3deb964ad26
parent 1021 bacf3584640e
child 1026 278253330b6a
--- a/Quot/Nominal/Abs.thy	Tue Feb 02 11:23:17 2010 +0100
+++ b/Quot/Nominal/Abs.thy	Tue Feb 02 11:56:37 2010 +0100
@@ -129,6 +129,26 @@
   apply(clarsimp)
   done
 
+lemma alpha_gen_atom_eqvt:
+  assumes a: "\<And>x. pi \<bullet> (f x) = f (pi \<bullet> x)"
+  and     b: "\<exists>pia. ({atom a}, t) \<approx>gen \<lambda>x1 x2. R x1 x2 \<and> R (pi \<bullet> x1) (pi \<bullet> x2) f pia ({atom b}, s)"
+  shows  "\<exists>pia. ({atom (pi \<bullet> a)}, pi \<bullet> t) \<approx>gen R f pia ({atom (pi \<bullet> b)}, pi \<bullet> s)"
+  using b apply -
+  apply(erule exE)
+  apply(rule_tac x="pi \<bullet> pia" in exI)
+  apply(simp add: alpha_gen.simps)
+  apply(erule conjE)+
+  apply(rule conjI)+
+  apply(rule_tac ?p1="- pi" in permute_eq_iff[THEN iffD1])
+  apply(simp add: a[symmetric] atom_eqvt eqvts)
+  apply(rule conjI)
+  apply(rule_tac ?p1="- pi" in fresh_star_permute_iff[THEN iffD1])
+  apply(simp add: a[symmetric] eqvts atom_eqvt)
+  apply(subst permute_eqvt[symmetric])
+  apply(simp)
+  done
+
+
 fun
   alpha_abs 
 where