(* "$Id: Crary.thy,v 1.24 2007/10/19 21:21:10 wenzelm Exp $" *) (* *) (* Formalisation of the chapter on Logical Relations *) (* and a Case Study in Equivalence Checking *) (* by Karl Crary from the book on Advanced Topics in *) (* Types and Programming Languages, MIT Press 2005 *) (* The formalisation was done by Julien Narboux and *) (* Christian Urban. *) theory Crary imports "../Nominal" begin atom_decl name nominal_datatype ty = TBase | TUnit | Arrow "ty" "ty" ("_->_" [100,100] 100) nominal_datatype trm = Unit | Var "name" ("Var _" [100] 100) | Lam "«name»trm" ("Lam [_]._" [100,100] 100) | App "trm" "trm" ("App _ _" [110,110] 100) | Const "nat" types Ctxt = "(name×ty) list" types Subst = "(name×trm) list" lemma perm_ty[simp]: fixes T::"ty" and pi::"name prm" shows "pi•T = T" by (induct T rule: ty.weak_induct) (simp_all) lemma fresh_ty[simp]: fixes x::"name" and T::"ty" shows "x\<sharp>T" by (simp add: fresh_def supp_def) lemma ty_cases: fixes T::ty shows "(∃ T1 T2. T=T1->T2) ∨ T=TUnit ∨ T=TBase" by (induct T rule:ty.weak_induct) (auto) instance ty :: size .. nominal_primrec "size (TBase) = 1" "size (TUnit) = 1" "size (T1->T2) = size T1 + size T2" by (rule TrueI)+ lemma ty_size_greater_zero[simp]: fixes T::"ty" shows "size T > 0" by (nominal_induct rule:ty.induct) (simp_all) section {* Substitutions *} fun lookup :: "Subst => name => trm" where "lookup [] x = Var x" | "lookup ((y,T)#ϑ) x = (if x=y then T else lookup ϑ x)" lemma lookup_eqvt[eqvt]: fixes pi::"name prm" shows "pi•(lookup ϑ x) = lookup (pi•ϑ) (pi•x)" by (induct ϑ) (auto simp add: perm_bij) lemma lookup_fresh: fixes z::"name" assumes a: "z\<sharp>ϑ" "z\<sharp>x" shows "z\<sharp> lookup ϑ x" using a by (induct rule: lookup.induct) (auto simp add: fresh_list_cons) lemma lookup_fresh': assumes a: "z\<sharp>ϑ" shows "lookup ϑ z = Var z" using a by (induct rule: lookup.induct) (auto simp add: fresh_list_cons fresh_prod fresh_atm) consts psubst :: "Subst => trm => trm" ("_<_>" [100,100] 130) nominal_primrec "ϑ<(Var x)> = (lookup ϑ x)" "ϑ<(App t1 t2)> = App ϑ<t1> ϑ<t2>" "x\<sharp>ϑ ==> ϑ<(Lam [x].t)> = Lam [x].(ϑ<t>)" "ϑ<(Const n)> = Const n" "ϑ<(Unit)> = Unit" apply(finite_guess)+ apply(rule TrueI)+ apply(simp add: abs_fresh)+ apply(fresh_guess)+ done abbreviation subst :: "trm => name => trm => trm" ("_[_::=_]" [100,100,100] 100) where "t[x::=t'] ≡ ([(x,t')])<t>" lemma subst[simp]: 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\<sharp>(y,t') ==> (Lam [x].t)[y::=t'] = Lam [x].(t[y::=t'])" and "Const n[y::=t'] = Const n" and "Unit [y::=t'] = Unit" by (simp_all add: fresh_list_cons fresh_list_nil) lemma subst_eqvt[eqvt]: fixes pi::"name prm" shows "pi•(t[x::=t']) = (pi•t)[(pi•x)::=(pi•t')]" by (nominal_induct t avoiding: x t' rule: trm.induct) (perm_simp add: fresh_bij)+ lemma subst_rename: fixes c::"name" assumes a: "c\<sharp>t1" shows "t1[a::=t2] = ([(c,a)]•t1)[c::=t2]" using a apply(nominal_induct t1 avoiding: a c t2 rule: trm.induct) apply(simp add: trm.inject calc_atm fresh_atm abs_fresh perm_nat_def)+ done lemma fresh_psubst: fixes z::"name" assumes a: "z\<sharp>t" "z\<sharp>ϑ" shows "z\<sharp>(ϑ<t>)" using a by (nominal_induct t avoiding: z ϑ t rule: trm.induct) (auto simp add: abs_fresh lookup_fresh) lemma fresh_subst'': fixes z::"name" assumes "z\<sharp>t2" shows "z\<sharp>t1[z::=t2]" using assms by (nominal_induct t1 avoiding: t2 z rule: trm.induct) (auto simp add: abs_fresh fresh_nat fresh_atm) lemma fresh_subst': fixes z::"name" assumes "z\<sharp>[y].t1" "z\<sharp>t2" shows "z\<sharp>t1[y::=t2]" using assms by (nominal_induct t1 avoiding: y t2 z rule: trm.induct) (auto simp add: abs_fresh fresh_nat fresh_atm) lemma fresh_subst: fixes z::"name" assumes a: "z\<sharp>t1" "z\<sharp>t2" shows "z\<sharp>t1[y::=t2]" using a by (auto simp add: fresh_subst' abs_fresh) lemma fresh_psubst_simp: assumes "x\<sharp>t" shows "((x,u)#ϑ)<t> = ϑ<t>" using assms proof (nominal_induct t avoiding: x u ϑ rule: trm.induct) case (Lam y t x u) have fs: "y\<sharp>ϑ" "y\<sharp>x" "y\<sharp>u" by fact+ moreover have "x\<sharp> Lam [y].t" by fact ultimately have "x\<sharp>t" by (simp add: abs_fresh fresh_atm) moreover have ih:"!!n T. n\<sharp>t ==> ((n,T)#ϑ)<t> = ϑ<t>" by fact ultimately have "((x,u)#ϑ)<t> = ϑ<t>" by auto moreover have "((x,u)#ϑ)<Lam [y].t> = Lam [y].(((x,u)#ϑ)<t>)" using fs by (simp add: fresh_list_cons fresh_prod) moreover have " ϑ<Lam [y].t> = Lam [y]. (ϑ<t>)" using fs by simp ultimately show "((x,u)#ϑ)<Lam [y].t> = ϑ<Lam [y].t>" by auto qed (auto simp add: fresh_atm abs_fresh) lemma forget: fixes x::"name" assumes a: "x\<sharp>t" shows "t[x::=t'] = t" using a by (nominal_induct t avoiding: x t' rule: trm.induct) (auto simp add: fresh_atm abs_fresh) lemma subst_fun_eq: fixes u::trm assumes h:"[x].t1 = [y].t2" shows "t1[x::=u] = t2[y::=u]" proof - { assume "x=y" and "t1=t2" then have ?thesis using h by simp } moreover { assume h1:"x ≠ y" and h2:"t1=[(x,y)] • t2" and h3:"x \<sharp> t2" then have "([(x,y)] • t2)[x::=u] = t2[y::=u]" by (simp add: subst_rename) then have ?thesis using h2 by simp } ultimately show ?thesis using alpha h by blast qed lemma psubst_empty[simp]: shows "[]<t> = t" by (nominal_induct t rule: trm.induct) (auto simp add: fresh_list_nil) lemma psubst_subst_psubst: assumes h:"c\<sharp>ϑ" shows "ϑ<t>[c::=s] = ((c,s)#ϑ)<t>" using h by (nominal_induct t avoiding: ϑ c s rule: trm.induct) (auto simp add: fresh_list_cons fresh_atm forget lookup_fresh lookup_fresh' fresh_psubst) lemma subst_fresh_simp: assumes a: "x\<sharp>ϑ" shows "ϑ<Var x> = Var x" using a by (induct ϑ arbitrary: x, auto simp add:fresh_list_cons fresh_prod fresh_atm) lemma psubst_subst_propagate: assumes "x\<sharp>ϑ" shows "ϑ<t[x::=u]> = ϑ<t>[x::=ϑ<u>]" using assms proof (nominal_induct t avoiding: x u ϑ rule: trm.induct) case (Var n x u ϑ) { assume "x=n" moreover have "x\<sharp>ϑ" by fact ultimately have "ϑ<Var n[x::=u]> = ϑ<Var n>[x::=ϑ<u>]" using subst_fresh_simp by auto } moreover { assume h:"x≠n" then have "x\<sharp>Var n" by (auto simp add: fresh_atm) moreover have "x\<sharp>ϑ" by fact ultimately have "x\<sharp>ϑ<Var n>" using fresh_psubst by blast then have " ϑ<Var n>[x::=ϑ<u>] = ϑ<Var n>" using forget by auto then have "ϑ<Var n[x::=u]> = ϑ<Var n>[x::=ϑ<u>]" using h by auto } ultimately show ?case by auto next case (Lam n t x u ϑ) have fs:"n\<sharp>x" "n\<sharp>u" "n\<sharp>ϑ" "x\<sharp>ϑ" by fact+ have ih:"!! y s ϑ. y\<sharp>ϑ ==> ((ϑ<(t[y::=s])>) = ((ϑ<t>)[y::=(ϑ<s>)]))" by fact have "ϑ <(Lam [n].t)[x::=u]> = ϑ<Lam [n]. (t [x::=u])>" using fs by auto then have "ϑ <(Lam [n].t)[x::=u]> = Lam [n]. ϑ<t [x::=u]>" using fs by auto moreover have "ϑ<t[x::=u]> = ϑ<t>[x::=ϑ<u>]" using ih fs by blast ultimately have "ϑ <(Lam [n].t)[x::=u]> = Lam [n].(ϑ<t>[x::=ϑ<u>])" by auto moreover have "Lam [n].(ϑ<t>[x::=ϑ<u>]) = (Lam [n].ϑ<t>)[x::=ϑ<u>]" using fs fresh_psubst by auto ultimately have "ϑ<(Lam [n].t)[x::=u]> = (Lam [n].ϑ<t>)[x::=ϑ<u>]" using fs by auto then show "ϑ<(Lam [n].t)[x::=u]> = ϑ<Lam [n].t>[x::=ϑ<u>]" using fs by auto qed (auto) section {* Typing *} inductive valid :: "Ctxt => bool" where v_nil[intro]: "valid []" | v_cons[intro]: "[|valid Γ;a\<sharp>Γ|] ==> valid ((a,T)#Γ)" equivariance valid inductive_cases valid_cons_elim_auto[elim]:"valid ((x,T)#Γ)" abbreviation "sub_context" :: "Ctxt => Ctxt => bool" (" _ ⊆ _ " [55,55] 55) where "Γ1 ⊆ Γ2 ≡ ∀a T. (a,T)∈set Γ1 --> (a,T)∈set Γ2" lemma valid_monotonicity[elim]: fixes Γ Γ' :: Ctxt assumes a: "Γ ⊆ Γ'" and b: "x\<sharp>Γ'" shows "(x,T1)#Γ ⊆ (x,T1)#Γ'" using a b by auto lemma fresh_context: fixes Γ :: "Ctxt" and a :: "name" assumes "a\<sharp>Γ" shows "¬(∃τ::ty. (a,τ)∈set Γ)" using assms by (induct Γ) (auto simp add: fresh_prod fresh_list_cons fresh_atm) lemma type_unicity_in_context: assumes a: "valid Γ" and b: "(x,T1) ∈ set Γ" and c: "(x,T2) ∈ set Γ" shows "T1=T2" using a b c by (induct Γ) (auto dest!: fresh_context) inductive typing :: "Ctxt=>trm=>ty=>bool" (" _ \<turnstile> _ : _ " [60,60,60] 60) where T_Var[intro]: "[|valid Γ; (x,T)∈set Γ|] ==> Γ \<turnstile> Var x : T" | T_App[intro]: "[|Γ \<turnstile> e1 : T1->T2; Γ \<turnstile> e2 : T1|] ==> Γ \<turnstile> App e1 e2 : T2" | T_Lam[intro]: "[|x\<sharp>Γ; (x,T1)#Γ \<turnstile> t : T2|] ==> Γ \<turnstile> Lam [x].t : T1->T2" | T_Const[intro]: "valid Γ ==> Γ \<turnstile> Const n : TBase" | T_Unit[intro]: "valid Γ ==> Γ \<turnstile> Unit : TUnit" equivariance typing nominal_inductive typing by (simp_all add: abs_fresh) lemma typing_implies_valid: assumes a: "Γ \<turnstile> t : T" shows "valid Γ" using a by (induct) (auto) declare trm.inject [simp add] declare ty.inject [simp add] inductive_cases typing_inv_auto[elim]: "Γ \<turnstile> Lam [x].t : T" "Γ \<turnstile> Var x : T" "Γ \<turnstile> App x y : T" "Γ \<turnstile> Const n : T" "Γ \<turnstile> Unit : TUnit" "Γ \<turnstile> s : TUnit" declare trm.inject [simp del] declare ty.inject [simp del] section {* Definitional Equivalence *} inductive def_equiv :: "Ctxt=>trm=>trm=>ty=>bool" ("_ \<turnstile> _ ≡ _ : _" [60,60] 60) where Q_Refl[intro]: "Γ \<turnstile> t : T ==> Γ \<turnstile> t ≡ t : T" | Q_Symm[intro]: "Γ \<turnstile> t ≡ s : T ==> Γ \<turnstile> s ≡ t : T" | Q_Trans[intro]: "[|Γ \<turnstile> s ≡ t : T; Γ \<turnstile> t ≡ u : T|] ==> Γ \<turnstile> s ≡ u : T" | Q_Abs[intro]: "[|x\<sharp>Γ; (x,T1)#Γ \<turnstile> s2 ≡ t2 : T2|] ==> Γ \<turnstile> Lam [x]. s2 ≡ Lam [x]. t2 : T1 -> T2" | Q_App[intro]: "[|Γ \<turnstile> s1 ≡ t1 : T1 -> T2 ; Γ \<turnstile> s2 ≡ t2 : T1|] ==> Γ \<turnstile> App s1 s2 ≡ App t1 t2 : T2" | Q_Beta[intro]: "[|x\<sharp>(Γ,s2,t2); (x,T1)#Γ \<turnstile> s1 ≡ t1 : T2 ; Γ \<turnstile> s2 ≡ t2 : T1|] ==> Γ \<turnstile> App (Lam [x]. s1) s2 ≡ t1[x::=t2] : T2" | Q_Ext[intro]: "[|x\<sharp>(Γ,s,t); (x,T1)#Γ \<turnstile> App s (Var x) ≡ App t (Var x) : T2|] ==> Γ \<turnstile> s ≡ t : T1 -> T2" | Q_Unit[intro]: "[|Γ \<turnstile> s : TUnit; Γ \<turnstile> t: TUnit|] ==> Γ \<turnstile> s ≡ t : TUnit" equivariance def_equiv nominal_inductive def_equiv by (simp_all add: abs_fresh fresh_subst'') lemma def_equiv_implies_valid: assumes a: "Γ \<turnstile> t ≡ s : T" shows "valid Γ" using a by (induct) (auto elim: typing_implies_valid) lemma test30: fixes x::"name" assumes a: "(x,T) ∈ set Γ" shows "x∈supp Γ" using a apply(induct Γ) apply(auto simp add: supp_list_cons supp_prod supp_atm) done lemma supp_ty[simp]: fixes T::"ty" shows "(supp T) = ({}::name set)" apply(simp add: supp_def) done lemma test3a: fixes Γ :: Ctxt and t :: trm assumes a: "Γ \<turnstile> t : T" shows "(supp t) ⊆ ((supp Γ)::name set)" using a apply(induct) apply(auto simp add: trm.supp supp_atm supp_list_cons abs_supp supp_prod) apply(simp add: test30) apply(simp add: supp_def perm_nat_def) done lemma test3b: shows "supp (t1[x::=t2]) ⊆ ((supp ([x].t1,t2))::name set)" apply(nominal_induct t1 avoiding: x t2 rule: trm.induct) apply(auto simp add: trm.supp supp_prod abs_supp supp_atm) apply(blast) apply(blast) apply(simp add: supp_def perm_nat_def) done lemma test3: fixes Γ :: Ctxt and s t :: trm assumes a: "Γ \<turnstile> s ≡ t : T" shows "(supp (s,t)) ⊆ ((supp Γ)::name set)" using a apply(induct) apply(auto simp add: supp_prod supp_list_cons trm.supp abs_supp supp_atm) apply(drule test3a) apply(blast) apply(subgoal_tac "supp (t1[x::=t2]) ⊆ ((supp ([x].t1,t2))::name set)") apply(auto simp add: supp_prod abs_supp)[1] apply(rule test3b) apply(case_tac "x=xa") apply(simp add: fresh_def supp_prod) apply(blast) apply(case_tac "x=xa") apply(simp add: fresh_def supp_prod) apply(blast) apply(drule test3a) apply(blast) apply(drule test3a)+ apply(blast) done lemma test0: assumes a: "(x,T1)#Γ \<turnstile> s1 ≡ t1 : T2" and b: "Γ \<turnstile> s2 ≡ t2 : T1" shows "Γ \<turnstile> App (Lam [x]. s1) s2 ≡ t1[x::=t2] : T2" using a b apply(rule_tac Q_Beta) apply(auto) apply(drule def_equiv_implies_valid) apply(drule valid.cases) apply(auto) apply(drule test3) apply(auto simp add: fresh_def supp_prod) done lemma test1: assumes a: "∀x. x\<sharp>Γ --> (x,T1)#Γ \<turnstile> App s (Var x) ≡ App t (Var x) : T2" shows "Γ \<turnstile> s ≡ t : T1 -> T2" using a apply - apply(generate_fresh "name") apply(rule_tac x="c" in Q_Ext) apply(simp) apply(simp add: fresh_prod) done lemma test2: assumes a: "x\<sharp>(Γ,s,t) ∧ (x,T1)#Γ \<turnstile> App s (Var x) ≡ App t (Var x) : T2" shows "∀x. x\<sharp>(Γ,s,t) --> (x,T1)#Γ \<turnstile> App s (Var x) ≡ App t (Var x) : T2" using a apply - apply(rule allI) apply(case_tac "xa=x") apply(simp) apply(rule impI) apply(erule conjE) apply(drule_tac pi="[(x,xa)]" in def_equiv.eqvt) apply(simp add: eqvts) apply(simp add: calc_atm) apply(perm_simp) done lemma test2: assumes a: "x\<sharp>(Γ,s,t) ∧ (x,T1)#Γ \<turnstile> App s (Var x) ≡ App t (Var x) : T2" shows "∀x. x\<sharp>Γ --> (x,T1)#Γ \<turnstile> App s (Var x) ≡ App t (Var x) : T2" using a apply - apply(rule allI) apply(case_tac "xa=x") apply(simp) apply(rule impI) apply(erule conjE) apply(frule test3) apply(simp add: supp_prod supp_list_cons supp_atm trm.supp) apply(subgoal_tac "xa\<sharp>(s,t)") apply(drule_tac pi="[(x,xa)]" in def_equiv.eqvt) apply(simp add: eqvts) apply(simp add: calc_atm) apply(perm_simp) apply(simp add: fresh_def supp_prod) apply(auto) done section {* Weak Head Reduction *} inductive whr_def :: "trm=>trm=>bool" ("_ \<leadsto> _" [80,80] 80) where QAR_Beta[intro]: "App (Lam [x]. t1) t2 \<leadsto> t1[x::=t2]" | QAR_App[intro]: "t1 \<leadsto> t1' ==> App t1 t2 \<leadsto> App t1' t2" declare trm.inject [simp add] declare ty.inject [simp add] inductive_cases whr_inv_auto[elim]: "t \<leadsto> t'" "Lam [x].t \<leadsto> t'" "App (Lam [x].t12) t2 \<leadsto> t" "Var x \<leadsto> t" "Const n \<leadsto> t" "App p q \<leadsto> t" "t \<leadsto> Const n" "t \<leadsto> Var x" "t \<leadsto> App p q" declare trm.inject [simp del] declare ty.inject [simp del] equivariance whr_def section {* Weak Head Normalisation *} abbreviation nf :: "trm => bool" ("_ \<leadsto>|" [100] 100) where "t\<leadsto>| ≡ ¬(∃ u. t \<leadsto> u)" inductive whn_def :: "trm=>trm=>bool" ("_ \<Down> _" [80,80] 80) where QAN_Reduce[intro]: "[|s \<leadsto> t; t \<Down> u|] ==> s \<Down> u" | QAN_Normal[intro]: "t\<leadsto>| ==> t \<Down> t" declare trm.inject[simp] inductive_cases whn_inv_auto[elim]: "t \<Down> t'" declare trm.inject[simp del] lemma whn_eqvt[eqvt]: fixes pi::"name prm" assumes a: "t \<Down> t'" shows "(pi•t) \<Down> (pi•t')" using a apply(induct) apply(rule QAN_Reduce) apply(rule whr_def.eqvt) apply(assumption)+ apply(rule QAN_Normal) apply(auto) apply(drule_tac pi="rev pi" in whr_def.eqvt) apply(perm_simp) done lemma red_unicity : assumes a: "x \<leadsto> a" and b: "x \<leadsto> b" shows "a=b" using a b apply (induct arbitrary: b) apply (erule whr_inv_auto(3)) apply (clarify) apply (rule subst_fun_eq) apply (simp) apply (force) apply (erule whr_inv_auto(6)) apply (blast)+ done lemma nf_unicity : assumes "x \<Down> a" and "x \<Down> b" shows "a=b" using assms proof (induct arbitrary: b) case (QAN_Reduce x t a b) have h:"x \<leadsto> t" "t \<Down> a" by fact+ have ih:"!!b. t \<Down> b ==> a = b" by fact have "x \<Down> b" by fact then obtain t' where "x \<leadsto> t'" and hl:"t' \<Down> b" using h by auto then have "t=t'" using h red_unicity by auto then show "a=b" using ih hl by auto qed (auto) lemma test4a: shows "s \<leadsto> t ==> (supp t) ⊆ ((supp s)::name set)" apply(induct s t rule: whr_def.induct) apply(insert test3b) apply(simp add: trm.supp supp_prod abs_supp) apply(auto simp add: trm.supp) done lemma test4b: shows "s \<Down> t ==> (supp t) ⊆ ((supp s)::name set)" apply(induct s t rule: whn_def.induct) apply(auto dest: test4a) done section {* Algorithmic Term Equivalence and Algorithmic Path Equivalence *} inductive alg_equiv :: "Ctxt=>trm=>trm=>ty=>bool" ("_ \<turnstile> _ \<Leftrightarrow> _ : _" [60,60,60,60] 60) and alg_path_equiv :: "Ctxt=>trm=>trm=>ty=>bool" ("_ \<turnstile> _ \<leftrightarrow> _ : _" [60,60,60,60] 60) where QAT_Base[intro]: "[|s \<Down> p; t \<Down> q; Γ \<turnstile> p \<leftrightarrow> q : TBase|] ==> Γ \<turnstile> s \<Leftrightarrow> t : TBase" | QAT_Arrow[intro]: "[|x\<sharp>(Γ,s,t); (x,T1)#Γ \<turnstile> App s (Var x) \<Leftrightarrow> App t (Var x) : T2|] ==> Γ \<turnstile> s \<Leftrightarrow> t : T1 -> T2" | QAT_One[intro]: "valid Γ ==> Γ \<turnstile> s \<Leftrightarrow> t : TUnit" | QAP_Var[intro]: "[|valid Γ;(x,T) ∈ set Γ|] ==> Γ \<turnstile> Var x \<leftrightarrow> Var x : T" | QAP_App[intro]: "[|Γ \<turnstile> p \<leftrightarrow> q : T1 -> T2; Γ \<turnstile> s \<Leftrightarrow> t : T1|] ==> Γ \<turnstile> App p s \<leftrightarrow> App q t : T2" | QAP_Const[intro]: "valid Γ ==> Γ \<turnstile> Const n \<leftrightarrow> Const n : TBase" equivariance alg_equiv nominal_inductive alg_equiv avoids QAT_Arrow: x by simp_all declare trm.inject [simp add] declare ty.inject [simp add] inductive_cases alg_equiv_inv_auto[elim]: "Γ \<turnstile> s\<Leftrightarrow>t : TBase" "Γ \<turnstile> s\<Leftrightarrow>t : T1 -> T2" "Γ \<turnstile> s\<leftrightarrow>t : TBase" "Γ \<turnstile> s\<leftrightarrow>t : TUnit" "Γ \<turnstile> s\<leftrightarrow>t : T1 -> T2" "Γ \<turnstile> Var x \<leftrightarrow> t : T" "Γ \<turnstile> Var x \<leftrightarrow> t : T'" "Γ \<turnstile> s \<leftrightarrow> Var x : T" "Γ \<turnstile> s \<leftrightarrow> Var x : T'" "Γ \<turnstile> Const n \<leftrightarrow> t : T" "Γ \<turnstile> s \<leftrightarrow> Const n : T" "Γ \<turnstile> App p s \<leftrightarrow> t : T" "Γ \<turnstile> s \<leftrightarrow> App q t : T" "Γ \<turnstile> Lam[x].s \<leftrightarrow> t : T" "Γ \<turnstile> t \<leftrightarrow> Lam[x].s : T" declare trm.inject [simp del] declare ty.inject [simp del] lemma Q_Arrow_strong_inversion: assumes fs: "x\<sharp>Γ" "x\<sharp>t" "x\<sharp>u" and h: "Γ \<turnstile> t \<Leftrightarrow> u : T1->T2" shows "(x,T1)#Γ \<turnstile> App t (Var x) \<Leftrightarrow> App u (Var x) : T2" proof - obtain y where fs2: "y\<sharp>(Γ,t,u)" and "(y,T1)#Γ \<turnstile> App t (Var y) \<Leftrightarrow> App u (Var y) : T2" using h by auto then have "([(x,y)]•((y,T1)#Γ)) \<turnstile> [(x,y)]• App t (Var y) \<Leftrightarrow> [(x,y)]• App u (Var y) : T2" using alg_equiv.eqvt[simplified] by blast then show ?thesis using fs fs2 by (perm_simp) qed (* Warning this lemma is false: lemma algorithmic_type_unicity: shows "[| Γ \<turnstile> s \<Leftrightarrow> t : T ; Γ \<turnstile> s \<Leftrightarrow> u : T' |] ==> T = T'" and "[| Γ \<turnstile> s \<leftrightarrow> t : T ; Γ \<turnstile> s \<leftrightarrow> u : T' |] ==> T = T'" Here is the counter example : Γ \<turnstile> Const n \<Leftrightarrow> Const n : Tbase and Γ \<turnstile> Const n \<Leftrightarrow> Const n : TUnit *) lemma algorithmic_path_type_unicity: shows "Γ \<turnstile> s \<leftrightarrow> t : T ==> Γ \<turnstile> s \<leftrightarrow> u : T' ==> T = T'" proof (induct arbitrary: u T' rule: alg_equiv_alg_path_equiv.inducts(2) [of _ _ _ _ _ "%a b c d . True" ]) case (QAP_Var Γ x T u T') have "Γ \<turnstile> Var x \<leftrightarrow> u : T'" by fact then have "u=Var x" and "(x,T') ∈ set Γ" by auto moreover have "valid Γ" "(x,T) ∈ set Γ" by fact+ ultimately show "T=T'" using type_unicity_in_context by auto next case (QAP_App Γ p q T1 T2 s t u T2') have ih:"!!u T. Γ \<turnstile> p \<leftrightarrow> u : T ==> T1->T2 = T" by fact have "Γ \<turnstile> App p s \<leftrightarrow> u : T2'" by fact then obtain r t T1' where "u = App r t" "Γ \<turnstile> p \<leftrightarrow> r : T1' -> T2'" by auto with ih have "T1->T2 = T1' -> T2'" by auto then show "T2=T2'" using ty.inject by auto qed (auto) lemma alg_path_equiv_implies_valid: shows "Γ \<turnstile> s \<Leftrightarrow> t : T ==> valid Γ" and "Γ \<turnstile> s \<leftrightarrow> t : T ==> valid Γ" by (induct rule : alg_equiv_alg_path_equiv.inducts, auto) lemma algorithmic_symmetry: shows "Γ \<turnstile> s \<Leftrightarrow> t : T ==> Γ \<turnstile> t \<Leftrightarrow> s : T" and "Γ \<turnstile> s \<leftrightarrow> t : T ==> Γ \<turnstile> t \<leftrightarrow> s : T" by (induct rule: alg_equiv_alg_path_equiv.inducts) (auto simp add: fresh_prod) lemma algorithmic_transitivity: shows "Γ \<turnstile> s \<Leftrightarrow> t : T ==> Γ \<turnstile> t \<Leftrightarrow> u : T ==> Γ \<turnstile> s \<Leftrightarrow> u : T" and "Γ \<turnstile> s \<leftrightarrow> t : T ==> Γ \<turnstile> t \<leftrightarrow> u : T ==> Γ \<turnstile> s \<leftrightarrow> u : T" proof (nominal_induct Γ s t T and Γ s t T avoiding: u rule: alg_equiv_alg_path_equiv.strong_inducts) case (QAT_Base s p t q Γ u) have "Γ \<turnstile> t \<Leftrightarrow> u : TBase" by fact then obtain r' q' where b1: "t \<Down> q'" and b2: "u \<Down> r'" and b3: "Γ \<turnstile> q' \<leftrightarrow> r' : TBase" by auto have ih: "Γ \<turnstile> q \<leftrightarrow> r' : TBase ==> Γ \<turnstile> p \<leftrightarrow> r' : TBase" by fact have "t \<Down> q" by fact with b1 have eq: "q=q'" by (simp add: nf_unicity) with ih b3 have "Γ \<turnstile> p \<leftrightarrow> r' : TBase" by simp moreover have "s \<Down> p" by fact ultimately show "Γ \<turnstile> s \<Leftrightarrow> u : TBase" using b2 by auto next case (QAT_Arrow x Γ s t T1 T2 u) have ih:"(x,T1)#Γ \<turnstile> App t (Var x) \<Leftrightarrow> App u (Var x) : T2 ==> (x,T1)#Γ \<turnstile> App s (Var x) \<Leftrightarrow> App u (Var x) : T2" by fact have fs: "x\<sharp>Γ" "x\<sharp>s" "x\<sharp>t" "x\<sharp>u" by fact+ have "Γ \<turnstile> t \<Leftrightarrow> u : T1->T2" by fact then have "(x,T1)#Γ \<turnstile> App t (Var x) \<Leftrightarrow> App u (Var x) : T2" using fs by (simp add: Q_Arrow_strong_inversion) with ih have "(x,T1)#Γ \<turnstile> App s (Var x) \<Leftrightarrow> App u (Var x) : T2" by simp then show "Γ \<turnstile> s \<Leftrightarrow> u : T1->T2" using fs by (auto simp add: fresh_prod) next case (QAP_App Γ p q T1 T2 s t u) have "Γ \<turnstile> App q t \<leftrightarrow> u : T2" by fact then obtain r T1' v where ha: "Γ \<turnstile> q \<leftrightarrow> r : T1'->T2" and hb: "Γ \<turnstile> t \<Leftrightarrow> v : T1'" and eq: "u = App r v" by auto have ih1: "Γ \<turnstile> q \<leftrightarrow> r : T1->T2 ==> Γ \<turnstile> p \<leftrightarrow> r : T1->T2" by fact have ih2:"Γ \<turnstile> t \<Leftrightarrow> v : T1 ==> Γ \<turnstile> s \<Leftrightarrow> v : T1" by fact have "Γ \<turnstile> p \<leftrightarrow> q : T1->T2" by fact then have "Γ \<turnstile> q \<leftrightarrow> p : T1->T2" by (simp add: algorithmic_symmetry) with ha have "T1'->T2 = T1->T2" using algorithmic_path_type_unicity by simp then have "T1' = T1" by (simp add: ty.inject) then have "Γ \<turnstile> s \<Leftrightarrow> v : T1" "Γ \<turnstile> p \<leftrightarrow> r : T1->T2" using ih1 ih2 ha hb by auto then show "Γ \<turnstile> App p s \<leftrightarrow> u : T2" using eq by auto qed (auto) lemma algorithmic_weak_head_closure: shows "Γ \<turnstile> s \<Leftrightarrow> t : T ==> s' \<leadsto> s ==> t' \<leadsto> t ==> Γ \<turnstile> s' \<Leftrightarrow> t' : T" apply (nominal_induct Γ s t T avoiding: s' t' rule: alg_equiv_alg_path_equiv.strong_inducts(1) [of _ _ _ _ "%a b c d e. True"]) apply(auto intro!: QAT_Arrow) done lemma algorithmic_monotonicity: shows "Γ \<turnstile> s \<Leftrightarrow> t : T ==> Γ ⊆ Γ' ==> valid Γ' ==> Γ' \<turnstile> s \<Leftrightarrow> t : T" and "Γ \<turnstile> s \<leftrightarrow> t : T ==> Γ ⊆ Γ' ==> valid Γ' ==> Γ' \<turnstile> s \<leftrightarrow> t : T" proof (nominal_induct Γ s t T and Γ s t T avoiding: Γ' rule: alg_equiv_alg_path_equiv.strong_inducts) case (QAT_Arrow x Γ s t T1 T2 Γ') have fs:"x\<sharp>Γ" "x\<sharp>s" "x\<sharp>t" "x\<sharp>Γ'" by fact+ have h2:"Γ ⊆ Γ'" by fact have ih:"!!Γ'. [|(x,T1)#Γ ⊆ Γ'; valid Γ'|] ==> Γ' \<turnstile> App s (Var x) \<Leftrightarrow> App t (Var x) : T2" by fact have "valid Γ'" by fact then have "valid ((x,T1)#Γ')" using fs by auto moreover have sub: "(x,T1)#Γ ⊆ (x,T1)#Γ'" using h2 by auto ultimately have "(x,T1)#Γ' \<turnstile> App s (Var x) \<Leftrightarrow> App t (Var x) : T2" using ih by simp then show "Γ' \<turnstile> s \<Leftrightarrow> t : T1->T2" using fs by (auto simp add: fresh_prod) qed (auto) lemma path_equiv_implies_nf: assumes "Γ \<turnstile> s \<leftrightarrow> t : T" shows "s \<leadsto>|" and "t \<leadsto>|" using assms by (induct rule: alg_equiv_alg_path_equiv.inducts(2)) (simp, auto) section {* Logical Equivalence *} function log_equiv :: "(Ctxt => trm => trm => ty => bool)" ("_ \<turnstile> _ is _ : _" [60,60,60,60] 60) where "Γ \<turnstile> s is t : TUnit = True" | "Γ \<turnstile> s is t : TBase = Γ \<turnstile> s \<Leftrightarrow> t : TBase" | "Γ \<turnstile> s is t : (T1 -> T2) = (∀Γ' s' t'. Γ⊆Γ' --> valid Γ' --> Γ' \<turnstile> s' is t' : T1 --> (Γ' \<turnstile> (App s s') is (App t t') : T2))" apply (auto simp add: ty.inject) apply (subgoal_tac "(∃T1 T2. b=T1 -> T2) ∨ b=TUnit ∨ b=TBase" ) apply (force) apply (rule ty_cases) done termination apply(relation "measure (λ(_,_,_,T). size T)") apply(auto) done lemma logical_monotonicity: fixes Γ Γ' :: Ctxt assumes a1: "Γ \<turnstile> s is t : T" and a2: "Γ ⊆ Γ'" and a3: "valid Γ'" shows "Γ' \<turnstile> s is t : T" using a1 a2 a3 proof (induct arbitrary: Γ' rule: log_equiv.induct) case (2 Γ s t Γ') then show "Γ' \<turnstile> s is t : TBase" using algorithmic_monotonicity by auto next case (3 Γ s t T1 T2 Γ') have "Γ \<turnstile> s is t : T1->T2" and "Γ ⊆ Γ'" and "valid Γ'" by fact+ then show "Γ' \<turnstile> s is t : T1->T2" by simp qed (auto) lemma main_lemma: shows "Γ \<turnstile> s is t : T ==> valid Γ ==> Γ \<turnstile> s \<Leftrightarrow> t : T" and "Γ \<turnstile> p \<leftrightarrow> q : T ==> Γ \<turnstile> p is q : T" proof (nominal_induct T arbitrary: Γ s t p q rule: ty.induct) case (Arrow T1 T2) { case (1 Γ s t) have ih1:"!!Γ s t. [|Γ \<turnstile> s is t : T2; valid Γ|] ==> Γ \<turnstile> s \<Leftrightarrow> t : T2" by fact have ih2:"!!Γ s t. Γ \<turnstile> s \<leftrightarrow> t : T1 ==> Γ \<turnstile> s is t : T1" by fact have h:"Γ \<turnstile> s is t : T1->T2" by fact obtain x::name where fs:"x\<sharp>(Γ,s,t)" by (erule exists_fresh[OF fs_name1]) have "valid Γ" by fact then have v: "valid ((x,T1)#Γ)" using fs by auto then have "(x,T1)#Γ \<turnstile> Var x \<leftrightarrow> Var x : T1" by auto then have "(x,T1)#Γ \<turnstile> Var x is Var x : T1" using ih2 by auto then have "(x,T1)#Γ \<turnstile> App s (Var x) is App t (Var x) : T2" using h v by auto then have "(x,T1)#Γ \<turnstile> App s (Var x) \<Leftrightarrow> App t (Var x) : T2" using ih1 v by auto then show "Γ \<turnstile> s \<Leftrightarrow> t : T1->T2" using fs by (auto simp add: fresh_prod) next case (2 Γ p q) have h: "Γ \<turnstile> p \<leftrightarrow> q : T1->T2" by fact have ih1:"!!Γ s t. Γ \<turnstile> s \<leftrightarrow> t : T2 ==> Γ \<turnstile> s is t : T2" by fact have ih2:"!!Γ s t. [|Γ \<turnstile> s is t : T1; valid Γ|] ==> Γ \<turnstile> s \<Leftrightarrow> t : T1" by fact { fix Γ' s t assume "Γ ⊆ Γ'" and hl:"Γ' \<turnstile> s is t : T1" and hk: "valid Γ'" then have "Γ' \<turnstile> p \<leftrightarrow> q : T1 -> T2" using h algorithmic_monotonicity by auto moreover have "Γ' \<turnstile> s \<Leftrightarrow> t : T1" using ih2 hl hk by auto ultimately have "Γ' \<turnstile> App p s \<leftrightarrow> App q t : T2" by auto then have "Γ' \<turnstile> App p s is App q t : T2" using ih1 by auto } then show "Γ \<turnstile> p is q : T1->T2" by simp } next case TBase { case 2 have h:"Γ \<turnstile> s \<leftrightarrow> t : TBase" by fact then have "s \<leadsto>|" and "t \<leadsto>|" using path_equiv_implies_nf by auto then have "s \<Down> s" and "t \<Down> t" by auto then have "Γ \<turnstile> s \<Leftrightarrow> t : TBase" using h by auto then show "Γ \<turnstile> s is t : TBase" by auto } qed (auto elim: alg_path_equiv_implies_valid) corollary corollary_main: assumes a: "Γ \<turnstile> s \<leftrightarrow> t : T" shows "Γ \<turnstile> s \<Leftrightarrow> t : T" using a main_lemma alg_path_equiv_implies_valid by blast lemma logical_symmetry: assumes a: "Γ \<turnstile> s is t : T" shows "Γ \<turnstile> t is s : T" using a by (nominal_induct arbitrary: Γ s t rule: ty.induct) (auto simp add: algorithmic_symmetry) lemma logical_transitivity: assumes "Γ \<turnstile> s is t : T" "Γ \<turnstile> t is u : T" shows "Γ \<turnstile> s is u : T" using assms proof (nominal_induct arbitrary: Γ s t u rule:ty.induct) case TBase then show "Γ \<turnstile> s is u : TBase" by (auto elim: algorithmic_transitivity) next case (Arrow T1 T2 Γ s t u) have h1:"Γ \<turnstile> s is t : T1 -> T2" by fact have h2:"Γ \<turnstile> t is u : T1 -> T2" by fact have ih1:"!!Γ s t u. [|Γ \<turnstile> s is t : T1; Γ \<turnstile> t is u : T1|] ==> Γ \<turnstile> s is u : T1" by fact have ih2:"!!Γ s t u. [|Γ \<turnstile> s is t : T2; Γ \<turnstile> t is u : T2|] ==> Γ \<turnstile> s is u : T2" by fact { fix Γ' s' u' assume hsub:"Γ ⊆ Γ'" and hl:"Γ' \<turnstile> s' is u' : T1" and hk: "valid Γ'" then have "Γ' \<turnstile> u' is s' : T1" using logical_symmetry by blast then have "Γ' \<turnstile> u' is u' : T1" using ih1 hl by blast then have "Γ' \<turnstile> App t u' is App u u' : T2" using h2 hsub hk by auto moreover have "Γ' \<turnstile> App s s' is App t u' : T2" using h1 hsub hl hk by auto ultimately have "Γ' \<turnstile> App s s' is App u u' : T2" using ih2 by blast } then show "Γ \<turnstile> s is u : T1 -> T2" by auto qed (auto) lemma logical_weak_head_closure: assumes a: "Γ \<turnstile> s is t : T" and b: "s' \<leadsto> s" and c: "t' \<leadsto> t" shows "Γ \<turnstile> s' is t' : T" using a b c algorithmic_weak_head_closure by (nominal_induct arbitrary: Γ s t s' t' rule: ty.induct) (auto, blast) lemma logical_weak_head_closure': assumes "Γ \<turnstile> s is t : T" and "s' \<leadsto> s" shows "Γ \<turnstile> s' is t : T" using assms proof (nominal_induct arbitrary: Γ s t s' rule: ty.induct) case (TBase Γ s t s') then show ?case by force next case (TUnit Γ s t s') then show ?case by auto next case (Arrow T1 T2 Γ s t s') have h1:"s' \<leadsto> s" by fact have ih:"!!Γ s t s'. [|Γ \<turnstile> s is t : T2; s' \<leadsto> s|] ==> Γ \<turnstile> s' is t : T2" by fact have h2:"Γ \<turnstile> s is t : T1->T2" by fact then have hb:"∀Γ' s' t'. Γ⊆Γ' --> valid Γ' --> Γ' \<turnstile> s' is t' : T1 --> (Γ' \<turnstile> (App s s') is (App t t') : T2)" by auto { fix Γ' s2 t2 assume "Γ ⊆ Γ'" and "Γ' \<turnstile> s2 is t2 : T1" and "valid Γ'" then have "Γ' \<turnstile> (App s s2) is (App t t2) : T2" using hb by auto moreover have "(App s' s2) \<leadsto> (App s s2)" using h1 by auto ultimately have "Γ' \<turnstile> App s' s2 is App t t2 : T2" using ih by auto } then show "Γ \<turnstile> s' is t : T1->T2" by auto qed abbreviation log_equiv_for_psubsts :: "Ctxt => Subst => Subst => Ctxt => bool" ("_ \<turnstile> _ is _ over _" [60,60] 60) where "Γ' \<turnstile> ϑ is ϑ' over Γ ≡ ∀x T. (x,T) ∈ set Γ --> Γ' \<turnstile> ϑ<Var x> is ϑ'<Var x> : T" lemma logical_pseudo_reflexivity: assumes "Γ' \<turnstile> t is s over Γ" shows "Γ' \<turnstile> s is s over Γ" proof - have "Γ' \<turnstile> t is s over Γ" by fact moreover then have "Γ' \<turnstile> s is t over Γ" using logical_symmetry by blast ultimately show "Γ' \<turnstile> s is s over Γ" using logical_transitivity by blast qed lemma logical_subst_monotonicity : fixes Γ Γ' Γ'' :: Ctxt assumes a: "Γ' \<turnstile> ϑ is ϑ' over Γ" and b: "Γ' ⊆ Γ''" and c: "valid Γ''" shows "Γ'' \<turnstile> ϑ is ϑ' over Γ" using a b c logical_monotonicity by blast lemma equiv_subst_ext : assumes h1: "Γ' \<turnstile> ϑ is ϑ' over Γ" and h2: "Γ' \<turnstile> s is t : T" and fs: "x\<sharp>Γ" shows "Γ' \<turnstile> (x,s)#ϑ is (x,t)#ϑ' over (x,T)#Γ" using assms proof - { fix y U assume "(y,U) ∈ set ((x,T)#Γ)" moreover { assume "(y,U) ∈ set [(x,T)]" with h2 have "Γ' \<turnstile> ((x,s)#ϑ)<Var y> is ((x,t)#ϑ')<Var y> : U" by auto } moreover { assume hl:"(y,U) ∈ set Γ" then have "¬ y\<sharp>Γ" by (induct Γ) (auto simp add: fresh_list_cons fresh_atm fresh_prod) then have hf:"x\<sharp> Var y" using fs by (auto simp add: fresh_atm) then have "((x,s)#ϑ)<Var y> = ϑ<Var y>" "((x,t)#ϑ')<Var y> = ϑ'<Var y>" using fresh_psubst_simp by blast+ moreover have "Γ' \<turnstile> ϑ<Var y> is ϑ'<Var y> : U" using h1 hl by auto ultimately have "Γ' \<turnstile> ((x,s)#ϑ)<Var y> is ((x,t)#ϑ')<Var y> : U" by auto } ultimately have "Γ' \<turnstile> ((x,s)#ϑ)<Var y> is ((x,t)#ϑ')<Var y> : U" by auto } then show "Γ' \<turnstile> (x,s)#ϑ is (x,t)#ϑ' over (x,T)#Γ" by auto qed theorem fundamental_theorem_1: assumes a1: "Γ \<turnstile> t : T" and a2: "Γ' \<turnstile> ϑ is ϑ' over Γ" and a3: "valid Γ'" shows "Γ' \<turnstile> ϑ<t> is ϑ'<t> : T" using a1 a2 a3 proof (nominal_induct Γ t T avoiding: ϑ ϑ' arbitrary: Γ' rule: typing.strong_induct) case (T_Lam x Γ T1 t2 T2 ϑ ϑ' Γ') have vc: "x\<sharp>ϑ" "x\<sharp>ϑ'" "x\<sharp>Γ" by fact+ have asm1: "Γ' \<turnstile> ϑ is ϑ' over Γ" by fact have ih:"!!ϑ ϑ' Γ'. [|Γ' \<turnstile> ϑ is ϑ' over (x,T1)#Γ; valid Γ'|] ==> Γ' \<turnstile> ϑ<t2> is ϑ'<t2> : T2" by fact show "Γ' \<turnstile> ϑ<Lam [x].t2> is ϑ'<Lam [x].t2> : T1->T2" using vc proof (simp, intro strip) fix Γ'' s' t' assume sub: "Γ' ⊆ Γ''" and asm2: "Γ''\<turnstile> s' is t' : T1" and val: "valid Γ''" from asm1 val sub have "Γ'' \<turnstile> ϑ is ϑ' over Γ" using logical_subst_monotonicity by blast with asm2 vc have "Γ'' \<turnstile> (x,s')#ϑ is (x,t')#ϑ' over (x,T1)#Γ" using equiv_subst_ext by blast with ih val have "Γ'' \<turnstile> ((x,s')#ϑ)<t2> is ((x,t')#ϑ')<t2> : T2" by auto with vc have "Γ''\<turnstile>ϑ<t2>[x::=s'] is ϑ'<t2>[x::=t'] : T2" by (simp add: psubst_subst_psubst) moreover have "App (Lam [x].ϑ<t2>) s' \<leadsto> ϑ<t2>[x::=s']" by auto moreover have "App (Lam [x].ϑ'<t2>) t' \<leadsto> ϑ'<t2>[x::=t']" by auto ultimately show "Γ''\<turnstile> App (Lam [x].ϑ<t2>) s' is App (Lam [x].ϑ'<t2>) t' : T2" using logical_weak_head_closure by auto qed qed (auto) theorem fundamental_theorem_2: assumes h1: "Γ \<turnstile> s ≡ t : T" and h2: "Γ' \<turnstile> ϑ is ϑ' over Γ" and h3: "valid Γ'" shows "Γ' \<turnstile> ϑ<s> is ϑ'<t> : T" using h1 h2 h3 proof (nominal_induct Γ s t T avoiding: Γ' ϑ ϑ' rule: def_equiv.strong_induct) case (Q_Refl Γ t T Γ' ϑ ϑ') have "Γ \<turnstile> t : T" and "valid Γ'" by fact+ moreover have "Γ' \<turnstile> ϑ is ϑ' over Γ" by fact ultimately show "Γ' \<turnstile> ϑ<t> is ϑ'<t> : T" using fundamental_theorem_1 by blast next case (Q_Symm Γ t s T Γ' ϑ ϑ') have "Γ' \<turnstile> ϑ is ϑ' over Γ" and "valid Γ'" by fact+ moreover have ih: "!! Γ' ϑ ϑ'. [|Γ' \<turnstile> ϑ is ϑ' over Γ; valid Γ'|] ==> Γ' \<turnstile> ϑ<t> is ϑ'<s> : T" by fact ultimately show "Γ' \<turnstile> ϑ<s> is ϑ'<t> : T" using logical_symmetry by blast next case (Q_Trans Γ s t T u Γ' ϑ ϑ') have ih1: "!! Γ' ϑ ϑ'. [|Γ' \<turnstile> ϑ is ϑ' over Γ; valid Γ'|] ==> Γ' \<turnstile> ϑ<s> is ϑ'<t> : T" by fact have ih2: "!! Γ' ϑ ϑ'. [|Γ' \<turnstile> ϑ is ϑ' over Γ; valid Γ'|] ==> Γ' \<turnstile> ϑ<t> is ϑ'<u> : T" by fact have h: "Γ' \<turnstile> ϑ is ϑ' over Γ" and v: "valid Γ'" by fact+ then have "Γ' \<turnstile> ϑ' is ϑ' over Γ" using logical_pseudo_reflexivity by auto then have "Γ' \<turnstile> ϑ'<t> is ϑ'<u> : T" using ih2 v by auto moreover have "Γ' \<turnstile> ϑ<s> is ϑ'<t> : T" using ih1 h v by auto ultimately show "Γ' \<turnstile> ϑ<s> is ϑ'<u> : T" using logical_transitivity by blast next case (Q_Abs x Γ T1 s2 t2 T2 Γ' ϑ ϑ') have fs:"x\<sharp>Γ" by fact have fs2: "x\<sharp>ϑ" "x\<sharp>ϑ'" by fact+ have h2: "Γ' \<turnstile> ϑ is ϑ' over Γ" and h3: "valid Γ'" by fact+ have ih:"!!Γ' ϑ ϑ'. [|Γ' \<turnstile> ϑ is ϑ' over (x,T1)#Γ; valid Γ'|] ==> Γ' \<turnstile> ϑ<s2> is ϑ'<t2> : T2" by fact { fix Γ'' s' t' assume "Γ' ⊆ Γ''" and hl:"Γ''\<turnstile> s' is t' : T1" and hk: "valid Γ''" then have "Γ'' \<turnstile> ϑ is ϑ' over Γ" using h2 logical_subst_monotonicity by blast then have "Γ'' \<turnstile> (x,s')#ϑ is (x,t')#ϑ' over (x,T1)#Γ" using equiv_subst_ext hl fs by blast then have "Γ'' \<turnstile> ((x,s')#ϑ)<s2> is ((x,t')#ϑ')<t2> : T2" using ih hk by blast then have "Γ''\<turnstile> ϑ<s2>[x::=s'] is ϑ'<t2>[x::=t'] : T2" using fs2 psubst_subst_psubst by auto moreover have "App (Lam [x]. ϑ<s2>) s' \<leadsto> ϑ<s2>[x::=s']" and "App (Lam [x].ϑ'<t2>) t' \<leadsto> ϑ'<t2>[x::=t']" by auto ultimately have "Γ'' \<turnstile> App (Lam [x]. ϑ<s2>) s' is App (Lam [x].ϑ'<t2>) t' : T2" using logical_weak_head_closure by auto } moreover have "valid Γ'" by fact ultimately have "Γ' \<turnstile> Lam [x].ϑ<s2> is Lam [x].ϑ'<t2> : T1->T2" by auto then show "Γ' \<turnstile> ϑ<Lam [x].s2> is ϑ'<Lam [x].t2> : T1->T2" using fs2 by auto next case (Q_App Γ s1 t1 T1 T2 s2 t2 Γ' ϑ ϑ') then show "Γ' \<turnstile> ϑ<App s1 s2> is ϑ'<App t1 t2> : T2" by auto next case (Q_Beta x Γ s2 t2 T1 s12 t12 T2 Γ' ϑ ϑ') have h: "Γ' \<turnstile> ϑ is ϑ' over Γ" and h': "valid Γ'" by fact+ have fs: "x\<sharp>Γ" by fact have fs2: " x\<sharp>ϑ" "x\<sharp>ϑ'" by fact+ have ih1: "!!Γ' ϑ ϑ'. [|Γ' \<turnstile> ϑ is ϑ' over Γ; valid Γ'|] ==> Γ' \<turnstile> ϑ<s2> is ϑ'<t2> : T1" by fact have ih2: "!!Γ' ϑ ϑ'. [|Γ' \<turnstile> ϑ is ϑ' over (x,T1)#Γ; valid Γ'|] ==> Γ' \<turnstile> ϑ<s12> is ϑ'<t12> : T2" by fact have "Γ' \<turnstile> ϑ<s2> is ϑ'<t2> : T1" using ih1 h' h by auto then have "Γ' \<turnstile> (x,ϑ<s2>)#ϑ is (x,ϑ'<t2>)#ϑ' over (x,T1)#Γ" using equiv_subst_ext h fs by blast then have "Γ' \<turnstile> ((x,ϑ<s2>)#ϑ)<s12> is ((x,ϑ'<t2>)#ϑ')<t12> : T2" using ih2 h' by auto then have "Γ' \<turnstile> ϑ<s12>[x::=ϑ<s2>] is ϑ'<t12>[x::=ϑ'<t2>] : T2" using fs2 psubst_subst_psubst by auto then have "Γ' \<turnstile> ϑ<s12>[x::=ϑ<s2>] is ϑ'<t12[x::=t2]> : T2" using fs2 psubst_subst_propagate by auto moreover have "App (Lam [x].ϑ<s12>) (ϑ<s2>) \<leadsto> ϑ<s12>[x::=ϑ<s2>]" by auto ultimately have "Γ' \<turnstile> App (Lam [x].ϑ<s12>) (ϑ<s2>) is ϑ'<t12[x::=t2]> : T2" using logical_weak_head_closure' by auto then show "Γ' \<turnstile> ϑ<App (Lam [x].s12) s2> is ϑ'<t12[x::=t2]> : T2" using fs2 by simp next case (Q_Ext x Γ s t T1 T2 Γ' ϑ ϑ') have h2: "Γ' \<turnstile> ϑ is ϑ' over Γ" and h2': "valid Γ'" by fact+ have fs:"x\<sharp>Γ" "x\<sharp>s" "x\<sharp>t" by fact+ have ih:"!!Γ' ϑ ϑ'. [|Γ' \<turnstile> ϑ is ϑ' over (x,T1)#Γ; valid Γ'|] ==> Γ' \<turnstile> ϑ<App s (Var x)> is ϑ'<App t (Var x)> : T2" by fact { fix Γ'' s' t' assume hsub: "Γ' ⊆ Γ''" and hl: "Γ''\<turnstile> s' is t' : T1" and hk: "valid Γ''" then have "Γ'' \<turnstile> ϑ is ϑ' over Γ" using h2 logical_subst_monotonicity by blast then have "Γ'' \<turnstile> (x,s')#ϑ is (x,t')#ϑ' over (x,T1)#Γ" using equiv_subst_ext hl fs by blast then have "Γ'' \<turnstile> ((x,s')#ϑ)<App s (Var x)> is ((x,t')#ϑ')<App t (Var x)> : T2" using ih hk by blast then have "Γ'' \<turnstile> App (((x,s')#ϑ)<s>) (((x,s')#ϑ)<(Var x)>) is App (((x,t')#ϑ')<t>) (((x,t')#ϑ')<(Var x)>) : T2" by auto then have "Γ'' \<turnstile> App ((x,s')#ϑ)<s> s' is App ((x,t')#ϑ')<t> t' : T2" by auto then have "Γ'' \<turnstile> App (ϑ<s>) s' is App (ϑ'<t>) t' : T2" using fs fresh_psubst_simp by auto } moreover have "valid Γ'" by fact ultimately show "Γ' \<turnstile> ϑ<s> is ϑ'<t> : T1->T2" by auto next case (Q_Unit Γ s t Γ' ϑ ϑ') then show "Γ' \<turnstile> ϑ<s> is ϑ'<t> : TUnit" by auto qed theorem completeness: assumes asm: "Γ \<turnstile> s ≡ t : T" shows "Γ \<turnstile> s \<Leftrightarrow> t : T" proof - have val: "valid Γ" using def_equiv_implies_valid asm by simp moreover { fix x T assume "(x,T) ∈ set Γ" "valid Γ" then have "Γ \<turnstile> Var x is Var x : T" using main_lemma(2) by blast } ultimately have "Γ \<turnstile> [] is [] over Γ" by auto then have "Γ \<turnstile> []<s> is []<t> : T" using fundamental_theorem_2 val asm by blast then have "Γ \<turnstile> s is t : T" by simp then show "Γ \<turnstile> s \<Leftrightarrow> t : T" using main_lemma(1) val by simp qed text {* We leave soundness as an exercise - like in the book :-) \\ @{prop[mode=IfThen] "[|Γ \<turnstile> s \<Leftrightarrow> t : T; Γ \<turnstile> t : T; Γ \<turnstile> s : T|] ==> Γ \<turnstile> s ≡ t : T"} \\ @{prop "[|Γ \<turnstile> s \<leftrightarrow> t : T; Γ \<turnstile> t : T; Γ \<turnstile> s : T|] ==> Γ \<turnstile> s ≡ t : T"} *} end
lemma perm_ty:
pi • T = T
lemma fresh_ty:
x \<sharp> T
lemma ty_cases:
(∃T1 T2. T = T1->T2) ∨ T = TUnit ∨ T = TBase
lemma ty_size_greater_zero:
0 < size T
lemma lookup_eqvt:
pi • lookup ϑ x = lookup (pi • ϑ) (pi • x)
lemma lookup_fresh:
[| z \<sharp> ϑ; z \<sharp> x |] ==> z \<sharp> lookup ϑ x
lemma lookup_fresh':
z \<sharp> ϑ ==> lookup ϑ z = Var z
lemma subst(1):
Var x[y::=t'] = (if x = y then t' else Var x)
and subst(2):
App t1 t2[y::=t'] = App (t1[y::=t']) (t2[y::=t'])
and subst(3):
x \<sharp> (y, t') ==> Lam [x].t[y::=t'] = Lam [x].t[y::=t']
and subst(4):
Const n[y::=t'] = Const n
and subst(5):
Unit[y::=t'] = Unit
lemma subst_eqvt:
pi • t[x::=t'] = (pi • t)[(pi • x)::=(pi • t')]
lemma subst_rename:
c \<sharp> t1 ==> t1[a::=t2] = ([(c, a)] • t1)[c::=t2]
lemma fresh_psubst:
[| z \<sharp> t; z \<sharp> ϑ |] ==> z \<sharp> ϑ<t>
lemma fresh_subst'':
z \<sharp> t2 ==> z \<sharp> t1[z::=t2]
lemma fresh_subst':
[| z \<sharp> [y].t1; z \<sharp> t2 |] ==> z \<sharp> t1[y::=t2]
lemma fresh_subst:
[| z \<sharp> t1; z \<sharp> t2 |] ==> z \<sharp> t1[y::=t2]
lemma fresh_psubst_simp:
x \<sharp> t ==> ((x, u) # ϑ)<t> = ϑ<t>
lemma forget:
x \<sharp> t ==> t[x::=t'] = t
lemma subst_fun_eq:
[x].t1 = [y].t2 ==> t1[x::=u] = t2[y::=u]
lemma psubst_empty:
[]<t> = t
lemma psubst_subst_psubst:
c \<sharp> ϑ ==> ϑ<t>[c::=s] = ((c, s) # ϑ)<t>
lemma subst_fresh_simp:
x \<sharp> ϑ ==> ϑ<Var x> = Var x
lemma psubst_subst_propagate:
x \<sharp> ϑ ==> ϑ<t[x::=u]> = ϑ<t>[x::=ϑ<u>]
lemma valid_monotonicity:
[| Γ ⊆ Γ' ; x \<sharp> Γ' |] ==> (x, T1) # Γ ⊆ (x, T1) # Γ'
lemma fresh_context:
a \<sharp> Γ ==> ¬ (∃τ. (a, τ) ∈ set Γ)
lemma type_unicity_in_context:
[| valid Γ; (x, T1) ∈ set Γ; (x, T2) ∈ set Γ |] ==> T1 = T2
lemma typing_implies_valid:
Γ \<turnstile> t : T ==> valid Γ
lemma def_equiv_implies_valid:
Γ \<turnstile> t ≡ s : T ==> valid Γ
lemma test30:
(x, T) ∈ set Γ ==> x ∈ supp Γ
lemma supp_ty:
supp T = {}
lemma test3a:
Γ \<turnstile> t : T ==> supp t ⊆ supp Γ
lemma test3b:
supp (t1[x::=t2]) ⊆ supp ([x].t1, t2)
lemma test3:
Γ \<turnstile> s ≡ t : T ==> supp (s, t) ⊆ supp Γ
lemma test0:
[| (x, T1) # Γ \<turnstile> s1 ≡ t1 : T2; Γ \<turnstile> s2 ≡ t2 : T1 |]
==> Γ \<turnstile> App (Lam [x].s1) s2 ≡ t1[x::=t2] : T2
lemma test1:
∀x. x \<sharp> Γ --> (x, T1) # Γ \<turnstile> App s (Var x) ≡ App t (Var x) : T2
==> Γ \<turnstile> s ≡ t : T1->T2
lemma test2:
x \<sharp> (Γ, s, t) ∧
(x, T1) # Γ \<turnstile> App s (Var x) ≡ App t (Var x) : T2
==> ∀x. x \<sharp> (Γ, s, t) -->
(x, T1) # Γ \<turnstile> App s (Var x) ≡ App t (Var x) : T2
lemma test2:
x \<sharp> (Γ, s, t) ∧
(x, T1) # Γ \<turnstile> App s (Var x) ≡ App t (Var x) : T2
==> ∀x. x \<sharp> Γ -->
(x, T1) # Γ \<turnstile> App s (Var x) ≡ App t (Var x) : T2
lemma whn_eqvt:
t \<Down> t' ==> pi • t \<Down> pi • t'
lemma red_unicity:
[| x \<leadsto> a; x \<leadsto> b |] ==> a = b
lemma nf_unicity:
[| x \<Down> a; x \<Down> b |] ==> a = b
lemma test4a:
s \<leadsto> t ==> supp t ⊆ supp s
lemma test4b:
s \<Down> t ==> supp t ⊆ supp s
lemma Q_Arrow_strong_inversion:
[| x \<sharp> Γ; x \<sharp> t; x \<sharp> u;
Γ \<turnstile> t \<Leftrightarrow> u : T1->T2 |]
==> (x, T1) # Γ \<turnstile> App t (Var x) \<Leftrightarrow> App u (Var x) : T2
lemma algorithmic_path_type_unicity:
[| Γ \<turnstile> s \<leftrightarrow> t : T;
Γ \<turnstile> s \<leftrightarrow> u : T' |]
==> T = T'
lemma alg_path_equiv_implies_valid(1):
Γ \<turnstile> s \<Leftrightarrow> t : T ==> valid Γ
and alg_path_equiv_implies_valid(2):
Γ \<turnstile> s \<leftrightarrow> t : T ==> valid Γ
lemma algorithmic_symmetry(1):
Γ \<turnstile> s \<Leftrightarrow> t : T
==> Γ \<turnstile> t \<Leftrightarrow> s : T
and algorithmic_symmetry(2):
Γ \<turnstile> s \<leftrightarrow> t : T
==> Γ \<turnstile> t \<leftrightarrow> s : T
lemma algorithmic_transitivity(1):
[| Γ \<turnstile> s \<Leftrightarrow> t : T;
Γ \<turnstile> t \<Leftrightarrow> u : T |]
==> Γ \<turnstile> s \<Leftrightarrow> u : T
and algorithmic_transitivity(2):
[| Γ \<turnstile> s \<leftrightarrow> t : T;
Γ \<turnstile> t \<leftrightarrow> u : T |]
==> Γ \<turnstile> s \<leftrightarrow> u : T
lemma algorithmic_weak_head_closure:
[| Γ \<turnstile> s \<Leftrightarrow> t : T; s' \<leadsto> s; t' \<leadsto> t |]
==> Γ \<turnstile> s' \<Leftrightarrow> t' : T
lemma algorithmic_monotonicity(1):
[| Γ \<turnstile> s \<Leftrightarrow> t : T; Γ ⊆ Γ' ; valid Γ' |]
==> Γ' \<turnstile> s \<Leftrightarrow> t : T
and algorithmic_monotonicity(2):
[| Γ \<turnstile> s \<leftrightarrow> t : T; Γ ⊆ Γ' ; valid Γ' |]
==> Γ' \<turnstile> s \<leftrightarrow> t : T
lemma path_equiv_implies_nf(1):
Γ \<turnstile> s \<leftrightarrow> t : T ==> s \<leadsto>|
and path_equiv_implies_nf(2):
Γ \<turnstile> s \<leftrightarrow> t : T ==> t \<leadsto>|
lemma logical_monotonicity:
[| Γ \<turnstile> s is t : T; Γ ⊆ Γ' ; valid Γ' |]
==> Γ' \<turnstile> s is t : T
lemma main_lemma(1):
[| Γ \<turnstile> s is t : T; valid Γ |]
==> Γ \<turnstile> s \<Leftrightarrow> t : T
and main_lemma(2):
Γ \<turnstile> p \<leftrightarrow> q : T ==> Γ \<turnstile> p is q : T
corollary corollary_main:
Γ \<turnstile> s \<leftrightarrow> t : T
==> Γ \<turnstile> s \<Leftrightarrow> t : T
lemma logical_symmetry:
Γ \<turnstile> s is t : T ==> Γ \<turnstile> t is s : T
lemma logical_transitivity:
[| Γ \<turnstile> s is t : T; Γ \<turnstile> t is u : T |]
==> Γ \<turnstile> s is u : T
lemma logical_weak_head_closure:
[| Γ \<turnstile> s is t : T; s' \<leadsto> s; t' \<leadsto> t |]
==> Γ \<turnstile> s' is t' : T
lemma logical_weak_head_closure':
[| Γ \<turnstile> s is t : T; s' \<leadsto> s |] ==> Γ \<turnstile> s' is t : T
lemma logical_pseudo_reflexivity:
Γ' \<turnstile> t is s over Γ ==> Γ' \<turnstile> s is s over Γ
lemma logical_subst_monotonicity:
[| Γ' \<turnstile> ϑ is ϑ' over Γ; Γ' ⊆ Γ'' ; valid Γ'' |]
==> Γ'' \<turnstile> ϑ is ϑ' over Γ
lemma equiv_subst_ext:
[| Γ' \<turnstile> ϑ is ϑ' over Γ; Γ' \<turnstile> s is t : T; x \<sharp> Γ |]
==> Γ' \<turnstile> (x, s) # ϑ is (x, t) # ϑ' over (x, T) # Γ
theorem fundamental_theorem_1:
[| Γ \<turnstile> t : T ; Γ' \<turnstile> ϑ is ϑ' over Γ; valid Γ' |]
==> Γ' \<turnstile> ϑ<t> is ϑ'<t> : T
theorem fundamental_theorem_2:
[| Γ \<turnstile> s ≡ t : T; Γ' \<turnstile> ϑ is ϑ' over Γ; valid Γ' |]
==> Γ' \<turnstile> ϑ<s> is ϑ'<t> : T
theorem completeness:
Γ \<turnstile> s ≡ t : T ==> Γ \<turnstile> s \<Leftrightarrow> t : T