# HG changeset patch # User Christian Urban # Date 1265291752 -3600 # Node ID d5d887bb986acd5c0a8b2302cc76de999c696a0a # Parent 090fa3f21380c521efe9fcada0c99bcae2c1a8c9# Parent afedef46d3abb10dbe3f4b6eb7fb0f398194494c merged diff -r 090fa3f21380 -r d5d887bb986a Quot/Nominal/Terms.thy --- a/Quot/Nominal/Terms.thy Thu Feb 04 14:55:21 2010 +0100 +++ b/Quot/Nominal/Terms.thy Thu Feb 04 14:55:52 2010 +0100 @@ -902,6 +902,20 @@ apply (simp add: permute_trm5_lts fresh_star_def) done +lemma lets_ok2: + "(Lt5 (Lcons x (Vr5 x) (Lcons y (Vr5 y) Lnil)) (Ap5 (Vr5 x) (Vr5 y))) = + (Lt5 (Lcons y (Vr5 y) (Lcons x (Vr5 x) Lnil)) (Ap5 (Vr5 x) (Vr5 y)))" +apply (subst alpha5_INJ) +apply (rule conjI) +apply (rule_tac x="0 :: perm" in exI) +apply (simp only: alpha_gen) +apply (simp add: permute_trm5_lts fresh_star_def) +apply (rule_tac x="(x \ y)" in exI) +apply (simp only: alpha_gen) +apply (simp add: permute_trm5_lts fresh_star_def) +done + + lemma lets_not_ok1: "x \ y \ (Lt5 (Lcons x (Vr5 x) (Lcons y (Vr5 y) Lnil)) (Ap5 (Vr5 x) (Vr5 y))) \ (Lt5 (Lcons y (Vr5 x) (Lcons x (Vr5 y) Lnil)) (Ap5 (Vr5 x) (Vr5 y)))" @@ -915,9 +929,27 @@ apply (simp only: alpha5_INJ(1)) done +lemma distinct_helper: + shows "\(rVr5 x \5 rAp5 y z)" + apply auto + apply (erule alpha5.cases) + apply (simp_all only: rtrm5.distinct) + done +lemma distinct_helper2: + shows "(Vr5 x) \ (Ap5 y z)" + by (lifting distinct_helper) - +lemma lets_nok: + "x \ y \ x \ z \ z \ y \ + (Lt5 (Lcons x (Ap5 (Vr5 z) (Vr5 z)) (Lcons y (Vr5 z) Lnil)) (Ap5 (Vr5 x) (Vr5 y))) \ + (Lt5 (Lcons y (Vr5 z) (Lcons x (Ap5 (Vr5 z) (Vr5 z)) Lnil)) (Ap5 (Vr5 x) (Vr5 y)))" +apply (subst alpha5_INJ) +apply (simp only: alpha_gen permute_trm5_lts fresh_star_def) +apply (subst alpha5_INJ(5)) +apply (subst alpha5_INJ(5)) +apply (simp add: distinct_helper2) +done text {* type schemes *} datatype ty =