(* ID: $Id: NBE.thy,v 1.5 2007/08/28 09:51:27 wenzelm Exp $ Author: Klaus Aehlig, Tobias Nipkow Work in progress *) theory NBE imports Main Executable_Set begin axiomatization where unproven: "PROP A" declare Let_def[simp] consts_code undefined ("(raise Match)") (*typedecl const_name*) types lam_var_name = nat ml_var_name = nat const_name = nat datatype tm = Ct const_name | Vt lam_var_name | Lam tm | At tm tm | term_of ml (* function 'to_term' *) and ml = (* rep of universal datatype *) C const_name "ml list" | V lam_var_name "ml list" | Fun ml "ml list" nat | "apply" ml ml (* function 'apply' *) (* ML *) | V_ML ml_var_name | A_ML ml "ml list" | Lam_ML ml | CC const_name (* ref to compiled code *) lemma [simp]: "x ∈ set vs ==> size x < Suc (ml_list_size1 vs)" by (induct vs) auto lemma [simp]: "x ∈ set vs ==> size x < Suc (ml_list_size2 vs)" by (induct vs) auto lemma [simp]:"x ∈ set vs ==> size x < Suc (size v + ml_list_size3 vs)" by (induct vs) auto lemma [simp]: "x ∈ set vs ==> size x < Suc (size v + ml_list_size4 vs)" by (induct vs) auto locale Vars = fixes r s t:: tm and rs ss ts :: "tm list" and u v w :: ml and us vs ws :: "ml list" and nm :: const_name and x :: lam_var_name and X :: ml_var_name inductive_set Pure_tms :: "tm set" where "Ct s : Pure_tms" | "Vt x : Pure_tms" | "t : Pure_tms ==> Lam t : Pure_tms" | "s : Pure_tms ==> t : Pure_tms ==> At s t : Pure_tms" consts R :: "(const_name * tm list * tm)set" (* reduction rules *) compR :: "(const_name * ml list * ml)set" (* compiled reduction rules *) fun lift_tm :: "nat => tm => tm" ("lift") and lift_ml :: "nat => ml => ml" ("lift") where "lift i (Ct nm) = Ct nm" | "lift i (Vt x) = Vt(if x < i then x else x+1)" | "lift i (Lam t) = Lam (lift (i+1) t)" | "lift i (At s t) = At (lift i s) (lift i t)" | "lift i (term_of v) = term_of (lift i v)" | "lift i (C nm vs) = C nm (map (lift i) vs)" | "lift i (V x vs) = V (if x < i then x else x+1) (map (lift i) vs)" | "lift i (Fun v vs n) = Fun (lift i v) (map (lift i) vs) n" | "lift i (apply u v) = apply (lift i u) (lift i v)" | "lift i (V_ML X) = V_ML X" | "lift i (A_ML v vs) = A_ML (lift i v) (map (lift i) vs)" | "lift i (Lam_ML v) = Lam_ML (lift i v)" | "lift i (CC nm) = CC nm" (* termination apply (relation "measure (sum_case (%(i,t). size t) (%(i,v). size v))") apply auto *) fun lift_tm_ML :: "nat => tm => tm" ("liftML") and lift_ml_ML :: "nat => ml => ml" ("liftML") where "liftML i (Ct nm) = Ct nm" | "liftML i (Vt x) = Vt x" | "liftML i (Lam t) = Lam (liftML i t)" | "liftML i (At s t) = At (liftML i s) (liftML i t)" | "liftML i (term_of v) = term_of (liftML i v)" | "liftML i (C nm vs) = C nm (map (liftML i) vs)" | "liftML i (V x vs) = V x (map (liftML i) vs)" | "liftML i (Fun v vs n) = Fun (liftML i v) (map (liftML i) vs) n" | "liftML i (apply u v) = apply (liftML i u) (liftML i v)" | "liftML i (V_ML X) = V_ML (if X < i then X else X+1)" | "liftML i (A_ML v vs) = A_ML (liftML i v) (map (liftML i) vs)" | "liftML i (Lam_ML v) = Lam_ML (liftML (i+1) v)" | "liftML i (CC nm) = CC nm" (* termination by (relation "measure (sum_case (%(i,t). size t) (%(i,v). size v))") auto *) constdefs cons :: "tm => (nat => tm) => (nat => tm)" (infix "##" 65) "t##f ≡ λi. case i of 0 => t | Suc j => lift 0 (f j)" cons_ML :: "ml => (nat => ml) => (nat => ml)" (infix "##" 65) "v##f ≡ λi. case i of 0 => v::ml | Suc j => liftML 0 (f j)" (* Only for pure terms! *) consts subst :: "(nat => tm) => tm => tm" primrec "subst f (Ct nm) = Ct nm" "subst f (Vt x) = f x" "subst f (Lam t) = Lam (subst (Vt 0 ## f) t)" "subst f (At s t) = At (subst f s) (subst f t)" lemma size_lift[simp]: shows "size(lift i t) = size(t::tm)" and "size(lift i (v::ml)) = size v" and "ml_list_size1 (map (lift i) vs) = ml_list_size1 vs" and "ml_list_size2 (map (lift i) vs) = ml_list_size2 vs" and "ml_list_size3 (map (lift i) vs) = ml_list_size3 vs" and "ml_list_size4 (map (lift i) vs) = ml_list_size4 vs" by (induct arbitrary: i and i and i and i and i and i rule: tm_ml.inducts) simp_all lemma size_lift_ML[simp]: shows "size(liftML i t) = size(t::tm)" and "size(liftML i (v::ml)) = size v" and "ml_list_size1 (map (liftML i) vs) = ml_list_size1 vs" and "ml_list_size2 (map (liftML i) vs) = ml_list_size2 vs" and "ml_list_size3 (map (liftML i) vs) = ml_list_size3 vs" and "ml_list_size4 (map (liftML i) vs) = ml_list_size4 vs" by (induct arbitrary: i and i and i and i and i and i rule: tm_ml.inducts) simp_all fun subst_ml_ML :: "(nat => ml) => ml => ml" ("substML") and subst_tm_ML :: "(nat => ml) => tm => tm" ("substML") where "substML f (Ct nm) = Ct nm" | "substML f (Vt x) = Vt x" | "substML f (Lam t) = Lam (substML (lift 0 o f) t)" | "substML f (At s t) = At (substML f s) (substML f t)" | "substML f (term_of v) = term_of (substML f v)" | "substML f (C nm vs) = C nm (map (substML f) vs)" | "substML f (V x vs) = V x (map (substML f) vs)" | "substML f (Fun v vs n) = Fun (substML f v) (map (substML f) vs) n" | "substML f (apply u v) = apply (substML f u) (substML f v)" | "substML f (V_ML X) = f X" | "substML f (A_ML v vs) = A_ML (substML f v) (map (substML f) vs)" | "substML f (Lam_ML v) = Lam_ML (substML (V_ML 0 ## f) v)" | "substML f (CC nm) = CC nm" (* FIXME currrently needed for code generator *) lemmas [code] = lift_tm_ML.simps lift_ml_ML.simps lemmas [code] = lift_tm.simps lift_ml.simps lemmas [code] = subst_tm_ML.simps subst_ml_ML.simps abbreviation subst_decr :: "nat => tm => nat => tm" where "subst_decr k t == %n. if n<k then Vt n else if n=k then t else Vt(n - 1)" abbreviation subst_decr_ML :: "nat => ml => nat => ml" where "subst_decr_ML k v == %n. if n<k then V_ML n else if n=k then v else V_ML(n - 1)" abbreviation subst1 :: "tm => tm => nat => tm" ("(_/[_'/_])" [300, 0, 0] 300) where "s[t/k] == subst (subst_decr k t) s" abbreviation subst1_ML :: "ml => ml => nat => ml" ("(_/[_'/_])" [300, 0, 0] 300) where "u[v/k] == substML (subst_decr_ML k v) u" lemma size_subst_ML[simp]: shows "(!x. size(f x) = 0) --> size(substML f t) = size(t::tm)" and "(!x. size(f x) = 0) --> size(substML f (v::ml)) = size v" and "(!x. size(f x) = 0) --> ml_list_size1 (map (substML f) vs) = ml_list_size1 vs" and "(!x. size(f x) = 0) --> ml_list_size2 (map (substML f) vs) = ml_list_size2 vs" and "(!x. size(f x) = 0) --> ml_list_size3 (map (substML f) vs) = ml_list_size3 vs" and "(!x. size(f x) = 0) --> ml_list_size4 (map (substML f) vs) = ml_list_size4 vs" apply (induct arbitrary: f and f and f and f and f and f rule: tm_ml.inducts) apply (simp_all add:cons_ML_def split:nat.split) done lemma lift_lift: includes Vars shows "i < k+1 ==> lift (Suc k) (lift i t) = lift i (lift k t)" and "i < k+1 ==> lift (Suc k) (lift i v) = lift i (lift k v)" apply(induct t and v arbitrary: i and i rule:lift_tm_lift_ml.induct) apply(simp_all add:map_compose[symmetric]) done corollary lift_o_lift: shows "i < k+1 ==> lift_tm (Suc k) o (lift_tm i) = lift_tm i o lift_tm k" and "i < k+1 ==> lift_ml (Suc k) o (lift_ml i) = lift_ml i o lift_ml k" by(rule ext, simp add:lift_lift)+ lemma lift_lift_ML: includes Vars shows "i < k+1 ==> liftML (Suc k) (liftML i t) = liftML i (liftML k t)" and "i < k+1 ==> liftML (Suc k) (liftML i v) = liftML i (liftML k v)" apply(induct t and v arbitrary: i and i rule:lift_tm_ML_lift_ml_ML.induct) apply(simp_all add:map_compose[symmetric]) done lemma lift_lift_ML_comm: includes Vars shows "lift j (liftML i t) = liftML i (lift j t)" and "lift j (liftML i v) = liftML i (lift j v)" apply(induct t and v arbitrary: i j and i j rule:lift_tm_ML_lift_ml_ML.induct) apply(simp_all add:map_compose[symmetric]) done lemma [simp]: "V_ML 0 ## subst_decr_ML k v = subst_decr_ML (Suc k) (liftML 0 v)" by(rule ext)(simp add:cons_ML_def split:nat.split) lemma [simp]: "lift 0 o subst_decr_ML k v = subst_decr_ML k (lift 0 v)" by(rule ext)(simp add:cons_ML_def split:nat.split) lemma subst_lift_id[simp]: includes Vars shows "substML (subst_decr_ML k v) (liftML k t) = t" and "(liftML k u)[v/k] = u" apply(induct k t and k u arbitrary: v and v rule: lift_tm_ML_lift_ml_ML.induct) apply (simp_all add:map_idI map_compose[symmetric]) apply (simp cong:if_cong) done inductive_set tRed :: "(tm * tm)set" (* beta + R reduction on pure terms *) and tred :: "[tm, tm] => bool" (infixl "->" 50) where "s -> t == (s, t) ∈ tRed" | "At (Lam t) s -> t[s/0]" | "(nm,ts,t) : R ==> foldl At (Ct nm) (map (subst rs) ts) -> subst rs t" | "t -> t' ==> Lam t -> Lam t'" | "s -> s' ==> At s t -> At s' t" | "t -> t' ==> At s t -> At s t'" abbreviation treds :: "[tm, tm] => bool" (infixl "->*" 50) where "s ->* t == (s, t) ∈ tRed^*" inductive_set tRed_list :: "(tm list * tm list) set" and treds_list :: "[tm list, tm list] => bool" (infixl "->*" 50) where "ss ->* ts == (ss, ts) ∈ tRed_list" | "[] ->* []" | "ts ->* ts' ==> t ->* t' ==> t#ts ->* t'#ts'" declare tRed_list.intros[simp] lemma tRed_list_refl[simp]: includes Vars shows "ts ->* ts" by(induct ts) auto fun ML_closed :: "nat => ml => bool" and ML_closed_t :: "nat => tm => bool" where "ML_closed i (C nm vs) = (ALL v:set vs. ML_closed i v)" | "ML_closed i (V nm vs) = (ALL v:set vs. ML_closed i v)" | "ML_closed i (Fun f vs n) = (ML_closed i f & (ALL v:set vs. ML_closed i v))" | "ML_closed i (A_ML v vs) = (ML_closed i v & (ALL v:set vs. ML_closed i v))" | "ML_closed i (apply v w) = (ML_closed i v & ML_closed i w)" | "ML_closed i (CC nm) = True" | "ML_closed i (V_ML X) = (X<i)" | "ML_closed i (Lam_ML v) = ML_closed (i+1) v" | "ML_closed_t i (term_of v) = ML_closed i v" | "ML_closed_t i (At r s) = (ML_closed_t i r & ML_closed_t i s)" | "ML_closed_t i (Lam t) = (ML_closed_t i t)" | "ML_closed_t i v = True" thm ML_closed.simps ML_closed_t.simps inductive_set Red :: "(ml * ml)set" and Redt :: "(tm * tm)set" and Redl :: "(ml list * ml list)set" and red :: "[ml, ml] => bool" (infixl "=>" 50) and redl :: "[ml list, ml list] => bool" (infixl "=>" 50) and redt :: "[tm, tm] => bool" (infixl "=>" 50) and reds :: "[ml, ml] => bool" (infixl "=>*" 50) and redts :: "[tm, tm] => bool" (infixl "=>*" 50) where "s => t == (s, t) ∈ Red" | "s => t == (s, t) ∈ Redl" | "s => t == (s, t) ∈ Redt" | "s =>* t == (s, t) ∈ Red^*" | "s =>* t == (s, t) ∈ Redt^*" (* ML *) | "A_ML (Lam_ML u) [v] => u[v/0]" (* compiled rules *) | "(nm,vs,v) : compR ==> ALL i. ML_closed 0 (f i) ==> A_ML (CC nm) (map (substML f) vs) => substML f v" (* apply *) | apply_Fun1: "apply (Fun f vs (Suc 0)) v => A_ML f (vs @ [v])" | apply_Fun2: "n > 0 ==> apply (Fun f vs (Suc n)) v => Fun f (vs @ [v]) n" | apply_C: "apply (C nm vs) v => C nm (vs @ [v])" | apply_V: "apply (V x vs) v => V x (vs @ [v])" (* term_of *) | term_of_C: "term_of (C nm vs) => foldl At (Ct nm) (map term_of vs)" | term_of_V: "term_of (V x vs) => foldl At (Vt x) (map term_of vs)" | term_of_Fun: "term_of(Fun vf vs n) => Lam (term_of ((apply (lift 0 (Fun vf vs n)) (V_ML 0))[V 0 []/0]))" (* Context *) | ctxt_Lam: "t => t' ==> Lam t => Lam t'" | ctxt_At1: "s => s' ==> At s t => At s' t" | ctxt_At2: "t => t' ==> At s t => At s t'" | ctxt_term_of: "v => v' ==> term_of v => term_of v'" | ctxt_C: "vs => vs' ==> C nm vs => C nm vs'" | ctxt_V: "vs => vs' ==> V x vs => V x vs'" | ctxt_Fun1: "f => f' ==> Fun f vs n => Fun f' vs n" | ctxt_Fun3: "vs => vs' ==> Fun f vs n => Fun f vs' n" | ctxt_apply1: "s => s' ==> apply s t => apply s' t" | ctxt_apply2: "t => t' ==> apply s t => apply s t'" | ctxt_A_ML1: "f => f' ==> A_ML f vs => A_ML f' vs" | ctxt_A_ML2: "vs => vs' ==> A_ML f vs => A_ML f vs'" | ctxt_list1: "v => v' ==> v#vs => v'#vs" | ctxt_list2: "vs => vs' ==> v#vs => v#vs'" consts ar :: "const_name => nat" axioms ar_pos: "ar nm > 0" types env = "ml list" consts eval :: "tm => env => ml" primrec "eval (Vt x) e = e!x" "eval (Ct nm) e = Fun (CC nm) [] (ar nm)" "eval (At s t) e = apply (eval s e) (eval t e)" "eval (Lam t) e = Fun (Lam_ML (eval t ((V_ML 0) # map (liftML 0) e))) [] 1" fun size' :: "ml => nat" where "size' (C nm vs) = (∑v\<leftarrow>vs. size' v)+1" | "size' (V nm vs) = (∑v\<leftarrow>vs. size' v)+1" | "size' (Fun f vs n) = (size' f + (∑v\<leftarrow>vs. size' v))+1" | "size' (A_ML v vs) = (size' v + (∑v\<leftarrow>vs. size' v))+1" | "size' (apply v w) = (size' v + size' w)+1" | "size' (CC nm) = 1" | "size' (V_ML X) = 1" | "size' (Lam_ML v) = size' v + 1" lemma listsum_size'[simp]: "v ∈ set vs ==> size' v < Suc(listsum (map size' vs))" by (rule unproven) corollary cor_listsum_size'[simp]: "v ∈ set vs ==> size' v < Suc(m + listsum (map size' vs))" using listsum_size'[of v vs] by arith lemma size_subst_ML[simp]: includes Vars assumes A: "!i. size(f i) = 0" shows "size(substML f t) = size(t)" and "size(substML f v) = size(v)" and "ml_list_size1 (map (substML f) vs) = ml_list_size1 vs" and "ml_list_size2 (map (substML f) vs) = ml_list_size2 vs" and "ml_list_size3 (map (substML f) vs) = ml_list_size3 vs" and "ml_list_size4 (map (substML f) vs) = ml_list_size4 vs" by (induct rule: tm_ml.inducts) (simp_all add: A cons_ML_def split:nat.split) lemma [simp]: "∀i j. size'(f i) = size'(V_ML j) ==> size' (substML f v) = size' v" by (rule unproven) lemma [simp]: "size' (lift i v) = size' v" by (rule unproven) (* the kernel function as in Section 4.1 of "Operational aspects…" *) function kernel :: "ml => tm" ("_!" 300) where "(C nm vs)! = foldl At (Ct nm) (map kernel vs)" | "(Lam_ML v)! = Lam (((lift 0 v)[V 0 []/0])!)" | "(Fun f vs n)! = foldl At (f!) (map kernel vs)" | "(A_ML v vs)! = foldl At (v!) (map kernel vs)" | "(apply v w)! = At (v!) (w!)" | "(CC nm)! = Ct nm" | "(V x vs)! = foldl At (Vt x) (map kernel vs)" | "(V_ML X)! = undefined" by pat_completeness auto termination by(relation "measure size'") auto consts kernelt :: "tm => tm" ("_!" 300) primrec "(Ct nm)! = Ct nm" "(term_of v)! = v!" "(Vt x)! = Vt x" "(At s t)! = At (s!) (t!)" "(Lam t)! = Lam (t!)" abbreviation kernels :: "ml list => tm list" ("_!" 300) where "vs ! == map kernel vs" (* soundness of the code generator *) axioms compiler_correct: "(nm, vs, v) : compR ==> ALL i. ML_closed 0 (f i) ==> (nm, (map (substML f) vs)!, (substML f v)!) : R" consts free_vars :: "tm => lam_var_name set" primrec "free_vars (Ct nm) = {}" "free_vars (Vt x) = {x}" "free_vars (Lam t) = {i. EX j : free_vars t. j = i+1}" "free_vars (At s t) = free_vars s ∪ free_vars t" lemma [simp]: "t : Pure_tms ==> liftML k t = t" by (erule Pure_tms.induct) simp_all lemma kernel_pure: includes Vars assumes "t : Pure_tms" shows "t! = t" using assms by (induct) simp_all lemma lift_eval: "t : Pure_tms ==> ALL e k. (ALL i : free_vars t. i < size e) --> lift k (eval t e) = eval t (map (lift k) e)" apply(induct set:Pure_tms) apply simp_all apply clarsimp apply(erule_tac x = "V_ML 0 # map (liftML 0) e" in allE) apply simp apply(erule impE) apply clarsimp apply(case_tac i)apply simp apply simp apply (simp add: map_compose[symmetric]) apply (simp add: o_def lift_lift_ML_comm) done lemma lift_ML_eval[rule_format]: "t : Pure_tms ==> ALL e k. (ALL i : free_vars t. i < size e) --> liftML k (eval t e) = eval t (map (liftML k) e)" apply(induct set:Pure_tms) apply simp_all apply clarsimp apply(erule_tac x = "V_ML 0 # map (liftML 0) e" in allE) apply simp apply(erule impE) apply clarsimp apply(case_tac i)apply simp apply simp apply (simp add: map_compose[symmetric]) apply (simp add:o_def lift_lift_ML) done lemma [simp]: includes Vars shows "(v ## f) 0 = v" by(simp add:cons_ML_def) lemma [simp]: includes Vars shows "(v ## f) (Suc n) = liftML 0 (f n)" by(simp add:cons_ML_def) lemma lift_o_shift: "lift k o (V_ML 0 ## f) = (V_ML 0 ## (lift k o f))" apply(rule ext) apply (simp add:cons_ML_def lift_lift_ML_comm split:nat.split) done lemma lift_subst_ML: shows "lift_tm k (substML f t) = substML (lift_ml k o f) (lift_tm k t)" and "lift_ml k (substML f v) = substML (lift_ml k o f) (lift_ml k v)" apply (induct t and v arbitrary: f k and f k rule: lift_tm_lift_ml.induct) apply (simp_all add:map_compose[symmetric] o_assoc lift_o_lift lift_o_shift) done corollary lift_subst_ML1: "∀v k. lift_ml 0 (u[v/k]) = (lift_ml 0 u)[lift 0 v/k]" apply(rule measure_induct[where f = "size" and a = u]) apply(case_tac x) apply(simp_all add:lift_lift map_compose[symmetric] lift_subst_ML) apply(subst lift_lift_ML_comm)apply simp done lemma lift_ML_lift_ML: includes Vars shows "i < k+1 ==> liftML (Suc k) (liftML i t) = liftML i (liftML k t)" and "i < k+1 ==> liftML (Suc k) (liftML i v) = liftML i (liftML k v)" apply (induct k t and k v arbitrary: i k and i k rule: lift_tm_ML_lift_ml_ML.induct) apply(simp_all add:map_compose[symmetric]) done corollary lift_ML_o_lift_ML: shows "i < k+1 ==> lift_tm_ML (Suc k) o (lift_tm_ML i) = lift_tm_ML i o lift_tm_ML k" and "i < k+1 ==> lift_ml_ML (Suc k) o (lift_ml_ML i) = lift_ml_ML i o lift_ml_ML k" by(rule ext, simp add:lift_ML_lift_ML)+ abbreviation insrt where "insrt k f == (%i. if i<k then lift_ml_ML k (f i) else if i=k then V_ML k else lift_ml_ML k (f(i - 1)))" lemma subst_insrt_lift: includes Vars shows "substML (insrt k f) (liftML k t) = liftML k (substML f t)" and "substML (insrt k f) (liftML k v) = liftML k (substML f v)" apply (induct k t and k v arbitrary: f k and f k rule: lift_tm_ML_lift_ml_ML.induct) apply (simp_all add:map_compose[symmetric] o_assoc lift_o_lift lift_o_shift) apply(subgoal_tac "lift 0 o insrt k f = insrt k (lift 0 o f)") apply simp apply(rule ext) apply (simp add:lift_lift_ML_comm) apply(subgoal_tac "V_ML 0 ## insrt k f = insrt (Suc k) (V_ML 0 ## f)") apply simp apply(rule ext) apply (simp add:lift_ML_lift_ML cons_ML_def split:nat.split) done corollary subst_cons_lift: includes Vars shows "substML (V_ML 0 ## f) o (lift_ml_ML 0) = lift_ml_ML 0 o (subst_ml_ML f)" apply(rule ext) apply(simp add: cons_ML_def subst_insrt_lift[symmetric]) apply(subgoal_tac "nat_case (V_ML 0) (λj. liftML 0 (f j)) = (λi. if i = 0 then V_ML 0 else liftML 0 (f (i - 1)))") apply simp apply(rule ext, simp split:nat.split) done lemma subst_eval[rule_format]: "t : Pure_tms ==> ALL f e. (ALL i : free_vars t. i < size e) --> substML f (eval t e) = eval t (map (substML f) e)" apply(induct set:Pure_tms) apply simp_all apply clarsimp apply(erule_tac x="V_ML 0 ## f" in allE) apply(erule_tac x= "(V_ML 0 # map (liftML 0) e)" in allE) apply(erule impE) apply clarsimp apply(case_tac i)apply simp apply simp apply (simp add:subst_cons_lift map_compose[symmetric]) done theorem kernel_eval[rule_format]: includes Vars shows "t : Pure_tms ==> ALL e. (ALL i : free_vars t. i < size e) --> (ALL i < size e. e!i = V i []) --> (eval t e)! = t!" apply(induct set:Pure_tms) apply simp_all apply clarsimp apply(subst lift_eval) apply simp apply clarsimp apply(case_tac i)apply simp apply simp apply(subst subst_eval) apply simp apply clarsimp apply(case_tac i)apply simp apply simp apply(erule_tac x="map (substML (λn. if n = 0 then V 0 [] else V_ML (n - 1))) (map (lift 0) (V_ML 0 # map (liftML 0) e))" in allE) apply(erule impE) apply(clarsimp) apply(case_tac i)apply simp apply simp apply(erule impE) apply(clarsimp) apply(case_tac i)apply simp apply simp apply simp done (* lemma subst_ML_compose: "subst_ml_ML f2 (subst_ml_ML f1 v) = subst_ml_ML (%i. subst_ml_ML f2 (f1 i)) v" by (rule unproven) *) lemma map_eq_iff_nth: "(map f xs = map g xs) = (!i<size xs. f(xs!i) = g(xs!i))" by (rule unproven) lemma [simp]: includes Vars shows "ML_closed k v ==> liftML k v = v" by (rule unproven) lemma [simp]: includes Vars shows "ML_closed 0 v ==> substML f v = v" by (rule unproven) lemma [simp]: includes Vars shows "ML_closed k v ==> ML_closed k (lift m v)" by (rule unproven) lemma red_Lam[simp]: includes Vars shows "t ->* t' ==> Lam t ->* Lam t'" apply(induct rule:rtrancl_induct) apply(simp_all) apply(blast intro: rtrancl_into_rtrancl tRed.intros) done lemma red_At1[simp]: includes Vars shows "t ->* t' ==> At t s ->* At t' s" apply(induct rule:rtrancl_induct) apply(simp_all) apply(blast intro: rtrancl_into_rtrancl tRed.intros) done lemma red_At2[simp]: includes Vars shows "t ->* t' ==> At s t ->* At s t'" apply(induct rule:rtrancl_induct) apply(simp_all) apply(blast intro:rtrancl_into_rtrancl tRed.intros) done lemma tRed_list_foldl_At: "ts ->* ts' ==> s ->* s' ==> foldl At s ts ->* foldl At s' ts'" apply(induct arbitrary:s s' rule:tRed_list.induct) apply simp apply simp apply(blast dest: red_At1 red_At2 intro:rtrancl_trans) done lemma [trans]: "s = t ==> t -> t' ==> s -> t'" by simp lemma subst_foldl[simp]: "subst f (foldl At s ts) = foldl At (subst f s) (map (subst f) ts)" by (induct ts arbitrary: s) auto lemma foldl_At_size: "size ts = size ts' ==> foldl At s ts = foldl At s' ts' <-> s = s' & ts = ts'" by (induct arbitrary: s s' rule:list_induct2) simp_all consts depth_At :: "tm => nat" primrec "depth_At(Ct cn) = 0" "depth_At(Vt x) = 0" "depth_At(Lam t) = 0" "depth_At(At s t) = depth_At s + 1" "depth_At(term_of v) = 0" lemma depth_At_foldl: "depth_At(foldl At s ts) = depth_At s + size ts" by (induct ts arbitrary: s) simp_all lemma foldl_At_eq_length: "foldl At s ts = foldl At s ts' ==> length ts = length ts'" apply(subgoal_tac "depth_At(foldl At s ts) = depth_At(foldl At s ts')") apply(erule thin_rl) apply (simp add:depth_At_foldl) apply simp done lemma foldl_At_eq[simp]: "foldl At s ts = foldl At s ts' <-> ts = ts'" apply(rule) prefer 2 apply simp apply(blast dest:foldl_At_size foldl_At_eq_length) done lemma [simp]: "foldl At s ts ! = foldl At (s!) (map kernelt ts)" by (induct ts arbitrary: s) simp_all lemma [simp]: "(kernelt o term_of) = kernel" by(rule ext) simp lemma shift_subst_decr: "Vt 0 ## subst_decr k t = subst_decr (Suc k) (lift 0 t)" apply(rule ext) apply (simp add:cons_def split:nat.split) done lemma [simp]: "lift k (foldl At s ts) = foldl At (lift k s) (map (lift k) ts)" by(induct ts arbitrary:s) simp_all subsection "Horrible detour" definition "liftn n == lift_ml 0 ^ n" lemma [simp]: "liftn n (C i vs) = C i (map (liftn n) vs)" apply(unfold liftn_def) apply(induct n) apply (simp_all add: map_compose[symmetric]) done lemma [simp]: "liftn n (CC nm) = CC nm" apply(unfold liftn_def) apply(induct n) apply (simp_all add: map_compose[symmetric]) done lemma [simp]: "liftn n (apply v w) = apply (liftn n v) (liftn n w)" apply(unfold liftn_def) apply(induct n) apply (simp_all add: map_compose[symmetric]) done lemma [simp]: "liftn n (A_ML v vs) = A_ML (liftn n v) (map (liftn n) vs)" apply(unfold liftn_def) apply(induct n) apply (simp_all add: map_compose[symmetric]) done lemma [simp]: "liftn n (Fun v vs i) = Fun (liftn n v) (map (liftn n) vs) i" apply(unfold liftn_def) apply(induct n) apply (simp_all add: map_compose[symmetric] id_def) done lemma [simp]: "liftn n (Lam_ML v) = Lam_ML (liftn n v)" apply(unfold liftn_def) apply(induct n) apply (simp_all add: map_compose[symmetric] id_def) done lemma liftn_liftn_add: "liftn m (liftn n v) = liftn (m+n) v" by(simp add:liftn_def funpow_add) lemma [simp]: "liftn n (V_ML k) = V_ML k" apply(unfold liftn_def) apply(induct n) apply (simp_all) done lemma liftn_lift_ML_comm: "liftn n (liftML 0 v) = liftML 0 (liftn n v)" apply(unfold liftn_def) apply(induct n) apply (simp_all add:lift_lift_ML_comm) done lemma liftn_cons: "liftn n ((V_ML 0 ## f) x) = (V_ML 0 ## (liftn n o f)) x" apply(simp add:cons_ML_def liftn_lift_ML_comm split:nat.split) done text{* End of horrible detour *} lemma kernel_subst1: "ML_closed 1 u ==> ML_closed 0 v ==> kernel( u[v/0]) = (kernel((lift 0 u)[V 0 []/0]))[kernel v/0]" by (rule unproven) lemma includes Vars shows foldl_Pure[simp]: "t : Pure_tms ==> ∀t∈set ts. t : Pure_tms ==> (!!s t. s : Pure_tms ==> t : Pure_tms ==> f s t : Pure_tms) ==> foldl f t ts ∈ Pure_tms" by(induct ts arbitrary: t) simp_all declare Pure_tms.intros[simp] lemma includes Vars shows "ML_closed 0 v ==> kernel v : Pure_tms" apply(induct rule:kernel.induct) apply simp_all apply(rule Pure_tms.intros); (* "ML_closed (Suc k) v ==> ML_closed k (lift 0 v)" *) by (rule unproven) lemma subst_Vt: includes Vars shows "subst Vt = id" by (rule unproven) (* apply(rule ext) apply(induct_tac x) apply simp_all done *) (* klappt noch nicht ganz *) theorem Red_sound: includes Vars shows "v => v' ==> ML_closed 0 v ==> v! ->* v'! & ML_closed 0 v'" and "t => t' ==> ML_closed_t 0 t ==> kernelt t ->* kernelt t' & ML_closed_t 0 t'" and "(vs :: ml list) => vs' ==> !v : set vs . ML_closed 0 v ==> map kernel vs ->* map kernel vs' & (! v':set vs'. ML_closed 0 v')" proof(induct rule:Red_Redt_Redl.inducts) fix u v let ?v = "A_ML (Lam_ML u) [v]" assume cl: "ML_closed 0 (A_ML (Lam_ML u) [v])" let ?u' = "(lift_ml 0 u)[V 0 []/0]" have "?v! = At (Lam ((?u')!)) (v !)" by simp also have "… -> (?u' !)[v!/0]" (is "_ -> ?R") by(rule tRed.intros) also have "?R = u[v/0]!" using cl apply(cut_tac u = "u" and v = "v" in kernel_subst1) apply(simp_all) done finally have "kernel(A_ML (Lam_ML u) [v]) ->* kernel(u[v/0])" (is ?A) by(rule r_into_rtrancl) moreover have "ML_closed 0 (u[v/0])" (is "?C") using cl apply simp by (rule unproven) ultimately show "?A & ?C" .. next case term_of_C thus ?case apply (auto simp:map_compose[symmetric])by (rule unproven) next fix f :: "nat => ml" and nm vs v assume f: "∀i. ML_closed 0 (f i)" and compR: "(nm, vs, v) ∈ compR" note tRed.intros(2)[OF compiler_correct[OF compR f], of Vt,simplified map_compose[symmetric]] hence red: "foldl At (Ct nm) (map (kernel o substML f) vs) -> (substML f v)!" (is "_ -> ?R") apply(simp add:map_compose) by (rule unproven) have "A_ML (CC nm) (map (substML f) vs)! = foldl At (Ct nm) (map (kernel o substML f) vs)" by (simp add:map_compose) also(* have "map (kernel o substML f) vs = map (subst (kernel o f)) (vs!)" using closed_subst_kernel(2)[OF compiled_V_free1[OF compR]] by (simp add:map_compose[symmetric]) also*) note red (*also have "?R = substML f v!" using closed_subst_kernel(2)[OF compiled_V_free2[OF compR]] by simp*) finally have "A_ML (CC nm) (map (substML f) vs)! ->* substML f v!" (is "?A") by(rule r_into_rtrancl) (* also have "?l = (substML fa (A_ML (CC nm) (map (substML f) vs)))!" (is "_ = ?l'") by (rule unproven) also have "?r = substML fa (substML f v)!" (is "_ = ?r'") by (rule unproven) finally have "?l' ->* ?r'" (is ?A) . *) moreover have "ML_closed 0 (substML f v)" (is "?C") by (rule unproven) ultimately show "?A & ?C" .. next case term_of_V thus ?case apply (auto simp:map_compose[symmetric]) by (rule unproven) next case (term_of_Fun vf vs n) hence "term_of (Fun vf vs n)! ->* Lam (term_of (apply (lift 0 (Fun vf vs n)) (V_ML 0)[V 0 []/0]))!" by - (rule unproven) moreover have "ML_closed_t 0 (Lam (term_of (apply (lift 0 (Fun vf vs n)) (V_ML 0)[V 0 []/0])))" by (rule unproven) ultimately show ?case .. next case apply_Fun1 thus ?case by simp next case apply_Fun2 thus ?case by simp next case apply_C thus ?case by simp next case apply_V thus ?case by simp next case ctxt_Lam thus ?case by(auto) next case ctxt_At1 thus ?case by(auto) next case ctxt_At2 thus ?case by (auto) next case ctxt_term_of thus ?case by (auto) next case ctxt_C thus ?case by (fastsimp simp:tRed_list_foldl_At) next case ctxt_V thus ?case by (fastsimp simp:tRed_list_foldl_At) next case ctxt_Fun1 thus ?case by (fastsimp simp:tRed_list_foldl_At) next case ctxt_Fun3 thus ?case by (fastsimp simp:tRed_list_foldl_At) next case ctxt_apply1 thus ?case by auto next case ctxt_apply2 thus ?case by auto next case ctxt_A_ML1 thus ?case by (fastsimp simp:tRed_list_foldl_At) next case ctxt_A_ML2 thus ?case by (fastsimp simp:tRed_list_foldl_At) next case ctxt_list1 thus ?case by simp next case ctxt_list2 thus ?case by simp qed inductive_cases tRedE: "Ct n -> u" thm tRedE lemma [simp]: "Ct n = foldl At t ts <-> t = Ct n & ts = []" by (induct ts arbitrary:t) auto corollary kernel_inv: includes Vars shows "(t :: tm) =>* t' ==> ML_closed_t 0 t ==> t! ->* t'!" by (rule unproven) theorem includes Vars assumes t: "t : Pure_tms" and t': "t' : Pure_tms" and closed: "free_vars t = {}" and reds: "term_of (eval t []) =>* t'" shows "t ->* t' " proof - have ML_cl: "ML_closed_t 0 (term_of (eval t []))" by (rule unproven) have "(eval t [])! = t!" using kernel_eval[OF t, where e="[]"] closed by simp hence "(term_of (eval t []))! = t!" by simp moreover have "term_of (eval t [])! ->* t'!" using kernel_inv[OF reds ML_cl] by auto ultimately have "t! ->* t'!" by simp thus ?thesis using kernel_pure t t' by auto qed end
lemma
x ∈ set vs ==> size x < Suc (ml_list_size1 vs)
lemma
x ∈ set vs ==> size x < Suc (ml_list_size2 vs)
lemma
x ∈ set vs ==> size x < Suc (size v + ml_list_size3 vs)
lemma
x ∈ set vs ==> size x < Suc (size v + ml_list_size4 vs)
lemma size_lift(1):
size (lift i t) = size t
and size_lift(2):
size (lift i v) = size v
and size_lift(3):
ml_list_size1 (map (lift i) vs) = ml_list_size1 vs
and size_lift(4):
ml_list_size2 (map (lift i) vs) = ml_list_size2 vs
and size_lift(5):
ml_list_size3 (map (lift i) vs) = ml_list_size3 vs
and size_lift(6):
ml_list_size4 (map (lift i) vs) = ml_list_size4 vs
lemma size_lift_ML(1):
size (liftML i t) = size t
and size_lift_ML(2):
size (liftML i v) = size v
and size_lift_ML(3):
ml_list_size1 (map (liftML i) vs) = ml_list_size1 vs
and size_lift_ML(4):
ml_list_size2 (map (liftML i) vs) = ml_list_size2 vs
and size_lift_ML(5):
ml_list_size3 (map (liftML i) vs) = ml_list_size3 vs
and size_lift_ML(6):
ml_list_size4 (map (liftML i) vs) = ml_list_size4 vs
lemma
liftML i (Ct nm) = Ct nm
liftML i (Vt x) = Vt x
liftML i (Lam t) = Lam (liftML i t)
liftML i (At s t) = At (liftML i s) (liftML i t)
liftML i (term_of v) = term_of (liftML i v)
liftML i (C nm vs) = C nm (map (liftML i) vs)
liftML i (V x vs) = V x (map (liftML i) vs)
liftML i (Fun v vs n) = Fun (liftML i v) (map (liftML i) vs) n
liftML i (apply u v) = apply (liftML i u) (liftML i v)
liftML i (V_ML X) = V_ML (if X < i then X else X + 1)
liftML i (A_ML v vs) = A_ML (liftML i v) (map (liftML i) vs)
liftML i (Lam_ML v) = Lam_ML (liftML (i + 1) v)
liftML i (CC nm) = CC nm
lemma
lift i (Ct nm) = Ct nm
lift i (Vt x) = Vt (if x < i then x else x + 1)
lift i (Lam t) = Lam (lift (i + 1) t)
lift i (At s t) = At (lift i s) (lift i t)
lift i (term_of v) = term_of (lift i v)
lift i (C nm vs) = C nm (map (lift i) vs)
lift i (V x vs) = V (if x < i then x else x + 1) (map (lift i) vs)
lift i (Fun v vs n) = Fun (lift i v) (map (lift i) vs) n
lift i (apply u v) = apply (lift i u) (lift i v)
lift i (V_ML X) = V_ML X
lift i (A_ML v vs) = A_ML (lift i v) (map (lift i) vs)
lift i (Lam_ML v) = Lam_ML (lift i v)
lift i (CC nm) = CC nm
lemma
substML f (Ct nm) = Ct nm
substML f (Vt x) = Vt x
substML f (Lam t) = Lam (substML (lift 0 o f) t)
substML f (At s t) = At (substML f s) (substML f t)
substML f (term_of v) = term_of (substML f v)
substML f (C nm vs) = C nm (map (substML f) vs)
substML f (V x vs) = V x (map (substML f) vs)
substML f (Fun v vs n) = Fun (substML f v) (map (substML f) vs) n
substML f (apply u v) = apply (substML f u) (substML f v)
substML f (V_ML X) = f X
substML f (A_ML v vs) = A_ML (substML f v) (map (substML f) vs)
substML f (Lam_ML v) = Lam_ML (substML (V_ML 0 ## f) v)
substML f (CC nm) = CC nm
lemma size_subst_ML(1):
(∀x. size (f x) = 0) --> size (substML f t) = size t
and size_subst_ML(2):
(∀x. size (f x) = 0) --> size (substML f v) = size v
and size_subst_ML(3):
(∀x. size (f x) = 0) --> ml_list_size1 (map (substML f) vs) = ml_list_size1 vs
and size_subst_ML(4):
(∀x. size (f x) = 0) --> ml_list_size2 (map (substML f) vs) = ml_list_size2 vs
and size_subst_ML(5):
(∀x. size (f x) = 0) --> ml_list_size3 (map (substML f) vs) = ml_list_size3 vs
and size_subst_ML(6):
(∀x. size (f x) = 0) --> ml_list_size4 (map (substML f) vs) = ml_list_size4 vs
lemma lift_lift(1):
i < k + 1 ==> lift (Suc k) (lift i t) = lift i (lift k t)
and lift_lift(2):
i < k + 1 ==> lift (Suc k) (lift i v) = lift i (lift k v)
corollary lift_o_lift(1):
i < k + 1 ==> lift (Suc k) o lift i = lift i o lift k
and lift_o_lift(2):
i < k + 1 ==> lift (Suc k) o lift i = lift i o lift k
lemma lift_lift_ML(1):
i < k + 1 ==> liftML (Suc k) (liftML i t) = liftML i (liftML k t)
and lift_lift_ML(2):
i < k + 1 ==> liftML (Suc k) (liftML i v) = liftML i (liftML k v)
lemma lift_lift_ML_comm(1):
lift j (liftML i t) = liftML i (lift j t)
and lift_lift_ML_comm(2):
lift j (liftML i v) = liftML i (lift j v)
lemma
V_ML 0 ## subst_decr_ML k v = subst_decr_ML (Suc k) (liftML 0 v)
lemma
lift 0 o subst_decr_ML k v = subst_decr_ML k (lift 0 v)
lemma subst_lift_id(1):
substML (subst_decr_ML k v) (liftML k t) = t
and subst_lift_id(2):
liftML k u[v/k] = u
lemma tRed_list_refl:
ts ->* ts
lemma listsum_size':
v ∈ set vs ==> size' v < Suc (listsum (map size' vs))
corollary cor_listsum_size':
v ∈ set vs ==> size' v < Suc (m + listsum (map size' vs))
lemma size_subst_ML(1):
∀i. size (f i) = 0 ==> size (substML f t) = size t
and size_subst_ML(2):
∀i. size (f i) = 0 ==> size (substML f v) = size v
and size_subst_ML(3):
∀i. size (f i) = 0 ==> ml_list_size1 (map (substML f) vs) = ml_list_size1 vs
and size_subst_ML(4):
∀i. size (f i) = 0 ==> ml_list_size2 (map (substML f) vs) = ml_list_size2 vs
and size_subst_ML(5):
∀i. size (f i) = 0 ==> ml_list_size3 (map (substML f) vs) = ml_list_size3 vs
and size_subst_ML(6):
∀i. size (f i) = 0 ==> ml_list_size4 (map (substML f) vs) = ml_list_size4 vs
lemma
∀i j. size' (f i) = size' (V_ML j) ==> size' (substML f v) = size' v
lemma
size' (lift i v) = size' v
lemma
t ∈ Pure_tms ==> liftML k t = t
lemma kernel_pure:
t ∈ Pure_tms ==> t! = t
lemma lift_eval:
t ∈ Pure_tms
==> ∀e k. (∀i∈free_vars t. i < length e) -->
lift k (eval t e) = eval t (map (lift k) e)
lemma lift_ML_eval:
[| t ∈ Pure_tms; !!i. i ∈ free_vars t ==> i < length e |]
==> liftML k (eval t e) = eval t (map (liftML k) e)
lemma
(v ## f) 0 = v
lemma
(v ## f) (Suc n) = liftML 0 (f n)
lemma lift_o_shift:
lift k o V_ML 0 ## f = V_ML 0 ## (lift k o f)
lemma lift_subst_ML(1):
lift k (substML f t) = substML (lift k o f) (lift k t)
and lift_subst_ML(2):
lift k (substML f v) = substML (lift k o f) (lift k v)
corollary lift_subst_ML1:
∀v k. lift 0 (u[v/k]) = lift 0 u[lift 0 v/k]
lemma lift_ML_lift_ML(1):
i < k + 1 ==> liftML (Suc k) (liftML i t) = liftML i (liftML k t)
and lift_ML_lift_ML(2):
i < k + 1 ==> liftML (Suc k) (liftML i v) = liftML i (liftML k v)
corollary lift_ML_o_lift_ML(1):
i < k + 1 ==> liftML (Suc k) o liftML i = liftML i o liftML k
and lift_ML_o_lift_ML(2):
i < k + 1 ==> liftML (Suc k) o liftML i = liftML i o liftML k
lemma subst_insrt_lift(1):
substML (insrt k f) (liftML k t) = liftML k (substML f t)
and subst_insrt_lift(2):
substML (insrt k f) (liftML k v) = liftML k (substML f v)
corollary subst_cons_lift:
substML (V_ML 0 ## f) o liftML 0 = liftML 0 o substML f
lemma subst_eval:
[| t ∈ Pure_tms; !!i. i ∈ free_vars t ==> i < length e |]
==> substML f (eval t e) = eval t (map (substML f) e)
theorem kernel_eval:
t ∈ Pure_tms
==> ∀e. (∀i∈free_vars t. i < length e) -->
(∀i<length e. e ! i = V i []) --> eval t e! = t!
lemma map_eq_iff_nth:
(map f xs = map g xs) = (∀i<length xs. f (xs ! i) = g (xs ! i))
lemma
ML_closed k v ==> liftML k v = v
lemma
ML_closed 0 v ==> substML f v = v
lemma
ML_closed k v ==> ML_closed k (lift m v)
lemma red_Lam:
t ->* t' ==> Lam t ->* Lam t'
lemma red_At1:
t ->* t' ==> At t s ->* At t' s
lemma red_At2:
t ->* t' ==> At s t ->* At s t'
lemma tRed_list_foldl_At:
[| ts ->* ts'; s ->* s' |] ==> foldl At s ts ->* foldl At s' ts'
lemma
[| s = t; t -> t' |] ==> s -> t'
lemma subst_foldl:
subst f (foldl At s ts) = foldl At (subst f s) (map (subst f) ts)
lemma foldl_At_size:
length ts = length ts'
==> (foldl At s ts = foldl At s' ts') = (s = s' ∧ ts = ts')
lemma depth_At_foldl:
depth_At (foldl At s ts) = depth_At s + length ts
lemma foldl_At_eq_length:
foldl At s ts = foldl At s ts' ==> length ts = length ts'
lemma foldl_At_eq:
(foldl At s ts = foldl At s ts') = (ts = ts')
lemma
foldl At s ts! = foldl At (s!) (map kernelt ts)
lemma
kernelt o term_of = kernel
lemma shift_subst_decr:
Vt 0 ## subst_decr k t = subst_decr (Suc k) (lift 0 t)
lemma
lift k (foldl At s ts) = foldl At (lift k s) (map (lift k) ts)
lemma
liftn n (C i vs) = C i (map (liftn n) vs)
lemma
liftn n (CC nm) = CC nm
lemma
liftn n (apply v w) = apply (liftn n v) (liftn n w)
lemma
liftn n (A_ML v vs) = A_ML (liftn n v) (map (liftn n) vs)
lemma
liftn n (Fun v vs i) = Fun (liftn n v) (map (liftn n) vs) i
lemma
liftn n (Lam_ML v) = Lam_ML (liftn n v)
lemma liftn_liftn_add:
liftn m (liftn n v) = liftn (m + n) v
lemma
liftn n (V_ML k) = V_ML k
lemma liftn_lift_ML_comm:
liftn n (liftML 0 v) = liftML 0 (liftn n v)
lemma liftn_cons:
liftn n ((V_ML 0 ## f) x) = (V_ML 0 ## (liftn n o f)) x
lemma kernel_subst1:
[| ML_closed 1 u; ML_closed 0 v |] ==> u[v/0]! = lift 0 u[V 0 []/0]![v!/0]
lemma foldl_Pure:
[| t ∈ Pure_tms; ∀t∈set ts. t ∈ Pure_tms;
!!s t. [| s ∈ Pure_tms; t ∈ Pure_tms |] ==> f s t ∈ Pure_tms |]
==> foldl f t ts ∈ Pure_tms
lemma
ML_closed 0 v ==> v! ∈ Pure_tms
lemma subst_Vt:
subst Vt = id
theorem Red_sound(1):
[| v => v'; ML_closed 0 v |] ==> v! ->* v'! ∧ ML_closed 0 v'
and Red_sound(2):
[| t => t'; ML_closed_t 0 t |] ==> t! ->* t'! ∧ ML_closed_t 0 t'
and Red_sound(3):
[| vs => vs'; ∀v∈set vs. ML_closed 0 v |]
==> vs! ->* vs'! ∧ (∀v'∈set vs'. ML_closed 0 v')
lemma
(Ct n = foldl At t ts) = (t = Ct n ∧ ts = [])
corollary kernel_inv:
[| t =>* t'; ML_closed_t 0 t |] ==> t! ->* t'!
theorem
[| t ∈ Pure_tms; t' ∈ Pure_tms; free_vars t = {}; term_of (eval t []) =>* t' |]
==> t ->* t'