# HG changeset patch # User Christian Urban # Date 1271276632 -7200 # Node ID b7423c6b55645a5b6d52cf85400cd4d85ceafa14 # Parent c123419a4eb01c86a7bdd271323cf113e179b94f deleted offending [eqvt]-attribute in Abs; Lambda works again, but there is now a problem in CoreHaskell diff -r c123419a4eb0 -r b7423c6b5564 Nominal/Abs.thy --- a/Nominal/Abs.thy Wed Apr 14 20:21:11 2010 +0200 +++ b/Nominal/Abs.thy Wed Apr 14 22:23:52 2010 +0200 @@ -762,7 +762,7 @@ by (simp_all add: fresh_plus_perm) -lemma alpha_gen_eqvt[eqvt]: +lemma alpha_gen_eqvt(*[eqvt]*): shows "(bs, x) \gen R f q (cs, y) \ (p \ bs, p \ x) \gen (p \ R) (p \ f) (p \ q) (p \ cs, p \ y)" and "(bs, x) \res R f q (cs, y) \ (p \ bs, p \ x) \res (p \ R) (p \ f) (p \ q) (p \ cs, p \ y)" and "(ds, x) \lst R f q (es, y) \ (p \ ds, p \ x) \lst (p \ R) (p \ f) (p \ q) (p \ es, p \ y)" diff -r c123419a4eb0 -r b7423c6b5564 Nominal/Ex/Lambda.thy --- a/Nominal/Ex/Lambda.thy Wed Apr 14 20:21:11 2010 +0200 +++ b/Nominal/Ex/Lambda.thy Wed Apr 14 22:23:52 2010 +0200 @@ -9,6 +9,8 @@ | App "lam" "lam" | Lam x::"name" l::"lam" bind x in l +thm lam.fv +thm lam.supp lemmas supp_fn' = lam.fv[simplified lam.supp] declare lam.perm[eqvt] @@ -151,24 +153,17 @@ | t'_App[intro]: "\\ |\ t1 : T1\T2; \ |\ t2 : T1\ \ \ |\ App t1 t2 : T2" | t'_Lam[intro]: "\atom x\\;(x,T1)#\ |\ t : T2\ \ \ |\ Lam x t : T1\T2" -use "../../Nominal-General/nominal_eqvt.ML" - equivariance valid equivariance typing +equivariance typing' +equivariance typing2' +equivariance typing'' thm valid.eqvt thm typing.eqvt thm eqvts thm eqvts_raw -equivariance typing' -equivariance typing2' -equivariance typing'' - -ML {* -fun mk_minus p = - Const (@{const_name "uminus"}, @{typ "perm => perm"}) $ p -*} inductive tt :: "('a \ 'a \ bool) \ ('a \ 'a \ bool)"