# HG changeset patch # User Christian Urban # Date 1274269443 -3600 # Node ID a5dc3558cdecbbc8312309e5c7b2a34f3b27b2f8 # Parent 5dc48e1af733d87191a50488142c66b8f59ae487# Parent d76667e51d307bc41dca2d176b763aaeee85e5fd merged diff -r 5dc48e1af733 -r a5dc3558cdec Nominal/Ex/Lambda.thy --- a/Nominal/Ex/Lambda.thy Wed May 19 12:43:38 2010 +0100 +++ b/Nominal/Ex/Lambda.thy Wed May 19 12:44:03 2010 +0100 @@ -472,6 +472,115 @@ nominal_inductive typing *) +(* Substitution *) + +definition new where + "new s = (THE x. \a \ s. atom x \ a)" + +lemma size_no_change: "size (p \ (t :: lam_raw)) = size t" + by (induct t) simp_all + +function + subst_raw :: "lam_raw \ name \ lam_raw \ lam_raw" +where + "subst_raw (Var_raw x) y s = (if x=y then s else (Var_raw x))" +| "subst_raw (App_raw l r) y s = App_raw (subst_raw l y s) (subst_raw r y s)" +| "subst_raw (Lam_raw x t) y s = + Lam_raw (new ({atom y} \ fv_lam_raw s \ fv_lam_raw t - {atom x})) + (subst_raw ((x \ (new ({atom y} \ fv_lam_raw s \ fv_lam_raw t - {atom x}))) \ t) y s)" +by (pat_completeness, auto) +termination + apply (relation "measure (\(t, y, s). (size t))") + apply (auto simp add: size_no_change) + done + +lemma fv_subst[simp]: "fv_lam_raw (subst_raw t y s) = + (if (atom y \ fv_lam_raw t) then fv_lam_raw s \ (fv_lam_raw t - {atom y}) else fv_lam_raw t)" + apply (induct t arbitrary: s) + apply (auto simp add: supp_at_base)[1] + apply (auto simp add: supp_at_base)[1] + apply (simp only: fv_lam_raw.simps) + apply simp + apply (rule conjI) + apply clarify + sorry + +thm supp_at_base +lemma new_eqvt[eqvt]: "p \ (new s) = new (p \ s)" + sorry + +lemma subst_var_raw_eqvt[eqvt]: "p \ (subst_raw t y s) = subst_raw (p \ t) (p \ y) (p \ s)" + apply (induct t arbitrary: p y s) + apply simp_all + apply(perm_simp) + apply simp + sorry + +lemma subst_id: "alpha_lam_raw (subst_raw x d (Var_raw d)) x" + apply (induct x arbitrary: d) + apply (simp_all add: alpha_lam_raw.intros) + apply (rule alpha_lam_raw.intros) + apply (rule_tac x="(name \ new (insert (atom d) (supp d)))" in exI) + apply (simp add: alphas) + oops + +quotient_definition + subst ("_ [ _ ::= _ ]" [100,100,100] 100) +where + "subst :: lam \ name \ lam \ lam" is "subst_raw" + +lemmas fv_rsp = quot_respect(10)[simplified] + +lemma subst_rsp_pre1: + assumes a: "alpha_lam_raw a b" + shows "alpha_lam_raw (subst_raw a y c) (subst_raw b y c)" + using a + apply (induct a b arbitrary: c y rule: alpha_lam_raw.induct) + apply (simp add: equivp_reflp[OF lam_equivp]) + apply (simp add: alpha_lam_raw.intros) + apply (simp only: alphas) + apply clarify + apply (simp only: subst_raw.simps) + apply (rule alpha_lam_raw.intros) + apply (simp only: alphas) + sorry + +lemma subst_rsp_pre2: + assumes a: "alpha_lam_raw a b" + shows "alpha_lam_raw (subst_raw c y a) (subst_raw c y b)" + using a + apply (induct c arbitrary: a b y) + apply (simp add: equivp_reflp[OF lam_equivp]) + apply (simp add: alpha_lam_raw.intros) + apply simp + apply (rule alpha_lam_raw.intros) + sorry + +lemma [quot_respect]: + "(alpha_lam_raw ===> op = ===> alpha_lam_raw ===> alpha_lam_raw) subst_raw subst_raw" + proof (intro fun_relI, simp) + fix a b c d :: lam_raw + fix y :: name + assume a: "alpha_lam_raw a b" + assume b: "alpha_lam_raw c d" + have c: "alpha_lam_raw (subst_raw a y c) (subst_raw b y c)" using subst_rsp_pre1 a by simp + then have d: "alpha_lam_raw (subst_raw b y c) (subst_raw b y d)" using subst_rsp_pre2 b by simp + show "alpha_lam_raw (subst_raw a y c) (subst_raw b y d)" + using c d equivp_transp[OF lam_equivp] by blast + qed + +lemma simp3: + "x \ y \ atom x \ fv_lam_raw s \ alpha_lam_raw (subst_raw (Lam_raw x t) y s) (Lam_raw x (subst_raw t y s))" + apply simp + apply (rule alpha_lam_raw.intros) + apply (rule_tac x ="(x \ (new (insert (atom y) (fv_lam_raw s \ fv_lam_raw t) - + {atom x})))" in exI) + apply (simp only: alphas) + apply simp + sorry + +lemmas subst_simps = subst_raw.simps(1-2)[quot_lifted,no_vars] + simp3[quot_lifted,simplified lam.supp,simplified fresh_def[symmetric], no_vars] end