# HG changeset patch # User Cezary Kaliszyk # Date 1269331997 -3600 # Node ID a7e60da429e27832c55ecf2f54f3be33d4cc467c # Parent 5f0bb35114c35539ecd062e617701cf394525bb7 Move Let properties to ExLet diff -r 5f0bb35114c3 -r a7e60da429e2 Nominal/ExLet.thy --- a/Nominal/ExLet.thy Tue Mar 23 09:06:28 2010 +0100 +++ b/Nominal/ExLet.thy Tue Mar 23 09:13:17 2010 +0100 @@ -11,15 +11,15 @@ Vr "name" | Ap "trm" "trm" | Lm x::"name" t::"trm" bind x in t -| Lt a::"lts" t::"trm" bind "bv a" in t +| Lt a::"lts" t::"trm" bind "bn a" in t and lts = - Nil -| Cons "name" "trm" "lts" + Lnil +| Lcons "name" "trm" "lts" binder bn where - "bn Nil = {}" -| "bn (Cons x t l) = {atom x} \ (bn l)" + "bn Lnil = {}" +| "bn (Lcons x t l) = {atom x} \ (bn l)" thm trm_lts.fv thm trm_lts.eq_iff @@ -29,6 +29,43 @@ thm trm_lts.distinct thm trm_lts.fv[simplified trm_lts.supp] +lemma lets_bla: + "x \ z \ y \ z \ x \ y \(Lt (Lcons x (Vr y) Lnil) (Vr x)) \ (Lt (Lcons x (Vr z) Lnil) (Vr x))" + by (simp add: trm_lts.eq_iff) + +lemma lets_ok: + "(Lt (Lcons x (Vr y) Lnil) (Vr x)) = (Lt (Lcons y (Vr y) Lnil) (Vr y))" + apply (simp add: trm_lts.eq_iff) + apply (rule_tac x="(x \ y)" in exI) + apply (simp_all add: alphas) + apply (simp add: fresh_star_def eqvts) + done + +lemma lets_ok3: + "x \ y \ + (Lt (Lcons x (Ap (Vr y) (Vr x)) (Lcons y (Vr y) Lnil)) (Ap (Vr x) (Vr y))) \ + (Lt (Lcons y (Ap (Vr x) (Vr y)) (Lcons x (Vr x) Lnil)) (Ap (Vr x) (Vr y)))" + apply (simp add: alphas trm_lts.eq_iff) + done + + +lemma lets_not_ok1: + "(Lt (Lcons x (Vr x) (Lcons y (Vr y) Lnil)) (Ap (Vr x) (Vr y))) = + (Lt (Lcons y (Vr x) (Lcons x (Vr y) Lnil)) (Ap (Vr x) (Vr y)))" + apply (simp add: alphas trm_lts.eq_iff) + apply (rule_tac x="0::perm" in exI) + apply (simp add: fresh_star_def eqvts) + apply blast + done + +lemma lets_nok: + "x \ y \ x \ z \ z \ y \ + (Lt (Lcons x (Ap (Vr z) (Vr z)) (Lcons y (Vr z) Lnil)) (Ap (Vr x) (Vr y))) \ + (Lt (Lcons y (Vr z) (Lcons x (Ap (Vr z) (Vr z)) Lnil)) (Ap (Vr x) (Vr y)))" + apply (simp add: alphas trm_lts.eq_iff fresh_star_def) + done + + end diff -r 5f0bb35114c3 -r a7e60da429e2 Nominal/Manual/Term5n.thy --- a/Nominal/Manual/Term5n.thy Tue Mar 23 09:06:28 2010 +0100 +++ b/Nominal/Manual/Term5n.thy Tue Mar 23 09:13:17 2010 +0100 @@ -174,55 +174,4 @@ lemmas fv_lts[simp] = fv_rlts.simps[quot_lifted] lemmas alpha5_INJ = alpha5_inj[unfolded alpha_gen, quot_lifted, folded alpha_gen] -lemma lets_bla: - "x \ z \ y \ z \ x \ y \(Lt5 (Lcons x (Vr5 y) Lnil) (Vr5 x)) \ (Lt5 (Lcons x (Vr5 z) Lnil) (Vr5 x))" -apply (simp only: alpha5_INJ) -apply (simp only: bv5) -apply simp -done - -lemma lets_ok: - "(Lt5 (Lcons x (Vr5 y) Lnil) (Vr5 x)) = (Lt5 (Lcons y (Vr5 y) Lnil) (Vr5 y))" -apply (simp add: alpha5_INJ) -apply (rule_tac x="(x \ y)" in exI) -apply (simp_all add: alpha_gen) -apply (simp add: permute_trm5_lts fresh_star_def eqvts) -done - -lemma lets_ok3: - "x \ y \ - (Lt5 (Lcons x (Ap5 (Vr5 y) (Vr5 x)) (Lcons y (Vr5 y) Lnil)) (Ap5 (Vr5 x) (Vr5 y))) \ - (Lt5 (Lcons y (Ap5 (Vr5 x) (Vr5 y)) (Lcons x (Vr5 x) Lnil)) (Ap5 (Vr5 x) (Vr5 y)))" -apply (simp add: permute_trm5_lts alpha_gen alpha5_INJ) -done - - -lemma lets_not_ok1: - "(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)))" -apply (simp add: alpha5_INJ alpha_gen) -apply (rule_tac x="0::perm" in exI) -apply (simp add: permute_trm5_lts fresh_star_def alpha5_INJ(5) alpha5_INJ(2) alpha5_INJ(1) eqvts) -apply blast -done - -lemma distinct_helper: - shows "\(rVr5 x \5 rAp5 y z)" - apply auto - apply (erule alpha_rtrm5.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 (simp only: alpha5_INJ(3) alpha5_INJ(5) alpha_gen permute_trm5_lts fresh_star_def) -apply (simp add: distinct_helper2 alpha5_INJ permute_trm5_lts) -done - end