diff -r 6088fea1c8b1 -r 8a1c8dc72b5c Quot/Examples/LamEx.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Quot/Examples/LamEx.thy Mon Dec 07 14:09:50 2009 +0100 @@ -0,0 +1,252 @@ +theory LamEx +imports Nominal "../QuotMain" +begin + +atom_decl name + +thm abs_fresh(1) + +nominal_datatype rlam = + rVar "name" +| rApp "rlam" "rlam" +| rLam "name" "rlam" + +print_theorems + +function + rfv :: "rlam \ name set" +where + rfv_var: "rfv (rVar a) = {a}" +| rfv_app: "rfv (rApp t1 t2) = (rfv t1) \ (rfv t2)" +| rfv_lam: "rfv (rLam a t) = (rfv t) - {a}" +sorry + +termination rfv sorry + +inductive + alpha :: "rlam \ rlam \ bool" ("_ \ _" [100, 100] 100) +where + a1: "a = b \ (rVar a) \ (rVar b)" +| a2: "\t1 \ t2; s1 \ s2\ \ rApp t1 s1 \ rApp t2 s2" +| a3: "\t \ ([(a,b)]\s); a \ rfv (rLam b t)\ \ rLam a t \ rLam b s" + +print_theorems + +lemma alpha_refl: + fixes t::"rlam" + shows "t \ t" + apply(induct t rule: rlam.induct) + apply(simp add: a1) + apply(simp add: a2) + apply(rule a3) + apply(subst pt_swap_bij'') + apply(rule pt_name_inst) + apply(rule at_name_inst) + apply(simp) + apply(simp) + done + +lemma alpha_equivp: + shows "equivp alpha" +sorry + +quotient lam = rlam / alpha + apply(rule alpha_equivp) + done + +print_quotients + +quotient_def + Var :: "name \ lam" +where + "Var \ rVar" + +quotient_def + App :: "lam \ lam \ lam" +where + "App \ rApp" + +quotient_def + Lam :: "name \ lam \ lam" +where + "Lam \ rLam" + +thm Var_def +thm App_def +thm Lam_def + +quotient_def + fv :: "lam \ name set" +where + "fv \ rfv" + +thm fv_def + +(* definition of overloaded permutation function *) +(* for the lifted type lam *) +overloading + perm_lam \ "perm :: 'x prm \ lam \ lam" (unchecked) +begin + +quotient_def + perm_lam :: "'x prm \ lam \ lam" +where + "perm_lam \ (perm::'x prm \ rlam \ rlam)" + +end + +(*quotient_def (for lam) + abs_fun_lam :: "'x prm \ lam \ lam" +where + "perm_lam \ (perm::'x prm \ rlam \ rlam)"*) + + +thm perm_lam_def + +(* lemmas that need to lift *) +lemma pi_var_com: + fixes pi::"'x prm" + shows "(pi\rVar a) \ rVar (pi\a)" + sorry + +lemma pi_app_com: + fixes pi::"'x prm" + shows "(pi\rApp t1 t2) \ rApp (pi\t1) (pi\t2)" + sorry + +lemma pi_lam_com: + fixes pi::"'x prm" + shows "(pi\rLam a t) \ rLam (pi\a) (pi\t)" + sorry + + + +lemma real_alpha: + assumes a: "t = [(a,b)]\s" "a\[b].s" + shows "Lam a t = Lam b s" +using a +unfolding fresh_def supp_def +sorry + +lemma perm_rsp[quotient_rsp]: + "(op = ===> alpha ===> alpha) op \ op \" + apply(auto) + (* this is propably true if some type conditions are imposed ;o) *) + sorry + +lemma fresh_rsp: + "(op = ===> alpha ===> op =) fresh fresh" + apply(auto) + (* this is probably only true if some type conditions are imposed *) + sorry + +lemma rVar_rsp[quotient_rsp]: + "(op = ===> alpha) rVar rVar" +by (auto intro:a1) + +lemma rApp_rsp[quotient_rsp]: "(alpha ===> alpha ===> alpha) rApp rApp" +by (auto intro:a2) + +lemma rLam_rsp[quotient_rsp]: "(op = ===> alpha ===> alpha) rLam rLam" + apply(auto) + apply(rule a3) + apply(rule_tac t="[(x,x)]\y" and s="y" in subst) + apply(rule sym) + apply(rule trans) + apply(rule pt_name3) + apply(rule at_ds1[OF at_name_inst]) + apply(simp add: pt_name1) + apply(assumption) + apply(simp add: abs_fresh) + done + +lemma rfv_rsp[quotient_rsp]: "(alpha ===> op =) rfv rfv" + sorry + +lemma rvar_inject: "rVar a \ rVar b = (a = b)" +apply (auto) +apply (erule alpha.cases) +apply (simp_all add: rlam.inject alpha_refl) +done + +ML {* val qty = @{typ "lam"} *} +ML {* val rsp_thms = @{thms perm_rsp fresh_rsp rVar_rsp rApp_rsp rLam_rsp rfv_rsp} *} + +ML {* val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} qty *} +ML {* val (trans2, reps_same, absrep, quot) = lookup_quot_thms @{context} "lam" *} +ML {* fun lift_tac_lam lthy t = lift_tac lthy t *} + +lemma pi_var: "(pi\('x \ 'x) list) \ Var a = Var (pi \ a)" +apply (tactic {* lift_tac_lam @{context} @{thm pi_var_com} 1 *}) +done + +lemma pi_app: "(pi\('x \ 'x) list) \ App (x\lam) (xa\lam) = App (pi \ x) (pi \ xa)" +apply (tactic {* lift_tac_lam @{context} @{thm pi_app_com} 1 *}) +done + +lemma pi_lam: "(pi\('x \ 'x) list) \ Lam (a\name) (x\lam) = Lam (pi \ a) (pi \ x)" +apply (tactic {* lift_tac_lam @{context} @{thm pi_lam_com} 1 *}) +done + +lemma fv_var: "fv (Var (a\name)) = {a}" +apply (tactic {* lift_tac_lam @{context} @{thm rfv_var} 1 *}) +done + +lemma fv_app: "fv (App (x\lam) (xa\lam)) = fv x \ fv xa" +apply (tactic {* lift_tac_lam @{context} @{thm rfv_app} 1 *}) +done + +lemma fv_lam: "fv (Lam (a\name) (x\lam)) = fv x - {a}" +apply (tactic {* lift_tac_lam @{context} @{thm rfv_lam} 1 *}) +done + +lemma a1: "(a\name) = (b\name) \ Var a = Var b" +apply (tactic {* lift_tac_lam @{context} @{thm a1} 1 *}) +done + +lemma a2: "\(x\lam) = (xa\lam); (xb\lam) = (xc\lam)\ \ App x xb = App xa xc" +apply (tactic {* lift_tac_lam @{context} @{thm a2} 1 *}) +done + +lemma a3: "\(x\lam) = [(a\name, b\name)] \ (xa\lam); a \ fv (Lam b x)\ \ Lam a x = Lam b xa" +apply (tactic {* lift_tac_lam @{context} @{thm a3} 1 *}) +done + +lemma alpha_cases: "\a1 = a2; \a b. \a1 = Var a; a2 = Var b; a = b\ \ P; + \x xa xb xc. \a1 = App x xb; a2 = App xa xc; x = xa; xb = xc\ \ P; + \x a b xa. \a1 = Lam a x; a2 = Lam b xa; x = [(a, b)] \ xa; a \ fv (Lam b x)\ \ P\ + \ P" +apply (tactic {* lift_tac_lam @{context} @{thm alpha.cases} 1 *}) +done + +lemma alpha_induct: "\(qx\lam) = (qxa\lam); \(a\name) b\name. a = b \ (qxb\lam \ lam \ bool) (Var a) (Var b); + \(x\lam) (xa\lam) (xb\lam) xc\lam. \x = xa; qxb x xa; xb = xc; qxb xb xc\ \ qxb (App x xb) (App xa xc); + \(x\lam) (a\name) (b\name) xa\lam. + \x = [(a, b)] \ xa; qxb x ([(a, b)] \ xa); a \ fv (Lam b x)\ \ qxb (Lam a x) (Lam b xa)\ + \ qxb qx qxa" +apply (tactic {* lift_tac_lam @{context} @{thm alpha.induct} 1 *}) +done + +lemma var_inject: "(Var a = Var b) = (a = b)" +apply (tactic {* lift_tac_lam @{context} @{thm rvar_inject} 1 *}) +done + +lemma lam_induct:" \\name. P (Var name); \lam1 lam2. \P lam1; P lam2\ \ P (App lam1 lam2); + \name lam. P lam \ P (Lam name lam)\ \ P lam" +apply (tactic {* lift_tac_lam @{context} @{thm rlam.induct} 1 *}) +done + +lemma var_supp: + shows "supp (Var a) = ((supp a)::name set)" + apply(simp add: supp_def) + apply(simp add: pi_var) + apply(simp add: var_inject) + done + +lemma var_fresh: + fixes a::"name" + shows "(a\(Var b)) = (a\b)" + apply(simp add: fresh_def) + apply(simp add: var_supp) + done +