# HG changeset patch # User urbanc # Date 1158672685 -7200 # Node ID fd9b0b78a7d3810819d8946cdb8ebacaa1981f51 # Parent 56e4bb01fd9976baea879a579c98d6d257888446 this file contains a compile-challenge suggested by Adam Chlipala; so far it contains only the definition and no proofs diff -r 56e4bb01fd99 -r fd9b0b78a7d3 src/HOL/Nominal/Examples/Compile.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Nominal/Examples/Compile.thy Tue Sep 19 15:31:25 2006 +0200 @@ -0,0 +1,759 @@ +(* $Id: *) + +(* A challenge suggested by Adam Chlipala *) + +theory Compile +imports "Nominal" +begin + +atom_decl name + +nominal_datatype data = DNat + | DProd "data" "data" + | DSum "data" "data" + +nominal_datatype ty = Data "data" + | Arrow "ty" "ty" ("_\_" [100,100] 100) + +nominal_datatype trm = Var "name" + | Lam "\name\trm" ("Lam [_]._" [100,100] 100) + | App "trm" "trm" + | Const "nat" + | Pr "trm" "trm" + | Fst "trm" + | Snd "trm" + | InL "trm" + | InR "trm" + | Case "trm" "\name\trm" "\name\trm" + ("Case _ of inl _ \ _ | inr _ \ _" [100,100,100,100,100] 100) + +nominal_datatype dataI = OneI | NatI + +nominal_datatype tyI = DataI "dataI" + | ArrowI "tyI" "tyI" ("_\_" [100,100] 100) + +nominal_datatype trmI = IVar "name" + | ILam "\name\trmI" ("ILam [_]._" [100,100] 100) + | IApp "trmI" "trmI" + | IUnit + | INat "nat" + | ISucc "trmI" + | IAss "trmI" "trmI" ("_\_" [100,100] 100) + | IRef "trmI" + | ISeq "trmI" "trmI" ("_;;_" [100,100] 100) + | Iif "trmI" "trmI" "trmI" + +text {* valid contexts *} + +consts + ctxts :: "((name\'a::pt_name) list) set" + valid :: "(name\'a::pt_name) list \ bool" +translations + "valid \" \ "\ \ ctxts" + +inductive ctxts +intros +v1[intro]: "valid []" +v2[intro]: "\valid \;a\\\\ valid ((a,\)#\)" + +text {* typing judgements for trms *} + +consts + typing :: "(((name\ty) list)\trm\ty) set" +syntax + "_typing_judge" :: "(name\ty) list\trm\ty\bool" (" _ \ _ : _ " [80,80,80] 80) +translations + "\ \ t : \" \ "(\,t,\) \ typing" + +inductive typing +intros +t0[intro]: "\valid \; (x,\)\set \\\ \ \ Var x : \" +t1[intro]: "\\ \ e1 : \1\\2; \ \ e2 : \1\\ \ \ App e1 e2 : \2" +t2[intro]: "\x\\;((x,\1)#\) \ t : \2\ \ \ \ Lam [x].t : \1\\2" +t3[intro]: "valid \ \ \ \ Const n : Data(DNat)" +t4[intro]: "\\ \ e1 : Data(\1); \ \ e2 : Data(\2)\ \ \ \ Pr e1 e2 : Data (DProd \1 \2)" +t5[intro]: "\\ \ e : Data(DProd \1 \2)\ \ \ \ Fst e : Data(\1)" +t6[intro]: "\\ \ e : Data(DProd \1 \2)\ \ \ \ Snd e : Data(\2)" +t7[intro]: "\\ \ e : Data(\1)\ \ \ \ InL e : Data(DSum \1 \2)" +t8[intro]: "\\ \ e : Data(\2)\ \ \ \ InR e : Data(DSum \1 \2)" +t9[intro]: "\x1\\; x2\\; \ \ e: Data(DSum \1 \2); + ((x1,Data(\1))#\) \ e1 : \; ((x2,Data(\2))#\) \ e2 : \\ + \ \ \ (Case e of inl x1 \ e1 | inr x2 \ e2) : \" + +text {* typing judgements for Itrms *} + +consts + Ityping :: "(((name\tyI) list)\trmI\tyI) set" +syntax + "_typing_judge" :: "(name\tyI) list\trmI\tyI\bool" (" _ I\ _ : _ " [80,80,80] 80) +translations + "\ I\ t : \" \ "(\,t,\) \ Ityping" + +inductive Ityping +intros +t0[intro]: "\valid \; (x,\)\set \\\ \ I\ IVar x : \" +t1[intro]: "\\ I\ e1 : \1\\2; \ I\ e2 : \1\\ \ I\ IApp e1 e2 : \2" +t2[intro]: "\x\\;((x,\1)#\) I\ t : \2\ \ \ I\ ILam [x].t : \1\\2" +t3[intro]: "valid \ \ \ I\ IUnit : DataI(OneI)" +t4[intro]: "valid \ \ \ I\ INat(n) : DataI(NatI)" +t5[intro]: "\ I\ e : DataI(NatI) \ \ I\ ISucc(e) : DataI(NatI)" +t6[intro]: "\\ I\ e : DataI(NatI)\ \ \ I\ IRef e : DataI (NatI)" +t7[intro]: "\\ I\ e1 : DataI(NatI); \ I\ e2 : DataI(NatI)\ \ \ I\ e1\e2 : DataI(OneI)" +t8[intro]: "\\ I\ e1 : DataI(NatI); \ I\ e2 : \\ \ \ I\ e1;;e2 : \" +t9[intro]: "\\ I\ e: DataI(NatI); \ I\ e1 : \; \ I\ e2 : \\ \ \ I\ Iif e e1 e2 : \" + +text {* capture-avoiding substitution *} + +consts + subst :: "'a \ name \ 'a \ 'a" ("_[_::=_]" [100,100,100] 100) + +constdefs + subst_Var :: "name \ trm \ name \ trm" + "subst_Var x t' \ \y. (if y=x then t' else (Var y))" + + subst_App :: "name \ trm \ trm \ trm \ trm \ trm \ trm" + "subst_App x t' \ \_ _ r1 r2. App r1 r2" + + subst_Lam :: "name \ trm \ name \ trm \ trm \ trm" + "subst_Lam x t' \ \a _ r. Lam [a].r" + + subst_Const :: "name \ trm \ nat \ trm" + "subst_Const x t' \ \n. Const n" + + subst_Pr :: "name \ trm \ trm \ trm \ trm \ trm \ trm" + "subst_Pr x t' \ \_ _ r1 r2. Pr r1 r2" + + subst_Fst :: "name \ trm \ trm \ trm \ trm" + "subst_Fst x t' \ \_ r. Fst r" + + subst_Snd :: "name \ trm \ trm \ trm \ trm" + "subst_Snd x t' \ \_ r. Snd r" + + subst_InL :: "name \ trm \ trm \ trm \ trm" + "subst_InL x t' \ \_ r. InL r" + + subst_InR :: "name \ trm \ trm \ trm \ trm" + "subst_InR x t' \ \_ r. InR r" + + subst_Case :: "name \ trm \ trm \ name \ trm \ name \ trm \ trm \ trm \ trm \ trm" + "subst_Case x t' \ \_ x _ y _ r r1 r2. Case r of inl x \ r1 | inr y \ r2" + +defs(overloaded) + subst_def: + "t[x::=t'] \ trm_rec (subst_Var x t') (subst_Lam x t') (subst_App x t') + (subst_Const x t') (subst_Pr x t') (subst_Fst x t') (subst_Snd x t') + (subst_InL x t') (subst_InR x t') (subst_Case x t') t" + +(* FIXME: the next two lemmas need to be in Nominal.thy *) +lemma perm_if: + fixes pi::"'a prm" + shows "pi\(if b then c1 else c2) = (if (pi\b) then (pi\c1) else (pi\c2))" +apply(simp add: perm_fun_def) +done + +lemma eq_eqvt: + fixes pi::"name prm" + and x::"'a::pt_name" + shows "pi\(x=y) = ((pi\x)=(pi\y))" + apply(simp add: perm_bool perm_bij) + done + +lemma fin_supp_subst: + shows "finite ((supp (subst_Var x t))::name set)" + and "finite ((supp (subst_Lam x t))::name set)" + and "finite ((supp (subst_App x t))::name set)" + and "finite ((supp (subst_Const x t))::name set)" + and "finite ((supp (subst_Pr x t))::name set)" + and "finite ((supp (subst_Fst x t))::name set)" + and "finite ((supp (subst_Snd x t))::name set)" + and "finite ((supp (subst_InL x t))::name set)" + and "finite ((supp (subst_InR x t))::name set)" + and "finite ((supp (subst_Case x t))::name set)" +apply - +apply(finite_guess add: fs_name1 subst_Var_def perm_if eq_eqvt) +apply(finite_guess add: fs_name1 subst_Lam_def) +apply(finite_guess add: fs_name1 subst_App_def) +apply(finite_guess add: fs_name1 subst_Const_def) +apply(finite_guess add: fs_name1 subst_Pr_def) +apply(finite_guess add: fs_name1 subst_Fst_def) +apply(finite_guess add: fs_name1 subst_Snd_def) +apply(finite_guess add: fs_name1 subst_InL_def) +apply(finite_guess add: fs_name1 subst_InR_def) +apply(finite_guess only: fs_name1 subst_Case_def) +apply(perm_simp) +apply(simp) +done + +lemma fcb_subst_Lam: + shows "x\(subst_Lam y t') x t r" + by (simp add: subst_Lam_def abs_fresh) + +lemma fcb_subst_Case: + assumes a: "x\r" "x\r2" "y\r" "y\r1" + shows "x\(subst_Case z t') e x e1 y e2 r r1 r2 \ y\(subst_Case z t') e x e1 y e2 r r1 r2" + using a + by (simp add: subst_Case_def abs_fresh) + +lemma subst: + shows "(Var x)[y::=t'] = (if x=y then t' else (Var x))" + and "(App t1 t2)[y::=t'] = App (t1[y::=t']) (t2[y::=t'])" + and "\x\y; x\t'\ \ (Lam [x].t)[y::=t'] = Lam [x].(t[y::=t'])" + and "(Const n)[y::=t'] = Const n" + and "(Pr e1 e2)[y::=t'] = Pr (e1[y::=t']) (e2[y::=t'])" + and "(Fst e)[y::=t'] = Fst (e[y::=t'])" + and "(Snd e)[y::=t'] = Snd (e[y::=t'])" + and "(InL e)[y::=t'] = InL (e[y::=t'])" + and "(InR e)[y::=t'] = InR (e[y::=t'])" + and "\z\x; x\y; x\e; x\e2; z\y; z\e; z\e1; x\t'; z\t'\ + \ (Case e of inl x \ e1 | inr z \ e2)[y::=t'] = + (Case (e[y::=t']) of inl x \ (e1[y::=t']) | inr z \ (e2[y::=t']))" +apply(unfold subst_def) +apply(rule trans) +apply(rule trm.recs[where P="\_. True", simplified]) +apply(rule fin_supp_subst)+ +apply(simp add: fcb_subst_Lam) +apply(simp add: fcb_subst_Case) +apply(simp add: subst_Var_def) +apply(rule trans) +apply(rule trm.recs[where P="\_. True", simplified]) +apply(rule fin_supp_subst)+ +apply(simp add: fcb_subst_Lam) +apply(simp add: fcb_subst_Case) +apply(simp add: subst_App_def) +apply(rule trans) +apply(rule trm.recs[where P="\_. True", simplified]) +apply(rule fin_supp_subst)+ +apply(simp add: fcb_subst_Lam) +apply(simp add: fcb_subst_Case) +apply(fresh_guess add: fs_name1 subst_Var_def perm_if eq_eqvt) +apply(fresh_guess add: fs_name1 subst_Lam_def) +apply(fresh_guess add: fs_name1 subst_App_def) +apply(fresh_guess add: fs_name1 subst_Const_def) +apply(fresh_guess add: fs_name1 subst_Pr_def) +apply(fresh_guess add: fs_name1 subst_Fst_def) +apply(fresh_guess add: fs_name1 subst_Snd_def) +apply(fresh_guess add: fs_name1 subst_InL_def) +apply(fresh_guess add: fs_name1 subst_InR_def) +apply(fresh_guess only: fs_name1 subst_Case_def) +apply(perm_simp) +apply(simp, simp) +apply(simp add: subst_Lam_def) +apply(rule trans) +apply(rule trm.recs[where P="\_. True", simplified]) +apply(rule fin_supp_subst)+ +apply(simp add: fcb_subst_Lam) +apply(simp add: fcb_subst_Case) +apply(simp add: subst_Const_def) +apply(rule trans) +apply(rule trm.recs[where P="\_. True", simplified]) +apply(rule fin_supp_subst)+ +apply(simp add: fcb_subst_Lam) +apply(simp add: fcb_subst_Case) +apply(simp add: subst_Pr_def) +apply(rule trans) +apply(rule trm.recs[where P="\_. True", simplified]) +apply(rule fin_supp_subst)+ +apply(simp add: fcb_subst_Lam) +apply(simp add: fcb_subst_Case) +apply(simp add: subst_Fst_def) +apply(rule trans) +apply(rule trm.recs[where P="\_. True", simplified]) +apply(rule fin_supp_subst)+ +apply(simp add: fcb_subst_Lam) +apply(simp add: fcb_subst_Case) +apply(simp add: subst_Snd_def) +apply(rule trans) +apply(rule trm.recs[where P="\_. True", simplified]) +apply(rule fin_supp_subst)+ +apply(simp add: fcb_subst_Lam) +apply(simp add: fcb_subst_Case) +apply(simp add: subst_InL_def) +apply(rule trans) +apply(rule trm.recs[where P="\_. True", simplified]) +apply(rule fin_supp_subst)+ +apply(simp add: fcb_subst_Lam) +apply(simp add: fcb_subst_Case) +apply(simp add: subst_InR_def) +apply(rule trans) +apply(rule trm.recs[where P="\_. True", simplified]) +apply(rule fin_supp_subst)+ +apply(simp add: fcb_subst_Lam) +apply(simp add: fcb_subst_Case) +apply(fresh_guess add: fs_name1 subst_Var_def perm_if eq_eqvt) +apply(fresh_guess add: fs_name1 subst_Var_def perm_if eq_eqvt) +apply(fresh_guess add: fs_name1 subst_Lam_def) +apply(fresh_guess add: fs_name1 subst_Lam_def) +apply(fresh_guess add: fs_name1 subst_App_def) +apply(fresh_guess add: fs_name1 subst_App_def) +apply(fresh_guess add: fs_name1 subst_Const_def) +apply(fresh_guess add: fs_name1 subst_Const_def) +apply(fresh_guess add: fs_name1 subst_Pr_def) +apply(fresh_guess add: fs_name1 subst_Pr_def) +apply(fresh_guess add: fs_name1 subst_Fst_def) +apply(fresh_guess add: fs_name1 subst_Fst_def) +apply(fresh_guess add: fs_name1 subst_Snd_def) +apply(fresh_guess add: fs_name1 subst_Snd_def) +apply(fresh_guess add: fs_name1 subst_InL_def) +apply(fresh_guess add: fs_name1 subst_InL_def) +apply(fresh_guess add: fs_name1 subst_InR_def) +apply(fresh_guess add: fs_name1 subst_InR_def) +apply(fresh_guess only: fs_name1 subst_Case_def) +apply(perm_simp) +apply(simp, simp) +apply(fresh_guess only: fs_name1 subst_Case_def) +apply(perm_simp) +apply(simp, simp) +apply(assumption)+ +apply(simp add: subst_Case_def) +done + +constdefs + subst_IVar :: "name \ trmI \ name \ trmI" + "subst_IVar x t' \ \y. (if y=x then t' else (IVar y))" + + subst_IApp :: "name \ trmI \ trmI \ trmI \ trmI \ trmI \ trmI" + "subst_IApp x t' \ \_ _ r1 r2. IApp r1 r2" + + subst_ILam :: "name \ trmI \ name \ trmI \ trmI \ trmI" + "subst_ILam x t' \ \a _ r. ILam [a].r" + + subst_IUnit :: "name \ trmI \ trmI" + "subst_IUnit x t' \ IUnit" + + subst_INat :: "name \ trmI \ nat \ trmI" + "subst_INat x t' \ \n. INat n" + + subst_ISucc :: "name \ trmI \ trmI \ trmI \ trmI" + "subst_ISucc x t' \ \_ r. ISucc r" + + subst_IAss :: "name \ trmI \ trmI \ trmI \ trmI \ trmI \ trmI" + "subst_IAss x t' \ \_ _ r1 r2. r1\r2" + + subst_IRef :: "name \ trmI \ trmI \ trmI \ trmI" + "subst_IRef x t' \ \_ r. IRef r" + + subst_ISeq :: "name \ trmI \ trmI \ trmI \ trmI \ trmI \ trmI" + "subst_ISeq x t' \ \_ _ r1 r2. ISeq r1 r2" + + subst_Iif :: "name \ trmI \ trmI \ trmI \ trmI \ trmI \ trmI \ trmI \ trmI" + "subst_Iif x t' \ \_ _ _ r r1 r2. Iif r r1 r2" + +defs(overloaded) + subst_trmI_def: + "t[x::=t'] \ trmI_rec (subst_IVar x t') (subst_ILam x t') (subst_IApp x t') + (subst_IUnit x t') (subst_INat x t') (subst_ISucc x t') (subst_IAss x t') + (subst_IRef x t') (subst_ISeq x t') (subst_Iif x t') t" + +lemma fin_supp_Isubst: + shows "finite ((supp (subst_IVar x t))::name set)" + and "finite ((supp (subst_ILam x t))::name set)" + and "finite ((supp (subst_IApp x t))::name set)" + and "finite ((supp (subst_INat x t))::name set)" + and "finite ((supp (subst_IUnit x t))::name set)" + and "finite ((supp (subst_ISucc x t))::name set)" + and "finite ((supp (subst_IAss x t))::name set)" + and "finite ((supp (subst_IRef x t))::name set)" + and "finite ((supp (subst_ISeq x t))::name set)" + and "finite ((supp (subst_Iif x t))::name set)" +apply - +apply(finite_guess add: fs_name1 subst_IVar_def perm_if eq_eqvt) +apply(finite_guess add: fs_name1 subst_ILam_def) +apply(finite_guess add: fs_name1 subst_IApp_def) +apply(finite_guess add: fs_name1 subst_INat_def) +apply(finite_guess add: fs_name1 subst_IUnit_def) +apply(finite_guess add: fs_name1 subst_ISucc_def) +apply(finite_guess add: fs_name1 subst_IAss_def) +apply(finite_guess add: fs_name1 subst_IRef_def) +apply(finite_guess add: fs_name1 subst_ISeq_def) +apply(finite_guess only: fs_name1 subst_Iif_def) +apply(perm_simp) +apply(simp) +done + +lemma fcb_subst_ILam: + shows "x\(subst_ILam y t') x t r" + by (simp add: subst_ILam_def abs_fresh) + +lemma Isubst: + shows "(IVar x)[y::=t'] = (if x=y then t' else (IVar x))" + and "(IApp t1 t2)[y::=t'] = IApp (t1[y::=t']) (t2[y::=t'])" + and "\x\y; x\t'\ \ (ILam [x].t)[y::=t'] = ILam [x].(t[y::=t'])" + and "(INat n)[y::=t'] = INat n" + and "(IUnit)[y::=t'] = IUnit" + and "(ISucc e)[y::=t'] = ISucc (e[y::=t'])" + and "(IAss e1 e2)[y::=t'] = IAss (e1[y::=t']) (e2[y::=t'])" + and "(IRef e)[y::=t'] = IRef (e[y::=t'])" + and "(ISeq e1 e2)[y::=t'] = ISeq (e1[y::=t']) (e2[y::=t'])" + and "(Iif e e1 e2)[y::=t'] = Iif (e[y::=t']) (e1[y::=t']) (e2[y::=t'])" +apply(unfold subst_trmI_def) +apply(rule trans) +apply(rule trmI.recs[where P="\_. True", simplified]) +apply(rule fin_supp_Isubst)+ +apply(simp add: fcb_subst_ILam) +apply(simp add: subst_IVar_def) +apply(rule trans) +apply(rule trmI.recs[where P="\_. True", simplified]) +apply(rule fin_supp_Isubst)+ +apply(simp add: fcb_subst_ILam) +apply(simp add: subst_IApp_def) +apply(rule trans) +apply(rule trmI.recs[where P="\_. True", simplified]) +apply(rule fin_supp_Isubst)+ +apply(simp add: fcb_subst_ILam) +apply(fresh_guess add: fs_name1 subst_IVar_def perm_if eq_eqvt) +apply(fresh_guess add: fs_name1 subst_ILam_def) +apply(fresh_guess add: fs_name1 subst_IApp_def) +apply(fresh_guess add: fs_name1 subst_IUnit_def) +apply(fresh_guess add: fs_name1 subst_INat_def) +apply(fresh_guess add: fs_name1 subst_ISucc_def) +apply(fresh_guess add: fs_name1 subst_IAss_def) +apply(fresh_guess add: fs_name1 subst_IRef_def) +apply(fresh_guess add: fs_name1 subst_ISeq_def) +apply(fresh_guess only: fs_name1 subst_Iif_def) +apply(perm_simp) +apply(simp, simp) +apply(simp add: subst_ILam_def) +apply(rule trans) +apply(rule trmI.recs[where P="\_. True", simplified]) +apply(rule fin_supp_Isubst)+ +apply(simp add: fcb_subst_ILam) +apply(simp add: subst_INat_def) +apply(rule trans) +apply(rule trmI.recs[where P="\_. True", simplified]) +apply(rule fin_supp_Isubst)+ +apply(simp add: fcb_subst_ILam) +apply(simp add: subst_IUnit_def) +apply(rule trans) +apply(rule trmI.recs[where P="\_. True", simplified]) +apply(rule fin_supp_Isubst)+ +apply(simp add: fcb_subst_ILam) +apply(simp add: subst_ISucc_def) +apply(rule trans) +apply(rule trmI.recs[where P="\_. True", simplified]) +apply(rule fin_supp_Isubst)+ +apply(simp add: fcb_subst_ILam) +apply(simp add: subst_IAss_def) +apply(rule trans) +apply(rule trmI.recs[where P="\_. True", simplified]) +apply(rule fin_supp_Isubst)+ +apply(simp add: fcb_subst_ILam) +apply(simp add: subst_IRef_def) +apply(rule trans) +apply(rule trmI.recs[where P="\_. True", simplified]) +apply(rule fin_supp_Isubst)+ +apply(simp add: fcb_subst_ILam) +apply(simp add: subst_ISeq_def) +apply(rule trans) +apply(rule trmI.recs[where P="\_. True", simplified]) +apply(rule fin_supp_Isubst)+ +apply(simp add: fcb_subst_ILam) +apply(simp add: subst_Iif_def) +done + +lemma Isubst_eqvt: + fixes pi::"name prm" + and t1::"trmI" + and t2::"trmI" + and x::"name" + shows "pi\(t1[x::=t2]) = ((pi\t1)[(pi\x)::=(pi\t2)])" + apply(nominal_induct t1 avoiding: x t2 rule: trmI.induct) + apply(simp_all add: Isubst perm_if eq_eqvt fresh_bij) + done + +lemma Isubst_supp: + fixes t1::"trmI" + and t2::"trmI" + and x::"name" + shows "((supp (t1[x::=t2]))::name set) \ (supp t2)\((supp t1)-{x})" +apply(nominal_induct t1 avoiding: x t2 rule: trmI.induct) +apply(auto simp add: Isubst trmI.supp supp_atm abs_supp supp_nat) +apply(blast)+ +done + +lemma Isubst_fresh: + fixes x::"name" + and y::"name" + and t1::"trmI" + and t2::"trmI" + assumes a: "x\[y].t1" "x\t2" + shows "x\(t1[y::=t2])" +using a +apply(auto simp add: fresh_def Isubst_supp) +apply(drule rev_subsetD) +apply(rule Isubst_supp) +apply(simp add: abs_supp) +done + +text {* big-step evaluation for trms *} + +consts + big :: "(trm\trm) set" +syntax + "_big_judge" :: "trm\trm\bool" ("_ \ _" [80,80] 80) +translations + "e1 \ e2" \ "(e1,e2) \ big" + +inductive big +intros +b0[intro]: "Lam [x].e \ Lam [x].e" +b1[intro]: "\e1\Lam [x].e; e2\e2'; e[x::=e2']\e'\ \ App e1 e2 \ e'" +b2[intro]: "Const n \ Const n" +b3[intro]: "\e1\e1'; e2\e2'\ \ Pr e1 e2 \ Pr e1' e2'" +b4[intro]: "e\Pr e1 e2 \ Fst e\e1" +b5[intro]: "e\Pr e1 e2 \ Snd e\e2" +b6[intro]: "e\e' \ InL e \ InL e'" +b7[intro]: "e\e' \ InR e \ InR e'" +b8[intro]: "\e\InL e'; e1[x::=e']\e''\ \ Case e of inl x1 \ e1 | inr x2 \ e2 \ e''" +b9[intro]: "\e\InR e'; e2[x::=e']\e''\ \ Case e of inl x1 \ e1 | inr x2 \ e2 \ e''" + +consts + Ibig :: "(((nat\nat)\trmI)\((nat\nat)\trmI)) set" +syntax + "_Ibig_judge" :: "((nat\nat)\trmI)\((nat\nat)\trmI)\bool" ("_ I\ _" [80,80] 80) +translations + "(m,e1) I\ (m',e2)" \ "((m,e1),(m',e2)) \ Ibig" + +inductive Ibig +intros +m0[intro]: "(m,ILam [x].e) I\ (m,ILam [x].e)" +m1[intro]: "\(m,e1)I\(m',ILam [x].e); (m',e2)I\(m'',e3); (m'',e[x::=e3])I\(m''',e4)\ + \ (m,IApp e1 e2) I\ (m''',e4)" +m2[intro]: "(m,IUnit) I\ (m,IUnit)" +m3[intro]: "(m,INat(n))I\(m,INat(n))" +m4[intro]: "(m,e)I\(m',INat(n)) \ (m,ISucc(e))I\(m',INat(n+1))" +m5[intro]: "(m,e)I\(m',INat(n)) \ (m,IRef(e))I\(m',INat(m' n))" +m6[intro]: "\(m,e1)I\(m',INat(n1)); (m',e2)I\(m'',INat(n2))\ \ (m,e1\e2)I\(m''(n1:=n2),IUnit)" +m7[intro]: "\(m,e1)I\(m',IUnit); (m',e2)I\(m'',e)\ \ (m,e1;;e2)I\(m'',e)" +m8[intro]: "\(m,e)I\(m',INat(n)); n\0; (m',e1)I\(m'',e)\ \ (m,Iif e e1 e2)I\(m'',e)" +m9[intro]: "\(m,e)I\(m',INat(0)); (m',e2)I\(m'',e)\ \ (m,Iif e e1 e2)I\(m'',e)" + +text {* Translation functions *} + +constdefs + trans_data :: "data \ tyI" + "trans_data \ \_. DataI(NatI)" + + trans_arrow :: "ty \ ty \ tyI \ tyI \ tyI" + "trans_arrow \ \_ _ r1 r2. ArrowI r1 r2" + + trans_type::"ty \ tyI" + "trans_type \ \ ty_rec (trans_data) (trans_arrow) \" + +lemma trans_type: + shows "trans_type (Data \) = DataI(NatI)" + and "trans_type (\1\\2) = (trans_type \1)\(trans_type \2)" +apply(unfold trans_type_def) +apply(rule trans) +apply(rule ty.recs[where P="\_. True", simplified]) +apply(simp add: trans_data_def) +apply(rule trans) +apply(rule ty.recs[where P="\_. True", simplified]) +apply(simp add: trans_arrow_def) +done + +constdefs + trans_Var :: "name \ trmI" + "trans_Var \ \x. IVar x" + + trans_App :: "trm \ trm \ trmI \ trmI \ trmI" + "trans_App \ \_ _ r1 r2. IApp r1 r2" + + trans_Lam :: "name \ trm \ trmI \ trmI" + "trans_Lam \ \a _ r. ILam [a].r" + + trans_Const :: "nat \ trmI" + "trans_Const \ \n. INat n" + + trans_Pr :: "trm \ trm \ trmI \ trmI \ trmI" + "trans_Pr \ \_ _ r1 r2. + let limit = IRef(INat 0) in + let v1 = r1 in + let v2 = r2 in + (((ISucc limit)\v1);;(ISucc(ISucc limit)\v2));;(INat 0 \ ISucc(ISucc(limit)))" + + trans_Fst :: "trm \ trmI \ trmI" + "trans_Fst \ \_ r. IRef (ISucc r)" + + trans_Snd :: "trm \ trmI \ trmI" + "trans_Snd \ \_ r. IRef (ISucc (ISucc r))" + + trans_InL :: "trm \ trmI \ trmI" + "trans_InL \ \_ r. + let limit = IRef(INat 0) in + let v = r in + (((ISucc limit)\INat(0));;(ISucc(ISucc limit)\v));;(INat 0 \ ISucc(ISucc(limit)))" + + trans_InR :: "trm \ trmI \ trmI" + "trans_InR \ \_ r. + let limit = IRef(INat 0) in + let v = r in + (((ISucc limit)\INat(1));;(ISucc(ISucc limit)\v));;(INat 0 \ ISucc(ISucc(limit)))" + + trans_Case :: "trm \ name \ trm \ name \ trm \ trmI \ trmI \ trmI \ trmI" + "trans_Case \ \_ x1 _ x2 _ r r1 r2. + let v = r in + let v1 = r1 in + let v2 = r2 in + Iif (IRef (ISucc v)) (v2[x2::=IRef (ISucc (ISucc v))]) (v1[x1::=IRef (ISucc (ISucc v))])" + + trans :: "trm \ trmI" + "trans t \ trm_rec (trans_Var) (trans_Lam) (trans_App) + (trans_Const) (trans_Pr) (trans_Fst) (trans_Snd) + (trans_InL) (trans_InR) (trans_Case) t" + +lemma fin_supp_trans: + shows "finite ((supp (trans_Var))::name set)" + and "finite ((supp (trans_Lam))::name set)" + and "finite ((supp (trans_App))::name set)" + and "finite ((supp (trans_Const))::name set)" + and "finite ((supp (trans_Pr))::name set)" + and "finite ((supp (trans_Fst))::name set)" + and "finite ((supp (trans_Snd))::name set)" + and "finite ((supp (trans_InL))::name set)" + and "finite ((supp (trans_InR))::name set)" + and "finite ((supp (trans_Case))::name set)" +apply - +apply(finite_guess add: fs_name1 trans_Var_def) +apply(finite_guess add: fs_name1 trans_Lam_def) +apply(finite_guess add: fs_name1 trans_App_def) +apply(finite_guess add: fs_name1 trans_Const_def) +apply(finite_guess add: fs_name1 trans_Pr_def Let_def perm_nat_def) +apply(finite_guess add: fs_name1 trans_Fst_def) +apply(finite_guess add: fs_name1 trans_Snd_def) +apply(finite_guess add: fs_name1 trans_InL_def Let_def perm_nat_def) +apply(finite_guess add: fs_name1 trans_InR_def Let_def perm_nat_def) +apply(finite_guess add: fs_name1 trans_Case_def Let_def Isubst_eqvt) +done + +lemma fcb_trans_Lam: + shows "x\(trans_Lam) x t r" + by (simp add: trans_Lam_def abs_fresh) + +lemma fcb_trans_Case: + assumes a: "x\r" "x\r2" "y\r" "y\r1" + shows "x\(trans_Case) e x e1 y e2 r r1 r2 \ y\(trans_Case) e x e1 y e2 r r1 r2" + using a + by (simp add: trans_Case_def abs_fresh Isubst_fresh) + +lemma trans: + shows "trans (Var x) = (IVar x)" + and "trans (App e1 e2) = IApp (trans e1) (trans e2)" + and "trans (Lam [x].e) = ILam [x].(trans e)" + and "trans (Const n) = INat n" + and "trans (Pr e1 e2) = + (let limit = IRef(INat 0) in + let v1 = (trans e1) in + let v2 = (trans e2) in + (((ISucc limit)\v1);;(ISucc(ISucc limit)\v2));;(INat 0 \ ISucc(ISucc(limit))))" + and "trans (Fst e) = IRef (ISucc (trans e))" + and "trans (Snd e) = IRef (ISucc (ISucc (trans e)))" + and "trans (InL e) = + (let limit = IRef(INat 0) in + let v = (trans e) in + (((ISucc limit)\INat(0));;(ISucc(ISucc limit)\v));;(INat 0 \ ISucc(ISucc(limit))))" + and "trans (InR e) = + (let limit = IRef(INat 0) in + let v = (trans e) in + (((ISucc limit)\INat(1));;(ISucc(ISucc limit)\v));;(INat 0 \ ISucc(ISucc(limit))))" + and "\x2\x1; x1\e; x1\e2; x2\e; x2\e1\ \ + trans (Case e of inl x1 \ e1 | inr x2 \ e2) = + (let v = (trans e) in + let v1 = (trans e1) in + let v2 = (trans e2) in + Iif (IRef (ISucc v)) (v2[x2::=IRef (ISucc (ISucc v))]) (v1[x1::=IRef (ISucc (ISucc v))]))" +apply(unfold trans_def) +apply(rule trans) +apply(rule trm.recs[where P="\_. True", simplified]) +apply(rule fin_supp_trans)+ +apply(simp add: fcb_trans_Lam) +apply(simp add: fcb_trans_Case) +apply(simp add: trans_Var_def) +apply(rule trans) +apply(rule trm.recs[where P="\_. True", simplified]) +apply(rule fin_supp_trans)+ +apply(simp add: fcb_trans_Lam) +apply(simp add: fcb_trans_Case) +apply(simp add: trans_App_def) +apply(rule trans) +apply(rule trm.recs[where P="\_. True", simplified]) +apply(rule fin_supp_trans)+ +apply(simp add: fcb_trans_Lam) +apply(simp add: fcb_trans_Case) +apply(fresh_guess add: fs_name1 trans_Var_def perm_if eq_eqvt) +apply(fresh_guess add: fs_name1 trans_Lam_def) +apply(fresh_guess add: fs_name1 trans_App_def) +apply(fresh_guess add: fs_name1 trans_Const_def) +apply(fresh_guess add: fs_name1 trans_Pr_def Let_def perm_nat_def) +apply(fresh_guess add: fs_name1 trans_Fst_def) +apply(fresh_guess add: fs_name1 trans_Snd_def) +apply(fresh_guess add: fs_name1 trans_InL_def Let_def perm_nat_def) +apply(fresh_guess add: fs_name1 trans_InR_def Let_def perm_nat_def) +apply(fresh_guess add: fs_name1 trans_Case_def Let_def Isubst_eqvt) +apply(simp add: trans_Lam_def) +apply(rule trans) +apply(rule trm.recs[where P="\_. True", simplified]) +apply(rule fin_supp_trans)+ +apply(simp add: fcb_trans_Lam) +apply(simp add: fcb_trans_Case) +apply(simp add: trans_Const_def) +apply(rule trans) +apply(rule trm.recs[where P="\_. True", simplified]) +apply(rule fin_supp_trans)+ +apply(simp add: fcb_trans_Lam) +apply(simp add: fcb_trans_Case) +apply(simp add: trans_Pr_def Let_def) +apply(rule trans) +apply(rule trm.recs[where P="\_. True", simplified]) +apply(rule fin_supp_trans)+ +apply(simp add: fcb_trans_Lam) +apply(simp add: fcb_trans_Case) +apply(simp add: trans_Fst_def) +apply(rule trans) +apply(rule trm.recs[where P="\_. True", simplified]) +apply(rule fin_supp_trans)+ +apply(simp add: fcb_trans_Lam) +apply(simp add: fcb_trans_Case) +apply(simp add: trans_Snd_def) +apply(rule trans) +apply(rule trm.recs[where P="\_. True", simplified]) +apply(rule fin_supp_trans)+ +apply(simp add: fcb_trans_Lam) +apply(simp add: fcb_trans_Case) +apply(simp add: trans_InL_def Let_def) +apply(rule trans) +apply(rule trm.recs[where P="\_. True", simplified]) +apply(rule fin_supp_trans)+ +apply(simp add: fcb_trans_Lam) +apply(simp add: fcb_trans_Case) +apply(simp add: trans_InR_def Let_def) +apply(rule trans) +apply(rule trm.recs[where P="\_. True", simplified]) +apply(rule fin_supp_trans)+ +apply(simp add: fcb_trans_Lam) +apply(simp add: fcb_trans_Case) +apply(fresh_guess add: fs_name1 trans_Var_def perm_if eq_eqvt) +apply(fresh_guess add: fs_name1 trans_Var_def perm_if eq_eqvt) +apply(fresh_guess add: fs_name1 trans_Lam_def) +apply(fresh_guess add: fs_name1 trans_Lam_def) +apply(fresh_guess add: fs_name1 trans_App_def) +apply(fresh_guess add: fs_name1 trans_App_def) +apply(fresh_guess add: fs_name1 trans_Const_def) +apply(fresh_guess add: fs_name1 trans_Const_def) +apply(fresh_guess add: fs_name1 trans_Pr_def Let_def perm_nat_def) +apply(fresh_guess add: fs_name1 trans_Pr_def Let_def perm_nat_def) +apply(fresh_guess add: fs_name1 trans_Fst_def) +apply(fresh_guess add: fs_name1 trans_Fst_def) +apply(fresh_guess add: fs_name1 trans_Snd_def) +apply(fresh_guess add: fs_name1 trans_Snd_def) +apply(fresh_guess add: fs_name1 trans_InL_def Let_def perm_nat_def) +apply(fresh_guess add: fs_name1 trans_InL_def Let_def perm_nat_def) +apply(fresh_guess add: fs_name1 trans_InR_def Let_def perm_nat_def) +apply(fresh_guess add: fs_name1 trans_InR_def Let_def perm_nat_def) +apply(fresh_guess add: fs_name1 trans_Case_def Let_def Isubst_eqvt) +apply(fresh_guess add: fs_name1 trans_Case_def Let_def Isubst_eqvt) +apply(assumption)+ +apply(simp add: trans_Case_def Let_def) +done + + +end \ No newline at end of file