diff -r fe84fcfab46f -r 687369ae8f81 Nominal/Ex/Lambda.thy --- a/Nominal/Ex/Lambda.thy Fri May 21 10:45:29 2010 +0200 +++ b/Nominal/Ex/Lambda.thy Fri May 21 10:47:07 2010 +0200 @@ -1,5 +1,5 @@ theory Lambda -imports "../NewParser" +imports "../NewParser" Quotient_Option begin atom_decl name @@ -477,6 +477,206 @@ definition new where "new s = (THE x. \a \ s. atom x \ a)" +primrec match_Var_raw where + "match_Var_raw (Var_raw x) = Some x" +| "match_Var_raw (App_raw x y) = None" +| "match_Var_raw (Lam_raw n t) = None" + +quotient_definition + "match_Var :: lam \ name option" +is match_Var_raw + +lemma [quot_respect]: "(alpha_lam_raw ===> op =) match_Var_raw match_Var_raw" + apply rule + apply (induct_tac a b rule: alpha_lam_raw.induct) + apply simp_all + done + +lemmas match_Var_simps = match_Var_raw.simps[quot_lifted] + +primrec match_App_raw where + "match_App_raw (Var_raw x) = None" +| "match_App_raw (App_raw x y) = Some (x, y)" +| "match_App_raw (Lam_raw n t) = None" + +quotient_definition + "match_App :: lam \ (lam \ lam) option" +is match_App_raw + +lemma [quot_respect]: + "(alpha_lam_raw ===> option_rel (prod_rel alpha_lam_raw alpha_lam_raw)) match_App_raw match_App_raw" + apply (intro fun_relI) + apply (induct_tac a b rule: alpha_lam_raw.induct) + apply simp_all + done + +lemmas match_App_simps = match_App_raw.simps[quot_lifted] + +primrec match_Lam_raw where + "match_Lam_raw (S :: atom set) (Var_raw x) = None" +| "match_Lam_raw S (App_raw x y) = None" +| "match_Lam_raw S (Lam_raw n t) = (let z = new (S \ (fv_lam_raw t - {atom n})) in Some (z, (n \ z) \ t))" + +quotient_definition + "match_Lam :: (atom set) \ lam \ (name \ lam) option" +is match_Lam_raw + +lemma [quot_respect]: + "(op = ===> alpha_lam_raw ===> option_rel (prod_rel op = alpha_lam_raw)) match_Lam_raw match_Lam_raw" + proof (intro fun_relI, clarify) + fix S t s + assume a: "alpha_lam_raw t s" + show "option_rel (prod_rel op = alpha_lam_raw) (match_Lam_raw S t) (match_Lam_raw S s)" + using a proof (induct t s rule: alpha_lam_raw.induct) + case goal1 show ?case by simp + next + case goal2 show ?case by simp + next + case (goal3 x t y s) + then obtain p where "({atom x}, t) \gen (\x1 x2. alpha_lam_raw x1 x2 \ + option_rel (prod_rel op = alpha_lam_raw) (match_Lam_raw S x1) + (match_Lam_raw S x2)) fv_lam_raw p ({atom y}, s)" .. + then have + c: "fv_lam_raw t - {atom x} = fv_lam_raw s - {atom y}" and + d: "(fv_lam_raw t - {atom x}) \* p" and + e: "alpha_lam_raw (p \ t) s" and + f: "option_rel (prod_rel op = alpha_lam_raw) (match_Lam_raw S (p \ t)) (match_Lam_raw S s)" and + g: "p \ {atom x} = {atom y}" unfolding alphas(1) by - (elim conjE, assumption)+ + let ?z = "new (S \ (fv_lam_raw t - {atom x}))" + have h: "?z = new (S \ (fv_lam_raw s - {atom y}))" using c by simp + show ?case + unfolding match_Lam_raw.simps Let_def option_rel.simps prod_rel.simps split_conv + proof + show "?z = new (S \ (fv_lam_raw s - {atom y}))" by (fact h) + next + have "atom y \ p" sorry + have "fv_lam_raw t \* ((x \ y) \ p)" sorry + then have "alpha_lam_raw (((x \ y) \ p) \ t) t" sorry + have "alpha_lam_raw (p \ t) ((x \ y) \ t)" sorry + have "alpha_lam_raw t ((x \ y) \ s)" sorry + then have "alpha_lam_raw ((x \ ?z) \ t) ((y \ ?z) \ s)" using eqvts(15) sorry + then show "alpha_lam_raw ((x \ new (S \ (fv_lam_raw t - {atom x}))) \ t) + ((y \ new (S \ (fv_lam_raw s - {atom y}))) \ s)" unfolding h . + qed + qed + qed + +lemmas match_Lam_simps = match_Lam_raw.simps[quot_lifted] + +lemma app_some: "match_App x = Some (a, b) \ x = App a b" +by (induct x rule: lam.induct) (simp_all add: match_App_simps) + +lemma lam_some: "match_Lam S x = Some (z, s) \ x = Lam z s \ atom z \ S" + apply (induct x rule: lam.induct) + apply (simp_all add: match_Lam_simps) + apply (simp add: Let_def) + apply (erule conjE) + apply (thin_tac "match_Lam S lam = Some (z, s) \ lam = Lam z s \ atom z \ S") + apply (rule conjI) + apply (simp add: lam.eq_iff) + apply (rule_tac x="(name \ z)" in exI) + apply (simp add: alphas) + apply (simp add: eqvts) + apply (simp only: lam.fv(3)[symmetric]) + apply (subgoal_tac "Lam name lam = Lam z s") + apply (simp del: lam.fv) + prefer 3 + apply (thin_tac "(name \ new (S \ (fv_lam lam - {atom name}))) \ lam = s") + apply (simp only: new_def) + apply (subgoal_tac "\a \ S. atom z \ a") + apply (simp only: fresh_def) + + thm new_def + apply simp + + +function subst where +"subst v s t = ( + case match_Var t of Some n \ if n = v then s else Var n | None \ + case match_App t of Some (l, r) \ App (subst v s l) (subst v s r) | None \ + case match_Lam (supp (v,s)) t of Some (n, t) \ Lam n (subst v s t) | None \ undefined)" +by pat_completeness auto + +termination apply (relation "measure (\(_, _, t). size t)") +apply auto[1] +defer +apply (case_tac a) apply simp +apply (frule app_some) apply simp +apply (case_tac a) apply simp +apply (frule app_some) apply simp +apply (case_tac a) apply simp +apply (frule lam_some) + apply simp +done + +lemmas lam_exhaust = lam_raw.exhaust[quot_lifted] + +lemma subst_eqvt: + "p \ (subst v s t) = subst (p \ v) (p \ s) (p \ t)" + proof (induct v s t rule: subst.induct) + case (1 v s t) + show ?case proof (cases t rule: lam_exhaust) + fix n + assume "t = Var n" + then show ?thesis by (simp add: match_Var_simps) + next + fix l r + assume "t = App l r" + then show ?thesis + apply (simp only:) + apply (subst subst.simps) + apply (subst match_Var_simps) + apply (simp only: option.cases) + apply (subst match_App_simps) + apply (simp only: option.cases) + apply (simp only: prod.cases) + apply (simp only: lam.perm) + apply (subst (3) subst.simps) + apply (subst match_Var_simps) + apply (simp only: option.cases) + apply (subst match_App_simps) + apply (simp only: option.cases) + apply (simp only: prod.cases) + apply (subst 1(2)[of "(l, r)" "l" "r"]) + apply (simp add: match_Var_simps) + apply (simp add: match_App_simps) + apply (rule refl) + apply (subst 1(3)[of "(l, r)" "l" "r"]) + apply (simp add: match_Var_simps) + apply (simp add: match_App_simps) + apply (rule refl) + apply (rule refl) + done + next + fix n t' + assume "t = Lam n t'" + then show ?thesis + apply (simp only: ) + apply (simp only: lam.perm) + apply (subst subst.simps) + apply (subst match_Var_simps) + apply (simp only: option.cases) + apply (subst match_App_simps) + apply (simp only: option.cases) + apply (subst match_Lam_simps) + apply (simp only: Let_def) + apply (simp only: option.cases) + apply (simp only: prod.cases) + apply (subst (2) subst.simps) + apply (subst match_Var_simps) + apply (simp only: option.cases) + apply (subst match_App_simps) + apply (simp only: option.cases) + apply (subst match_Lam_simps) + apply (simp only: Let_def) + apply (simp only: option.cases) + apply (simp only: prod.cases) + apply (simp only: lam.perm) + apply (simp only: lam.eq_iff) + sorry + qed + qed + lemma size_no_change: "size (p \ (t :: lam_raw)) = size t" by (induct t) simp_all @@ -488,7 +688,7 @@ | "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) + by (pat_completeness, auto) termination apply (relation "measure (\(t, y, s). (size t))") apply (auto simp add: size_no_change) @@ -523,9 +723,9 @@ oops quotient_definition - subst ("_ [ _ ::= _ ]" [100,100,100] 100) + subst2 ("_ [ _ ::= _ ]" [100,100,100] 100) where - "subst :: lam \ name \ lam \ lam" is "subst_raw" + "subst2 :: lam \ name \ lam \ lam" is "subst_raw" lemmas fv_rsp = quot_respect(10)[simplified] @@ -585,6 +785,9 @@ lemmas subst_simps = subst_raw.simps(1-2)[quot_lifted,no_vars] simp3[quot_lifted,simplified lam.supp,simplified fresh_def[symmetric], no_vars] + +thm subst_raw.simps(3)[quot_lifted,no_vars] + end