diff -r 9cec4269b7f9 -r c9d3dda79fe3 Nominal/Manual/LamEx.thy --- a/Nominal/Manual/LamEx.thy Fri Mar 26 10:55:13 2010 +0100 +++ b/Nominal/Manual/LamEx.thy Fri Mar 26 16:20:39 2010 +0100 @@ -95,8 +95,8 @@ \ rLam a t \a rLam b s" lemma a3_inverse: - assumes "rLam a t \ rLam b s" - shows "\pi. (rfv t - {atom a} = rfv s - {atom b} \ (rfv t - {atom a})\* pi \ (pi \ t) \ s)" + assumes "rLam a t \a rLam b s" + shows "\pi. (rfv t - {atom a} = rfv s - {atom b} \ (rfv t - {atom a})\* pi \ (pi \ t) \a s)" using assms apply(erule_tac alpha.cases) apply(auto) @@ -104,7 +104,7 @@ text {* should be automatic with new version of eqvt-machinery *} lemma alpha_eqvt: - shows "t \ s \ (pi \ t) \ (pi \ s)" + shows "t \a s \ (pi \ t) \a (pi \ s)" apply(induct rule: alpha.induct) apply(simp add: a1) apply(simp add: a2) @@ -126,7 +126,7 @@ done lemma alpha_refl: - shows "t \ t" + shows "t \a t" apply(induct t rule: rlam.induct) apply(simp add: a1) apply(simp add: a2) @@ -136,7 +136,7 @@ done lemma alpha_sym: - shows "t \ s \ s \ t" + shows "t \a s \ s \a t" apply(induct rule: alpha.induct) apply(simp add: a1) apply(simp add: a2) @@ -152,7 +152,7 @@ done lemma alpha_trans: - shows "t1 \ t2 \ t2 \ t3 \ t1 \ t3" + shows "t1 \a t2 \ t2 \a t3 \ t1 \a t3" apply(induct arbitrary: t3 rule: alpha.induct) apply(erule alpha.cases) apply(simp_all) @@ -188,22 +188,22 @@ done lemma alpha_rfv: - shows "t \ s \ rfv t = rfv s" + shows "t \a s \ rfv t = rfv s" apply(induct rule: alpha.induct) apply(simp_all) done inductive - alpha2 :: "rlam \ rlam \ bool" ("_ \2 _" [100, 100] 100) + alpha2 :: "rlam \ rlam \ bool" ("_ \a2 _" [100, 100] 100) where - a21: "a = b \ (rVar a) \2 (rVar b)" -| a22: "\t1 \2 t2; s1 \2 s2\ \ rApp t1 s1 \2 rApp t2 s2" -| a23: "(a = b \ t \2 s) \ (a \ b \ ((a \ b) \ t) \2 s \ atom b \ rfv t)\ rLam a t \2 rLam b s" + a21: "a = b \ (rVar a) \a2 (rVar b)" +| a22: "\t1 \a2 t2; s1 \a2 s2\ \ rApp t1 s1 \a2 rApp t2 s2" +| a23: "(a = b \ t \a2 s) \ (a \ b \ ((a \ b) \ t) \a2 s \ atom b \ rfv t)\ rLam a t \a2 rLam b s" lemma fv_vars: fixes a::name assumes a1: "\x \ rfv t - {atom a}. pi \ x = x" - shows "(pi \ t) \2 ((a \ (pi \ a)) \ t)" + shows "(pi \ t) \a2 ((a \ (pi \ a)) \ t)" using a1 apply(induct t) apply(auto) @@ -222,8 +222,8 @@ lemma - assumes a1: "t \2 s" - shows "t \ s" + assumes a1: "t \a2 s" + shows "t \a s" using a1 apply(induct) apply(rule alpha.intros) @@ -254,8 +254,8 @@ sorry lemma - assumes a1: "t \ s" - shows "t \2 s" + assumes a1: "t \a s" + shows "t \a2 s" using a1 apply(induct) apply(rule alpha2.intros) @@ -445,12 +445,12 @@ done lemma rlam_distinct: - shows "\(rVar nam \ rApp rlam1' rlam2')" - and "\(rApp rlam1' rlam2' \ rVar nam)" - and "\(rVar nam \ rLam nam' rlam')" - and "\(rLam nam' rlam' \ rVar nam)" - and "\(rApp rlam1 rlam2 \ rLam nam' rlam')" - and "\(rLam nam' rlam' \ rApp rlam1 rlam2)" + shows "\(rVar nam \a rApp rlam1' rlam2')" + and "\(rApp rlam1' rlam2' \a rVar nam)" + and "\(rVar nam \a rLam nam' rlam')" + and "\(rLam nam' rlam' \a rVar nam)" + and "\(rApp rlam1 rlam2 \a rLam nam' rlam')" + and "\(rLam nam' rlam' \a rApp rlam1 rlam2)" apply auto apply (erule alpha.cases) apply (simp_all only: rlam.distinct)