Nominal/Abs.thy
changeset 1845 b7423c6b5564
parent 1841 fcc660ece040
child 1857 591cc76da570
equal deleted inserted replaced
1844:c123419a4eb0 1845:b7423c6b5564
   760   using a
   760   using a
   761   unfolding alphas fresh_star_def
   761   unfolding alphas fresh_star_def
   762   by (simp_all add: fresh_plus_perm)
   762   by (simp_all add: fresh_plus_perm)
   763 
   763 
   764 
   764 
   765 lemma alpha_gen_eqvt[eqvt]:
   765 lemma alpha_gen_eqvt(*[eqvt]*):
   766   shows "(bs, x) \<approx>gen R f q (cs, y) \<Longrightarrow> (p \<bullet> bs, p \<bullet> x) \<approx>gen (p \<bullet> R) (p \<bullet> f) (p \<bullet> q) (p \<bullet> cs, p \<bullet> y)"
   766   shows "(bs, x) \<approx>gen R f q (cs, y) \<Longrightarrow> (p \<bullet> bs, p \<bullet> x) \<approx>gen (p \<bullet> R) (p \<bullet> f) (p \<bullet> q) (p \<bullet> cs, p \<bullet> y)"
   767   and   "(bs, x) \<approx>res R f q (cs, y) \<Longrightarrow> (p \<bullet> bs, p \<bullet> x) \<approx>res (p \<bullet> R) (p \<bullet> f) (p \<bullet> q) (p \<bullet> cs, p \<bullet> y)"
   767   and   "(bs, x) \<approx>res R f q (cs, y) \<Longrightarrow> (p \<bullet> bs, p \<bullet> x) \<approx>res (p \<bullet> R) (p \<bullet> f) (p \<bullet> q) (p \<bullet> cs, p \<bullet> y)"
   768   and   "(ds, x) \<approx>lst R f q (es, y) \<Longrightarrow> (p \<bullet> ds, p \<bullet> x) \<approx>lst (p \<bullet> R) (p \<bullet> f) (p \<bullet> q) (p \<bullet> es, p \<bullet> y)" 
   768   and   "(ds, x) \<approx>lst R f q (es, y) \<Longrightarrow> (p \<bullet> ds, p \<bullet> x) \<approx>lst (p \<bullet> R) (p \<bullet> f) (p \<bullet> q) (p \<bullet> es, p \<bullet> y)" 
   769   unfolding alphas
   769   unfolding alphas
   770   unfolding permute_eqvt[symmetric]
   770   unfolding permute_eqvt[symmetric]