Theory SOS

Up to index of Isabelle/HOL/HOL-Nominal/Examples

theory SOS
imports Nominal
begin

(* "$Id: SOS.thy,v 1.17 2007/09/23 20:11:50 urbanc Exp $" *)
(*                                                   *)
(* Formalisation of some typical SOS-proofs          *)
(*                                                   *) 
(* This work arose from challenge suggested by Adam  *)
(* Chlipala suggested on the POPLmark mailing list.  *)
(*                                                   *) 
(* We thank Nick Benton for helping us with the      *) 
(* termination-proof for evaluation.                 *)
(*                                                   *)
(* The formalisation was done by Julien Narboux and  *)
(* Christian Urban.                                  *)

theory SOS
  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)

lemma in_eqvt[eqvt]:
  fixes pi::"name prm"
  and   x::"'a::pt_name"
  assumes "x∈X"
  shows "pi•x ∈ pi•X"
  using assms by (perm_simp add: pt_set_bij1a[OF pt_name_inst, OF at_name_inst])

lemma perm_data[simp]: 
  fixes D::"data"
  and   pi::"name prm"
  shows "pi•D = D"
  by (induct D rule: data.weak_induct) (simp_all)

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)

text {* substitution *}

fun
  lookup :: "(name×trm) list => name => trm"   
where
  "lookup [] x        = Var x"
| "lookup ((y,e)#ϑ) x = (if x=y then e else lookup ϑ x)"

lemma lookup_eqvt:
  fixes pi::"name prm"
  and   ϑ::"(name×trm) list"
  and   X::"name"
  shows "pi•(lookup ϑ X) = lookup (pi•ϑ) (pi•X)"
by (induct ϑ, auto simp add: perm_bij)

lemma lookup_fresh:
  fixes z::"name"
  assumes "z\<sharp>ϑ" and "z\<sharp>x"
  shows "z \<sharp>lookup ϑ x"
using assms 
by (induct rule: lookup.induct) (auto simp add: fresh_list_cons)

lemma lookup_fresh':
  assumes "z\<sharp>ϑ"
  shows "lookup ϑ z = Var z"
using assms 
by (induct rule: lookup.induct)
   (auto simp add: fresh_list_cons fresh_prod fresh_atm)

text {* Parallel Substitution *}

consts
  psubst :: "(name×trm) list => trm => trm"  ("_<_>" [95,95] 105)
 
nominal_primrec
  "ϑ<(Var x)> = (lookup ϑ x)"
  "ϑ<(App e1 e2)> = App (ϑ<e1>) (ϑ<e2>)"
  "x\<sharp>ϑ ==> ϑ<(Lam [x].e)> = Lam [x].(ϑ<e>)"
  "ϑ<(Const n)> = Const n"
  "ϑ<(Pr e1 e2)> = Pr (ϑ<e1>) (ϑ<e2>)"
  "ϑ<(Fst e)> = Fst (ϑ<e>)"
  "ϑ<(Snd e)> = Snd (ϑ<e>)"
  "ϑ<(InL e)> = InL (ϑ<e>)"
  "ϑ<(InR e)> = InR (ϑ<e>)"
  "[|y≠x; x\<sharp>(e,e2,ϑ); y\<sharp>(e,e1,ϑ)|] 
   ==> ϑ<Case e of inl x -> e1 | inr y -> e2> = (Case (ϑ<e>) of inl x -> (ϑ<e1>) | inr y -> (ϑ<e2>))"
apply(finite_guess add: lookup_eqvt)+
apply(rule TrueI)+
apply(simp add: abs_fresh)+
apply(fresh_guess add: fs_name1 lookup_eqvt)+
done

lemma psubst_eqvt[eqvt]:
  fixes pi::"name prm" 
  and   t::"trm"
  shows "pi•(ϑ<t>) = (pi•ϑ)<(pi•t)>"
by (nominal_induct t avoiding: ϑ rule: trm.induct)
   (perm_simp add: fresh_bij lookup_eqvt)+

lemma fresh_psubst: 
  fixes z::"name"
  and   t::"trm"
  assumes "z\<sharp>t" and "z\<sharp>ϑ"
  shows "z\<sharp>(ϑ<t>)"
using assms
by (nominal_induct t avoiding: z ϑ t rule: trm.induct)
   (auto simp add: abs_fresh lookup_fresh)

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   "(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\<sharp>(y,e,e2,t'); z\<sharp>(y,e,e1,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']))"
by (simp_all add: fresh_list_cons fresh_list_nil)

lemma subst_eqvt[eqvt]:
  fixes pi::"name prm" 
  and   t::"trm"
  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 fresh_subst:
  fixes z::"name"
  and   t1::"trm"
  and   t2::"trm"
  assumes "z\<sharp>t1" and "z\<sharp>t2"
  shows "z\<sharp>t1[y::=t2]"
using assms 
by (nominal_induct t1 avoiding: z y t2 rule: trm.induct)
   (auto simp add: abs_fresh fresh_atm)

lemma fresh_subst':
  fixes z::"name"
  and   t1::"trm"
  and   t2::"trm"
  assumes "z\<sharp>[y].t1" and "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 forget: 
  fixes x::"name"
  and   L::"trm"
  assumes "x\<sharp>L" 
  shows "L[x::=P] = L"
  using assms
  by (nominal_induct L avoiding: x P rule: trm.induct)
     (auto simp add: fresh_atm abs_fresh)

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:"x\<sharp>ϑ"
shows "ϑ<e>[x::=e'] = ((x,e')#ϑ)<e>"
using h
apply(nominal_induct e avoiding: ϑ x e' rule: trm.induct)
apply(auto simp add: fresh_list_cons fresh_atm forget lookup_fresh lookup_fresh' fresh_psubst)
done

lemma fresh_subst_fresh:
    assumes "a\<sharp>e"
    shows "a\<sharp>t[a::=e]"
using assms 
by (nominal_induct t avoiding: a e rule: trm.induct)
   (auto simp add: fresh_atm abs_fresh fresh_nat) 

text {* Typing-Judgements *}

inductive
  valid :: "(name × 'a::pt_name) list => bool"
where
    v_nil[intro]:  "valid []"
  | v_cons[intro]: "[|valid Γ;x\<sharp>Γ|] ==> valid ((x,T)#Γ)"

equivariance valid 

inductive_cases  
  valid_cons_inv_auto[elim]:"valid ((x,T)#Γ)"

abbreviation
  "sub" :: "(name×ty) list => (name×ty) list => bool" ("_ ⊆ _" [55,55] 55)
where
  "Γ1 ⊆ Γ2 ≡ ∀x T. (x,T)∈set Γ1 --> (x,T)∈set Γ2"

lemma type_unicity_in_context:
  assumes asm1: "(x,t2) ∈ set ((x,t1)#Γ)" 
  and     asm2: "valid ((x,t1)#Γ)"
  shows "t1=t2"
proof -
  from asm2 have "x\<sharp>Γ" by (cases, auto)
  then have "(x,t2) ∉ set Γ"
    by (induct Γ) (auto simp add: fresh_list_cons fresh_prod fresh_atm)
  then have "(x,t2) = (x,t1)" using asm1 by auto
  then show "t1 = t2" by auto
qed

lemma case_distinction_on_context:
  fixes Γ::"(name × ty) list"
  assumes asm1: "valid ((m,t)#Γ)" 
  and     asm2: "(n,U) ∈ set ((m,T)#Γ)"
  shows "(n,U) = (m,T) ∨ ((n,U) ∈ set Γ ∧ n ≠ m)"
proof -
from asm2 have "(n,U) ∈ set [(m,T)] ∨ (n,U) ∈ set Γ" by auto
moreover
{ assume eq: "m=n"
  assume "(n,U) ∈ set Γ" 
  then have "¬ n\<sharp>Γ" 
    by (induct Γ) (auto simp add: fresh_list_cons fresh_prod fresh_atm)
  moreover have "m\<sharp>Γ" using asm1 by auto
  ultimately have False using eq by auto
}
ultimately show ?thesis by auto
qed

inductive
  typing :: "(name×ty) list=>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> e : T2|] ==> Γ \<turnstile> Lam [x].e : T1->T2"
| t_Const[intro]: "valid Γ ==> Γ \<turnstile> Const n : Data(DNat)"
| t_Pr[intro]:    "[|Γ \<turnstile> e1 : Data(S1); Γ \<turnstile> e2 : Data(S2)|] ==> Γ \<turnstile> Pr e1 e2 : Data (DProd S1 S2)"
| t_Fst[intro]:   "[|Γ \<turnstile> e : Data(DProd S1 S2)|] ==> Γ \<turnstile> Fst e : Data(S1)"
| t_Snd[intro]:   "[|Γ \<turnstile> e : Data(DProd S1 S2)|] ==> Γ \<turnstile> Snd e : Data(S2)"
| t_InL[intro]:   "[|Γ \<turnstile> e : Data(S1)|] ==> Γ \<turnstile> InL e : Data(DSum S1 S2)"
| t_InR[intro]:   "[|Γ \<turnstile> e : Data(S2)|] ==> Γ \<turnstile> InR e : Data(DSum S1 S2)"
| t_Case[intro]:  "[|x1\<sharp>(Γ,e,e2,x2); x2\<sharp>(Γ,e,e1,x1); Γ \<turnstile> e: Data(DSum S1 S2); 
                   (x1,Data(S1))#Γ \<turnstile> e1 : T; (x2,Data(S2))#Γ \<turnstile> e2 : T|] 
                   ==> Γ \<turnstile> (Case e of inl x1 -> e1 | inr x2 -> e2) : T"

equivariance typing

nominal_inductive typing
  by (simp_all add: abs_fresh fresh_prod fresh_atm)

lemmas typing_eqvt' = typing.eqvt[simplified]

lemma typing_implies_valid:
  assumes "Γ \<turnstile> t : T"
  shows "valid Γ"
  using assms
  by (induct) (auto)

declare trm.inject [simp add]
declare ty.inject  [simp add]
declare data.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> Fst x : T"
  "Γ \<turnstile> Snd x : T"
  "Γ \<turnstile> InL x : T"
  "Γ \<turnstile> InL x : Data (DSum T1 T2)"
  "Γ \<turnstile> InR x : T"
  "Γ \<turnstile> InR x : Data (DSum T1 T2)"
  "Γ \<turnstile> Pr x y : T"
  "Γ \<turnstile> Pr e1 e2 : Data (DProd σ1 σ2)"
  "Γ \<turnstile> Case e of inl x1 -> e1 | inr x2 -> e2 : T"

declare trm.inject [simp del]
declare ty.inject [simp del]
declare data.inject [simp del]

lemma t_Lam_elim[elim]: 
  assumes a1:"Γ \<turnstile> Lam [x].t : T" 
  and     a2: "x\<sharp>Γ"
  obtains T1 and T2 where "(x,T1)#Γ \<turnstile> t : T2" and "T=T1->T2"
proof -
  from a1 obtain x' t' T1 T2 
    where b1: "x'\<sharp>Γ" and b2: "(x',T1)#Γ \<turnstile> t' : T2" and b3: "[x'].t' = [x].t" and b4: "T=T1->T2"
    by auto
  obtain c::"name" where "c\<sharp>(Γ,x,x',t,t')" by (erule exists_fresh[OF fs_name1])
  then have fs: "c\<sharp>Γ" "c≠x" "c≠x'" "c\<sharp>t" "c\<sharp>t'" by (simp_all add: fresh_atm[symmetric]) 
  then have b5: "[(x',c)]•t'=[(x,c)]•t" using b3 fs by (simp add: alpha_fresh)
  have "([(x,c)]•[(x',c)]•((x',T1)#Γ)) \<turnstile> ([(x,c)]•[(x',c)]•t') : T2" using b2
    by (simp only: typing_eqvt')
  then have "(x,T1)#Γ \<turnstile> t : T2" using fs b1 a2 b5 by (perm_simp add: calc_atm)
  then show ?thesis using prems b4 by simp
qed

lemma t_Case_elim[elim]: 
  assumes "Γ \<turnstile> Case e of inl x1 -> e1 | inr x2 -> e2 : T" and "x1\<sharp>Γ" and "x2\<sharp>Γ" 
  obtains σ1 σ2 where "Γ \<turnstile> e : Data (DSum σ1 σ2)" 
                  and "(x1, Data σ1)#Γ \<turnstile> e1 : T" 
                  and "(x2, Data σ2)#Γ \<turnstile> e2 : T"
proof -
  have f:"x1\<sharp>Γ" "x2\<sharp>Γ" by fact+
  have "Γ \<turnstile> Case e of inl x1 -> e1 | inr x2 -> e2 : T" by fact
  then obtain σ1 σ2 x1' x2' e1' e2' where 
    h:"Γ \<turnstile> e : Data (DSum σ1 σ2)" and 
    h1:"(x1',Data σ1)#Γ \<turnstile> e1' : T" and 
    h2:"(x2',Data σ2)#Γ \<turnstile> e2' : T" and
    e1:"[x1].e1=[x1'].e1'" and e2:"[x2].e2=[x2'].e2'"
    by auto
  obtain c::name where f':"c \<sharp> (x1,x1',e1,e1',Γ)" by (erule exists_fresh[OF fs_name1])
  have e1':"[(x1,c)]•e1 = [(x1',c)]•e1'" using e1 f' by (auto simp add: alpha_fresh fresh_prod fresh_atm)
  have "[(x1',c)]•((x1',Data σ1)# Γ) \<turnstile> [(x1',c)]•e1' : T" using h1 typing_eqvt' by blast
  then have x:"(c,Data σ1)#( [(x1',c)]•Γ) \<turnstile> [(x1',c)]•e1': T" using f' 
    by (auto simp add: fresh_atm calc_atm)
  have "x1' \<sharp> Γ" using h1 typing_implies_valid by auto
  then have "(c,Data σ1)#Γ \<turnstile> [(x1 ,c)]•e1 : T" using f' x e1' by (auto simp add: perm_fresh_fresh)
  then have "[(x1,c)]•((c,Data σ1)#Γ) \<turnstile> [(x1,c)]•[(x1 ,c)]•e1 : T" using typing_eqvt' by blast 
  then have "([(x1,c)]•(c,Data σ1)) #Γ \<turnstile> [(x1,c)]•[(x1 ,c)]•e1 : T" using f f' 
    by (auto simp add: perm_fresh_fresh)
  then have "([(x1,c)]•(c,Data σ1)) #Γ \<turnstile> e1 : T" by perm_simp
  then have g1:"(x1, Data σ1)#Γ \<turnstile> e1 : T"  using f' by (auto simp add: fresh_atm calc_atm fresh_prod)
    (* The second part of the proof is the same *)
  obtain c::name where f':"c \<sharp> (x2,x2',e2,e2',Γ)" by (erule exists_fresh[OF fs_name1])
  have e2':"[(x2,c)]•e2 = [(x2',c)]•e2'" using e2 f' by (auto simp add: alpha_fresh fresh_prod fresh_atm)
  have "[(x2',c)]•((x2',Data σ2)# Γ) \<turnstile> [(x2',c)]•e2' : T" using h2 typing_eqvt' by blast
  then have x:"(c,Data σ2)#([(x2',c)]•Γ) \<turnstile> [(x2',c)]•e2': T" using f' 
    by (auto simp add: fresh_atm calc_atm)
  have "x2' \<sharp> Γ" using h2 typing_implies_valid by auto
  then have "(c,Data σ2)#Γ \<turnstile> [(x2 ,c)]•e2 : T" using f' x e2' by (auto simp add: perm_fresh_fresh)
  then have "[(x2,c)]•((c,Data σ2)#Γ) \<turnstile> [(x2,c)]•[(x2 ,c)]•e2 : T" using typing_eqvt' by blast 
  then have "([(x2,c)]•(c,Data σ2))#Γ \<turnstile> [(x2,c)]•[(x2 ,c)]•e2 : T" using f f' 
    by (auto simp add: perm_fresh_fresh)
  then have "([(x2,c)]•(c,Data σ2)) #Γ \<turnstile> e2 : T" by perm_simp
  then have g2:"(x2,Data σ2)#Γ \<turnstile> e2 : T"  using f' by (auto simp add: fresh_atm calc_atm fresh_prod)
  show ?thesis using g1 g2 prems by auto 
qed

lemma weakening: 
  fixes Γ1 Γ2 :: "(name×ty) list"
  assumes "Γ1 \<turnstile> e: T" and "valid Γ2" and "Γ1 ⊆ Γ2"
  shows "Γ2 \<turnstile> e: T"
  using assms
proof(nominal_induct Γ1 e T avoiding: Γ2 rule: typing.strong_induct)
  case (t_Lam x Γ1 T1 t T2 Γ2)
  have ih: "[|valid ((x,T1)#Γ2); (x,T1)#Γ1 ⊆ (x,T1)#Γ2|] ==> (x,T1)#Γ2 \<turnstile> t : T2" by fact
  have H1: "valid Γ2" by fact
  have H2: "Γ1 ⊆ Γ2" by fact
  have fs: "x\<sharp>Γ2" by fact
  then have "valid ((x,T1)#Γ2)" using H1 by auto
  moreover have "(x,T1)#Γ1 ⊆ (x,T1)#Γ2" using H2 by auto
  ultimately have "(x,T1)#Γ2 \<turnstile> t : T2" using ih by simp 
  thus "Γ2 \<turnstile> Lam [x].t : T1->T2" using fs by auto
next
  case (t_Case x1 Γ1 e e2 x2 e1 S1 S2 T Γ2)
  then have ih1: "valid ((x1,Data S1)#Γ2) ==> (x1,Data S1)#Γ2 \<turnstile> e1 : T" 
       and  ih2: "valid ((x2,Data S2)#Γ2) ==> (x2,Data S2)#Γ2 \<turnstile> e2 : T" 
       and  ih3: "Γ2 \<turnstile> e : Data (DSum S1 S2)" by auto 
  have fs1: "x1\<sharp>Γ2" "x1\<sharp>e" "x1\<sharp>e2" "x1\<sharp>x2" by fact+
  have fs2: "x2\<sharp>Γ2" "x2\<sharp>e" "x2\<sharp>e1" "x2\<sharp>x1" by fact+
  have "valid Γ2" by fact
  then have "valid ((x1,Data S1)#Γ2)" and "valid ((x2,Data S2)#Γ2)" using fs1 fs2 by auto
  then have "(x1, Data S1)#Γ2 \<turnstile> e1 : T" and "(x2, Data S2)#Γ2 \<turnstile> e2 : T" using ih1 ih2 by simp_all
  with ih3 show "Γ2 \<turnstile> Case e of inl x1 -> e1 | inr x2 -> e2 : T" using fs1 fs2 by auto
qed (auto)

lemma context_exchange:
  assumes a: "(x1,T1)#(x2,T2)#Γ \<turnstile> e : T"
  shows "(x2,T2)#(x1,T1)#Γ \<turnstile> e : T"
proof -
  from  a have "valid ((x1,T1)#(x2,T2)#Γ)" by (simp add: typing_implies_valid)
  then have "x1≠x2" "x1\<sharp>Γ" "x2\<sharp>Γ" "valid Γ"
    by (auto simp: fresh_list_cons fresh_atm[symmetric])
  then have "valid ((x2,T2)#(x1,T1)#Γ)"
    by (auto simp: fresh_list_cons fresh_atm)
  moreover 
  have "(x1,T1)#(x2,T2)#Γ ⊆ (x2,T2)#(x1,T1)#Γ" by auto
  ultimately show "(x2,T2)#(x1,T1)#Γ \<turnstile> e : T" using a by (auto intro: weakening)
qed

lemma typing_var_unicity: 
  assumes "(x,t1)#Γ \<turnstile> Var x : t2"
  shows "t1=t2"
proof - 
  have "(x,t2) ∈ set ((x,t1)#Γ)" and "valid ((x,t1)#Γ)" using assms by auto
  thus "t1=t2" by (simp only: type_unicity_in_context)
qed

lemma typing_substitution: 
  fixes Γ::"(name × ty) list"
  assumes "(x,T')#Γ \<turnstile> e : T" 
  and     "Γ \<turnstile> e': T'" 
  shows "Γ \<turnstile> e[x::=e'] : T"
  using assms
proof (nominal_induct e avoiding: Γ e' x arbitrary: T rule: trm.induct)
  case (Var y Γ e' x T)
  have h1: "(x,T')#Γ \<turnstile> Var y : T" by fact
  have h2: "Γ \<turnstile> e' : T'" by fact
  show "Γ \<turnstile> (Var y)[x::=e'] : T"
  proof (cases "x=y")
    case True
    assume as: "x=y"
    then have "T=T'" using h1 typing_var_unicity by auto
    then show "Γ \<turnstile> (Var y)[x::=e'] : T" using as h2 by simp
  next
    case False
    assume as: "x≠y" 
    have "(y,T) ∈ set ((x,T')#Γ)" using h1 by auto
    then have "(y,T) ∈ set Γ" using as by auto
    moreover 
    have "valid Γ" using h2 by (simp only: typing_implies_valid)
    ultimately show "Γ \<turnstile> (Var y)[x::=e'] : T" using as by (simp add: t_Var)
  qed
next
  case (Lam y t Γ e' x T)
  have vc: "y\<sharp>Γ" "y\<sharp>x" "y\<sharp>e'"  by fact+
  have pr1: "Γ \<turnstile> e' : T'" by fact
  have pr2: "(x,T')#Γ \<turnstile> Lam [y].t : T" by fact
  then obtain T1 T2 where pr2': "(y,T1)#(x,T')#Γ \<turnstile> t : T2" and eq: "T = T1->T2" 
    using vc by (auto simp add: fresh_list_cons)
  then have pr2'':"(x,T')#(y,T1)#Γ \<turnstile> t : T2" by (simp add: context_exchange)
  have ih: "[|(x,T')#(y,T1)#Γ \<turnstile> t : T2; (y,T1)#Γ \<turnstile> e' : T'|] ==> (y,T1)#Γ \<turnstile> t[x::=e'] : T2" by fact
  have "valid Γ" using pr1 by (simp add: typing_implies_valid)
  then have "valid ((y,T1)#Γ)" using vc by auto
  then have "(y,T1)#Γ \<turnstile> e' : T'" using pr1 by (auto intro: weakening)
  then have "(y,T1)#Γ \<turnstile> t[x::=e'] : T2" using ih pr2'' by simp
  then have "Γ \<turnstile> Lam [y].(t[x::=e']) : T1->T2" using vc by (auto intro: t_Lam)
  thus "Γ \<turnstile> (Lam [y].t)[x::=e'] : T" using vc eq by simp
next
  case (Case t1 x1 t2 x2 t3 Γ e' x T)
  have vc: "x1\<sharp>Γ" "x1\<sharp>e'" "x1\<sharp>x""x1\<sharp>t1" "x1\<sharp>t3" "x2\<sharp>Γ" 
           "x2\<sharp>e'" "x2\<sharp>x"  "x2\<sharp>t1" "x2\<sharp>t2" "x2≠x1" by fact+
  have as1: "Γ \<turnstile> e' : T'" by fact
  have as2: "(x,T')#Γ \<turnstile> Case t1 of inl x1 -> t2 | inr x2 -> t3 : T" by fact
  then obtain S1 S2 where 
    h1:"(x,T')#Γ \<turnstile> t1 : Data (DSum S1 S2)" and
    h2:"(x1,Data S1)#(x,T')#Γ \<turnstile> t2 : T" and
    h3:"(x2,Data S2)#(x,T')#Γ \<turnstile> t3 : T"
    using vc by (auto simp add: fresh_list_cons)
  have ih1: "[|(x,T')#Γ \<turnstile> t1 : Data (DSum S1 S2); Γ \<turnstile> e' : T'|] ==> Γ \<turnstile> t1[x::=e'] : Data (DSum S1 S2)"
  and  ih2: "[|(x,T')#(x1,Data S1)#Γ \<turnstile> t2:T; (x1,Data S1)#Γ \<turnstile> e':T'|] ==> (x1,Data S1)#Γ \<turnstile> t2[x::=e']:T" 
  and  ih3: "[|(x,T')#(x2,Data S2)#Γ \<turnstile> t3:T; (x2,Data S2)#Γ \<turnstile> e':T'|] ==> (x2,Data S2)#Γ \<turnstile> t3[x::=e']:T"
    by fact+
  from h2 have h2': "(x,T')#(x1,Data S1)#Γ \<turnstile> t2 : T" by (rule context_exchange)
  from h3 have h3': "(x,T')#(x2,Data S2)#Γ \<turnstile> t3 : T" by (rule context_exchange)
  have "Γ \<turnstile> t1[x::=e'] : Data (DSum S1 S2)" using h1 ih1 as1 by simp
  moreover
  have "valid ((x1,Data S1)#Γ)" using h2' by (auto dest: typing_implies_valid)
  then have "(x1,Data S1)#Γ \<turnstile> e' : T'" using as1 by (auto simp add: weakening)
  then have "(x1,Data S1)#Γ \<turnstile> t2[x::=e'] : T" using ih2 h2' by simp
  moreover 
  have "valid ((x2,Data S2)#Γ)" using h3' by (auto dest: typing_implies_valid)
  then have "(x2,Data S2)#Γ \<turnstile> e' : T'" using as1 by (auto simp add: weakening)
  then have "(x2,Data S2)#Γ \<turnstile> t3[x::=e'] : T" using ih3 h3' by simp
  ultimately have "Γ \<turnstile> Case (t1[x::=e']) of inl x1 -> (t2[x::=e']) | inr x2 -> (t3[x::=e']) : T"
    using vc by (auto simp add: fresh_atm fresh_subst)
  thus "Γ \<turnstile> (Case t1 of inl x1 -> t2 | inr x2 -> t3)[x::=e'] : T" using vc by simp
qed (simp, fast)+

text {* Big-Step Evaluation *}

inductive
  big :: "trm=>trm=>bool" ("_ \<Down> _" [80,80] 80) 
where
  b_Lam[intro]:   "Lam [x].e \<Down> Lam [x].e"
| b_App[intro]:   "[|x\<sharp>(e1,e2,e'); e1\<Down>Lam [x].e; e2\<Down>e2'; e[x::=e2']\<Down>e'|] ==> App e1 e2 \<Down> e'"
| b_Const[intro]: "Const n \<Down> Const n"
| b_Pr[intro]:    "[|e1\<Down>e1'; e2\<Down>e2'|] ==> Pr e1 e2 \<Down> Pr e1' e2'"
| b_Fst[intro]:   "e\<Down>Pr e1 e2 ==> Fst e\<Down>e1"
| b_Snd[intro]:   "e\<Down>Pr e1 e2 ==> Snd e\<Down>e2"
| b_InL[intro]:   "e\<Down>e' ==> InL e \<Down> InL e'"
| b_InR[intro]:   "e\<Down>e' ==> InR e \<Down> InR e'"
| b_CaseL[intro]: "[|x1\<sharp>(e,e2,e'',x2); x2\<sharp>(e,e1,e'',x1) ; e\<Down>InL e'; e1[x1::=e']\<Down>e''|] 
                   ==> Case e of inl x1 -> e1 | inr x2 -> e2 \<Down> e''"
| b_CaseR[intro]: "[|x1\<sharp>(e,e2,e'',x2); x2\<sharp>(e,e1,e'',x1) ; e\<Down>InR e'; e2[x2::=e']\<Down>e''|] 
                   ==> Case e of inl x1 -> e1 | inr x2 -> e2 \<Down> e''"

equivariance big

nominal_inductive big
  by (simp_all add: abs_fresh fresh_prod fresh_atm)

lemma big_eqvt':
  fixes pi::"name prm"
  assumes a: "(pi•t) \<Down> (pi•t')"
  shows "t \<Down> t'"
using a
apply -
apply(drule_tac pi="rev pi" in big.eqvt)
apply(perm_simp)
done

lemma fresh_preserved:
  fixes x::name
  fixes t::trm
  fixes t'::trm
  assumes "e \<Down> e'" and "x\<sharp>e" 
  shows "x\<sharp>e'"
  using assms by (induct) (auto simp add:fresh_subst')

declare trm.inject  [simp add]
declare ty.inject  [simp add]
declare data.inject [simp add]

inductive_cases b_inv_auto[elim]: 
  "App e1 e2 \<Down> t" 
  "Case e of inl x1 -> e1 | inr x2 -> e2 \<Down> t"
  "Lam[x].t \<Down> t"
  "Const n \<Down> t"
  "Fst e \<Down> t"
  "Snd e \<Down> t"
  "InL e \<Down> t"
  "InR e \<Down> t"
  "Pr e1 e2 \<Down> t"

declare trm.inject  [simp del]
declare ty.inject  [simp del]
declare data.inject [simp del]

lemma b_App_elim[elim]:
  assumes "App e1 e2 \<Down> e'" and "x\<sharp>(e1,e2,e')"
  obtains f1 and f2 where "e1 \<Down> Lam [x]. f1" "e2 \<Down> f2" "f1[x::=f2] \<Down> e'"
  using assms
  apply -
  apply(erule b_inv_auto)
  apply(drule_tac pi="[(xa,x)]" in big.eqvt)
  apply(drule_tac pi="[(xa,x)]" in big.eqvt)
  apply(drule_tac pi="[(xa,x)]" in big.eqvt)
  apply(perm_simp add: calc_atm eqvts)
  done

lemma  b_CaseL_elim[elim]: 
  assumes "Case e of inl x1 -> e1 | inr x2 -> e2 \<Down> e''" 
  and     "!! t. ¬  e \<Down> InR t"
  and     "x1\<sharp>e''" "x1\<sharp>e" "x2\<sharp>e''" "x1\<sharp>e"
  obtains e' where "e \<Down> InL e'" and "e1[x1::=e'] \<Down> e''"
  using assms 
  apply -
  apply(rule b_inv_auto(2))
  apply(auto)
  apply(simp add: alpha)
  apply(auto)
  apply(drule_tac x="[(x1,x1')]•e'" in meta_spec)
  apply(drule meta_mp)
  apply(rule_tac pi="[(x1,x1')]" in big_eqvt')
  apply(perm_simp add: fresh_prod)
  apply(drule meta_mp)
  apply(rule_tac pi="[(x1,x1')]" in big_eqvt')
  apply(perm_simp add: eqvts calc_atm)
  apply(assumption)
  apply(drule_tac x="[(x1,x1')]•e'" in meta_spec)
  apply(drule meta_mp)
  apply(rule_tac pi="[(x1,x1')]" in big_eqvt')
  apply(perm_simp add: fresh_prod)
  apply(drule meta_mp)
  apply(rule_tac pi="[(x1,x1')]" in big_eqvt')
  apply(perm_simp add: eqvts calc_atm)
  apply(assumption)
done

lemma b_CaseR_elim[elim]: 
  assumes "Case e of inl x1 -> e1 | inr x2 -> e2 \<Down> e''" 
  and     "!! t. ¬ e \<Down> InL t"
  and     "x1\<sharp>e''" "x1\<sharp>e" "x2\<sharp>e''" "x2\<sharp>e"
  obtains e' where "e \<Down> InR e'" and "e2[x2::=e'] \<Down> e''"
  using assms 
  apply -
  apply(rule b_inv_auto(2))
  apply(auto)
  apply(simp add: alpha)
  apply(auto)
  apply(drule_tac x="[(x2,x2')]•e'" in meta_spec)
  apply(drule meta_mp)
  apply(rule_tac pi="[(x2,x2')]" in big_eqvt')
  apply(perm_simp add: fresh_prod)
  apply(drule meta_mp)
  apply(rule_tac pi="[(x2,x2')]" in big_eqvt')
  apply(perm_simp add: eqvts calc_atm)
  apply(assumption)
  apply(drule_tac x="[(x2,x2')]•e'" in meta_spec)
  apply(drule meta_mp)
  apply(rule_tac pi="[(x2,x2')]" in big_eqvt')
  apply(perm_simp add: fresh_prod)
  apply(drule meta_mp)
  apply(rule_tac pi="[(x2,x2')]" in big_eqvt')
  apply(perm_simp add: eqvts calc_atm)
  apply(assumption)
done

inductive
  val :: "trm=>bool" 
where
  v_Lam[intro]:   "val (Lam [x].e)"
| v_Const[intro]: "val (Const n)"
| v_Pr[intro]:    "[|val e1; val e2|] ==> val (Pr e1 e2)"
| v_InL[intro]:   "val e ==> val (InL e)"
| v_InR[intro]:   "val e ==> val (InR e)"

declare trm.inject  [simp add]
declare ty.inject  [simp add]
declare data.inject [simp add]

inductive_cases v_inv_auto[elim]: 
  "val (Const n)"
  "val (Pr e1 e2)"
  "val (InL e)"
  "val (InR e)"
  "val (Fst e)"
  "val (Snd e)"
  "val (Case e of inl x1 -> e1 | inr x2 -> e2)"
  "val (Var x)"
  "val (Lam [x].e)"
  "val (App e1 e2)"

declare trm.inject  [simp del]
declare ty.inject  [simp del]
declare data.inject [simp del]

lemma subject_reduction:
  assumes a: "e \<Down> e'" 
  and     b: "Γ \<turnstile> e : T"
  shows "Γ \<turnstile> e' : T"
  using a b
proof (nominal_induct avoiding: Γ arbitrary: T rule: big.strong_induct) 
  case (b_App x e1 e2 e' e e2' Γ T)
  have vc: "x\<sharp>Γ" by fact
  have "Γ \<turnstile> App e1 e2 : T" by fact
  then obtain T' where 
    a1: "Γ \<turnstile> e1 : T'->T" and  
    a2: "Γ \<turnstile> e2 : T'" by auto
  have ih1: "Γ \<turnstile> e1 : T' -> T ==> Γ \<turnstile> Lam [x].e : T' -> T" by fact
  have ih2: "Γ \<turnstile> e2 : T' ==> Γ \<turnstile> e2' : T'" by fact 
  have ih3: "Γ \<turnstile> e[x::=e2'] : T ==> Γ \<turnstile> e' : T" by fact
  have "Γ \<turnstile> Lam [x].e : T'->T" using ih1 a1 by simp 
  then have "((x,T')#Γ) \<turnstile> e : T" using vc by (auto simp add: ty.inject)
  moreover
  have "Γ \<turnstile> e2': T'" using ih2 a2 by simp
  ultimately have "Γ \<turnstile> e[x::=e2'] : T" by (simp add: typing_substitution)
  thus "Γ \<turnstile> e' : T" using ih3 by simp
next
  case (b_CaseL x1 e e2 e'' x2 e1 e' Γ)
  have vc: "x1\<sharp>Γ" "x2\<sharp>Γ" by fact+
  have "Γ \<turnstile> Case e of inl x1 -> e1 | inr x2 -> e2 : T" by fact 
  then obtain S1 S2 e1' e2' where 
    a1: "Γ \<turnstile> e : Data (DSum S1 S2)" and 
    a2: "((x1,Data S1)#Γ) \<turnstile> e1 : T" using vc by auto 
  have ih1:"Γ \<turnstile> e : Data (DSum S1 S2) ==> Γ \<turnstile> InL e' : Data (DSum S1 S2)" by fact 
  have ih2:"Γ \<turnstile> e1[x1::=e'] : T ==> Γ \<turnstile> e'' : T " by fact 
  have "Γ \<turnstile> InL e' : Data (DSum S1 S2)" using ih1 a1 by simp
  then have "Γ \<turnstile> e' : Data S1" by auto
  then have "Γ \<turnstile> e1[x1::=e'] : T" using a2 by (simp add: typing_substitution)
  then show "Γ \<turnstile> e'' : T" using ih2 by simp
next
 case (b_CaseR x1 e e2 e'' x2 e1 e' Γ T)
 then show "Γ \<turnstile> e'' : T" by (blast intro: typing_substitution)
qed (blast)+

lemma unicity_of_evaluation:
  assumes a: "e \<Down> e1" 
  and     b: "e \<Down> e2"
  shows "e1 = e2"
  using a b
proof (nominal_induct e e1 avoiding: e2 rule: big.strong_induct)
  case (b_Lam x e t2)
  have "Lam [x].e \<Down> t2" by fact
  thus "Lam [x].e = t2" by (cases, simp_all add: trm.inject)
next
  case (b_App x e1 e2 e' e1' e2' t2)
  have ih1: "!!t. e1 \<Down> t ==> Lam [x].e1' = t" by fact
  have ih2:"!!t. e2 \<Down> t ==> e2' = t" by fact
  have ih3: "!!t. e1'[x::=e2'] \<Down> t ==> e' = t" by fact
  have app: "App e1 e2 \<Down> t2" by fact
  have vc: "x\<sharp>e1" "x\<sharp>e2" by fact+
  then have "x \<sharp> App e1 e2" by auto
  then have vc': "x\<sharp>t2" using fresh_preserved app by blast
  from vc vc' obtain f1 f2 where x1: "e1 \<Down> Lam [x]. f1" and x2: "e2 \<Down> f2" and x3: "f1[x::=f2] \<Down> t2"
    using app by (auto simp add: fresh_prod)
  then have "Lam [x]. f1 = Lam [x]. e1'" using ih1 by simp
  then 
  have "f1 = e1'" by (auto simp add: trm.inject alpha) 
  moreover 
  have "f2 = e2'" using x2 ih2 by simp
  ultimately have "e1'[x::=e2'] \<Down> t2" using x3 by simp
  thus ?case using ih3 by simp
next
  case (b_CaseL  x1 e e2 e'' x2 e1 e' t2)
  have fs: "x1\<sharp>e" "x1\<sharp>t2" "x2\<sharp>e" "x2\<sharp>t2" by fact+
  have ih1:"!!t. e \<Down> t ==> InL e' = t" by fact 
  have ih2:"!!t. e1[x1::=e'] \<Down> t ==> e'' = t" by fact
  have ha: "¬(∃t. e \<Down> InR t)" using ih1 by force
  have "Case e of inl x1 -> e1 | inr x2 -> e2 \<Down> t2" by fact
  then obtain f' where "e \<Down> InL f'" and h: "e1[x1::=f']\<Down>t2" using ha fs by auto
  then have "InL f' = InL e'" using ih1 by simp
  then have "f' = e'" by (simp add: trm.inject)
  then have "e1[x1::=e'] \<Down> t2" using h by simp
  then show "e'' = t2" using ih2 by simp
next 
  case (b_CaseR x1 e e2 e'' x2 e1 e' t2 )
  have fs: "x1\<sharp>e" "x1\<sharp>t2" "x2\<sharp>e" "x2\<sharp>t2" by fact+
  have ih1: "!!t. e \<Down> t ==> InR e' = t" by fact
  have ih2: "!!t. e2[x2::=e'] \<Down> t ==> e'' = t" by fact
  have ha: "¬(∃t. e \<Down> InL t)" using ih1 by force
  have "Case e of inl x1 -> e1 | inr x2 -> e2 \<Down> t2" by fact
  then obtain f' where "e \<Down> InR f'" and h: "e2[x2::=f']\<Down>t2"  using ha fs by auto
  then have "InR f' = InR e'" using ih1 by simp
  then have "e2[x2::=e'] \<Down> t2" using h by (simp add: trm.inject)
  thus "e'' = t2" using ih2 by simp
next
  case (b_Fst e e1 e2 e2')
  have "e \<Down> Pr e1 e2" by fact
  have "!! b. e \<Down> b ==> Pr e1 e2 = b" by fact
  have "Fst e \<Down> e2'" by fact
  show "e1 = e2'" using prems by (force simp add: trm.inject)
next
  case (b_Snd e e1 e2 e2')
  have "e \<Down> Pr e1 e2" by fact
  have "!! b. e \<Down> b ==> Pr e1 e2 = b" by fact
  have "Snd e \<Down> e2'" by fact
  show "e2 = e2'" using prems by (force simp add: trm.inject)
qed (blast)+

lemma not_val_App[simp]:
  shows 
  "¬ val (App e1 e2)" 
  "¬ val (Fst e)"
  "¬ val (Snd e)"
  "¬ val (Var x)"
  "¬ val (Case e of inl x1 -> e1 | inr x2 -> e2)"
by auto

lemma reduces_evaluates_to_values:
  assumes h:"t \<Down> t'"
  shows "val t'"
  using h by (induct) (auto)

lemma type_prod_evaluates_to_pairs:
  assumes a: "Γ \<turnstile> t : Data (DProd S1 S2)" 
  and     b: "t \<Down> t'"
  obtains t1 t2 where "t' = Pr t1 t2"
proof -
   have "Γ \<turnstile> t' : Data (DProd S1 S2)" using assms subject_reduction by simp
   moreover
   have "val t'" using reduces_evaluates_to_values assms by simp
   ultimately obtain t1 t2 where "t' = Pr t1 t2" by (cases, auto simp add:ty.inject data.inject)
   thus ?thesis using prems by auto
qed

lemma type_sum_evaluates_to_ins:
  assumes "Γ \<turnstile> t : Data (DSum σ1 σ2)" and "t \<Down> t'"
  shows "(∃t''. t' = InL t'') ∨ (∃t''. t' = InR t'')"
proof -
  have "Γ \<turnstile> t' : Data (DSum σ1 σ2)" using assms subject_reduction by simp
  moreover
  have "val t'" using reduces_evaluates_to_values assms by simp
  ultimately obtain t'' where "t' = InL t'' ∨  t' = InR t''"
    by (cases, auto simp add:ty.inject data.inject)
  thus ?thesis by auto
qed

lemma type_arrow_evaluates_to_lams:
  assumes "Γ \<turnstile> t : σ -> τ" and "t \<Down> t'"
  obtains  x t'' where "t' = Lam [x]. t''"
proof -
  have "Γ \<turnstile> t' : σ -> τ" using assms subject_reduction by simp
  moreover
  have "val t'" using reduces_evaluates_to_values assms by simp
  ultimately obtain x t'' where "t' = Lam [x]. t''" by (cases, auto simp add:ty.inject data.inject)
  thus ?thesis using prems by auto
qed

lemma type_nat_evaluates_to_consts:
  assumes "Γ \<turnstile> t : Data DNat" and "t \<Down> t'"
  obtains n where "t' = Const n"
proof -
  have "Γ \<turnstile> t' : Data DNat " using assms subject_reduction by simp
  moreover have "val t'" using reduces_evaluates_to_values assms by simp
  ultimately obtain n where "t' = Const n" by (cases, auto simp add:ty.inject data.inject)
  thus ?thesis using prems by auto
qed

consts
  V' :: "data => trm set" 

nominal_primrec    
  "V' (DNat) = {Const n | n. n ∈ (UNIV::nat set)}"
  "V' (DProd S1 S2) = {Pr x y | x y. x ∈ V' S1 ∧ y ∈ V' S2}"
  "V' (DSum S1 S2) = {InL x | x. x ∈ V' S1} ∪ {InR y | y. y ∈ V' S2}"
apply(rule TrueI)+
done

lemma Vprimes_are_values :
  fixes S::"data"
  assumes h: "e ∈ V' S"
  shows "val e"
using h 
by (nominal_induct S arbitrary: e rule:data.induct)
   (auto)

lemma V'_eqvt:
  fixes pi::"name prm"
  assumes a: "v ∈ V' S"
  shows "(pi•v) ∈ V' S"
using a
by (nominal_induct S arbitrary: v rule: data.induct)
   (auto simp add: trm.inject)

consts
  V :: "ty => trm set" 

nominal_primrec
  "V (Data S) = V' S"
  "V (T1 -> T2) = {Lam [x].e | x e. ∀ v ∈ (V T1). ∃ v'. e[x::=v] \<Down> v' ∧ v' ∈ V T2}"
apply(rule TrueI)+ 
done

lemma V_eqvt:
  fixes pi::"name prm"
  assumes a: "x∈V T"
  shows "(pi•x)∈V T"
using a
apply(nominal_induct T arbitrary: pi x rule: ty.induct)
apply(auto simp add: trm.inject perm_set_def)
apply(perm_simp add: V'_eqvt)
apply(rule_tac x="pi•xa" in exI)
apply(rule_tac x="pi•e" in exI)
apply(simp)
apply(auto)
apply(drule_tac x="(rev pi)•v" in bspec)
apply(force)
apply(auto)
apply(rule_tac x="pi•v'" in exI)
apply(auto)
apply(drule_tac pi="pi" in big.eqvt)
apply(perm_simp add: eqvts)
done

lemma V_arrow_elim_weak[elim] :
  assumes h:"u ∈ (V (T1 -> T2))"
  obtains a t where "u = Lam[a].t" and "∀ v ∈ (V T1). ∃ v'. t[a::=v] \<Down> v' ∧ v' ∈ V T2"
using h by (auto)

lemma V_arrow_elim_strong[elim]:
  fixes c::"'a::fs_name"
  assumes h: "u ∈ V (T1 -> T2)"
  obtains a t where "a\<sharp>c" "u = Lam[a].t" "∀v ∈ (V T1). ∃ v'. t[a::=v] \<Down> v' ∧ v' ∈ V T2"
using h
apply -
apply(erule V_arrow_elim_weak)
apply(subgoal_tac "∃a'::name. a'\<sharp>(a,t,c)") (*A*)
apply(erule exE)
apply(drule_tac x="a'" in meta_spec)
apply(drule_tac x="[(a,a')]•t" in meta_spec)
apply(drule meta_mp)
apply(simp)
apply(drule meta_mp)
apply(simp add: trm.inject alpha fresh_left fresh_prod calc_atm fresh_atm)
apply(perm_simp)
apply(force)
apply(drule meta_mp)
apply(rule ballI)
apply(drule_tac x="[(a,a')]•v" in bspec)
apply(simp add: V_eqvt)
apply(auto)
apply(rule_tac x="[(a,a')]•v'" in exI)
apply(auto)
apply(drule_tac pi="[(a,a')]" in big.eqvt)
apply(perm_simp add: eqvts calc_atm)
apply(simp add: V_eqvt)
(*A*)
apply(rule exists_fresh')
apply(simp add: fin_supp)
done

lemma V_are_values :
  fixes T::"ty"
  assumes h:"e ∈ V T"
  shows "val e"
using h by (nominal_induct T arbitrary: e rule:ty.induct, auto simp add: Vprimes_are_values)

lemma values_reduce_to_themselves:
  assumes h:"val v"
  shows "v \<Down> v"
using h by (induct,auto)

lemma Vs_reduce_to_themselves[simp]:
  assumes h:"v ∈ V T"
  shows "v \<Down> v"
using h by (simp add: values_reduce_to_themselves V_are_values)

lemma V_sum:
  assumes h:"x ∈ V (Data (DSum S1 S2))"
  shows "(∃ y. x= InL y ∧  y ∈ V' S1) ∨ (∃ y. x= InR y ∧ y ∈ V' S2)"
using h by simp

abbreviation 
 mapsto :: "(name×trm) list => name => trm => bool" ("_ maps _ to _" [55,55,55] 55) 
where
 "ϑ maps x to e≡ (lookup ϑ x) = e"

abbreviation 
  v_closes :: "(name×trm) list => (name×ty) list => bool" ("_ Vcloses _" [55,55] 55) 
where
  "ϑ Vcloses Γ ≡ ∀x T. ((x,T) ∈ set Γ --> (∃v. ϑ maps x to v ∧ v ∈ (V T)))"

lemma monotonicity:
  fixes m::"name"
  fixes ϑ::"(name × trm) list" 
  assumes h1: "ϑ Vcloses Γ"
  and     h2: "e ∈ V T" 
  and     h3: "valid ((x,T)#Γ)"
  shows "(x,e)#ϑ Vcloses (x,T)#Γ"
proof(intro strip)
  fix x' T'
  assume "(x',T') ∈ set ((x,T)#Γ)"
  then have "((x',T')=(x,T)) ∨ ((x',T')∈set Γ ∧ x'≠x)" using h3 
    by (rule_tac case_distinction_on_context)
  moreover
  { (* first case *)
    assume "(x',T') = (x,T)"
    then have "∃e'. ((x,e)#ϑ) maps x to e' ∧ e' ∈ V T'" using h2 by auto
  }
  moreover
  { (* second case *)
    assume "(x',T') ∈ set Γ" and neq:"x' ≠ x"
      then have "∃e'. ϑ maps x' to e' ∧ e' ∈ V T'" using h1 by auto
      then have "∃e'. ((x,e)#ϑ) maps x' to e' ∧ e' ∈ V T'" using neq by auto
  }
  ultimately show "∃e'.  ((x,e)#ϑ) maps x' to e'  ∧ e' ∈ V T'" by blast
qed

lemma termination_aux:
  fixes T :: "ty"
  fixes Γ :: "(name × ty) list"
  fixes ϑ :: "(name × trm) list"
  fixes e :: "trm"
  assumes h1: "Γ \<turnstile> e : T"
  and     h2: "ϑ Vcloses Γ"
  shows "∃v. ϑ<e> \<Down> v ∧ v ∈ V T" 
using h2 h1
proof(nominal_induct e avoiding: Γ ϑ arbitrary: T rule: trm.induct)
  case (App e1 e2 Γ ϑ T)
  have ih1:"!!ϑ Γ T. [|ϑ Vcloses Γ; Γ \<turnstile> e1 : T|] ==> ∃v. ϑ<e1> \<Down> v ∧ v ∈ V T" by fact
  have ih2:"!!ϑ Γ T. [|ϑ Vcloses Γ; Γ \<turnstile> e2 : T|] ==> ∃v. ϑ<e2> \<Down> v ∧ v ∈ V T" by fact
  have as1:"ϑ Vcloses Γ" by fact 
  have as2: "Γ \<turnstile> App e1 e2 : T" by fact
  from as2 obtain T' where "Γ \<turnstile> e1 : T' -> T" and "Γ \<turnstile> e2 : T'" by auto
  then obtain v1 v2 where "(i)": "ϑ<e1> \<Down> v1" "v1 ∈ V (T' -> T)"
                      and "(ii)":"ϑ<e2> \<Down> v2" "v2 ∈ V T'" using ih1 ih2 as1 by blast
  from "(i)" obtain x e' 
            where "v1 = Lam[x].e'" 
            and "(iii)": "(∀v ∈ (V T').∃ v'. e'[x::=v] \<Down> v' ∧ v' ∈ V T)"
            and "(iv)":  "ϑ<e1> \<Down> Lam [x].e'" 
            and fr: "x\<sharp>(ϑ,e1,e2)" by blast
  from fr have fr1: "x\<sharp>ϑ<e1>" and fr2: "x\<sharp>ϑ<e2>" by (simp_all add: fresh_psubst)
  from "(ii)" "(iii)" obtain v3 where "(v)": "e'[x::=v2] \<Down> v3" "v3 ∈ V T" by auto
  from fr2 "(ii)" have "x\<sharp>v2" by (simp add: fresh_preserved)
  then have "x\<sharp>e'[x::=v2]" by (simp add: fresh_subst_fresh)
  then have fr3: "x\<sharp>v3" using "(v)" by (simp add: fresh_preserved)
  from fr1 fr2 fr3 have "x\<sharp>(ϑ<e1>,ϑ<e2>,v3)" by simp
  with "(iv)" "(ii)" "(v)" have "App (ϑ<e1>) (ϑ<e2>) \<Down> v3" by auto
  then show "∃v. ϑ<App e1 e2> \<Down> v ∧ v ∈ V T" using "(v)" by auto
next
  case (Pr t1 t2 Γ ϑ T)
  have "Γ \<turnstile> Pr t1 t2 : T" by fact
  then obtain Ta Tb where ta:"Γ \<turnstile> t1 : Data Ta" and "Γ \<turnstile> t2 : Data Tb" 
                      and eq:"T=Data (DProd Ta Tb)" by auto
  have h:"ϑ Vcloses Γ" by fact
  then obtain v1 v2 where "ϑ<t1> \<Down> v1 ∧ v1 ∈ V (Data Ta)" "ϑ<t2> \<Down> v2 ∧ v2 ∈ V (Data Tb)" 
    using prems by blast
  thus "∃v. ϑ<Pr t1 t2> \<Down> v ∧ v ∈ V T" using eq by auto
next 
  case (Lam x e Γ ϑ T)
  have ih:"!!ϑ Γ T. [|ϑ Vcloses Γ; Γ \<turnstile> e : T|] ==> ∃v. ϑ<e> \<Down> v ∧ v ∈ V T" by fact
  have as1: "ϑ Vcloses Γ" by fact
  have as2: "Γ \<turnstile> Lam [x].e : T" by fact
  have fs: "x\<sharp>Γ" "x\<sharp>ϑ" by fact+
  from as2 fs obtain T1 T2 
    where "(i)": "(x,T1)#Γ \<turnstile> e:T2" and "(ii)": "T = T1 -> T2" by auto
  from "(i)" have "(iii)": "valid ((x,T1)#Γ)" by (simp add: typing_implies_valid)
  have "∀v ∈ (V T1). ∃v'. (ϑ<e>)[x::=v] \<Down> v' ∧ v' ∈ V T2"
  proof
    fix v
    assume "v ∈ (V T1)"
    with "(iii)" as1 have "(x,v)#ϑ Vcloses (x,T1)#Γ" using monotonicity by auto
    with ih "(i)" obtain v' where "((x,v)#ϑ)<e> \<Down> v' ∧ v' ∈ V T2" by blast
    then have "ϑ<e>[x::=v] \<Down> v' ∧ v' ∈ V T2" using fs by (simp add: psubst_subst_psubst)
    then show "∃v'. ϑ<e>[x::=v] \<Down> v' ∧ v' ∈ V T2" by auto
  qed
  then have "Lam[x].ϑ<e> ∈ V (T1 -> T2)" by auto
  then have "ϑ<Lam [x].e> \<Down> Lam[x].ϑ<e> ∧ Lam[x].ϑ<e> ∈ V (T1->T2)" using fs by auto
  thus "∃v. ϑ<Lam [x].e> \<Down> v ∧ v ∈ V T" using "(ii)" by auto
next
  case (Case t' n1 t1 n2 t2 Γ ϑ T)
  have f: "n1\<sharp>Γ" "n1\<sharp>ϑ" "n2\<sharp>Γ" "n2\<sharp>ϑ" "n2≠n1" "n1\<sharp>t'"
  "n1\<sharp>t2" "n2\<sharp>t'" "n2\<sharp>t1" by fact+
  have h:"ϑ Vcloses Γ" by fact
  have th:"Γ \<turnstile> Case t' of inl n1 -> t1 | inr n2 -> t2 : T" by fact
  then obtain S1 S2 where 
    hm:"Γ \<turnstile> t' : Data (DSum S1 S2)" and
    hl:"(n1,Data S1)#Γ \<turnstile> t1 : T" and
    hr:"(n2,Data S2)#Γ \<turnstile> t2 : T" using f by auto
  then obtain v0 where ht':"ϑ<t'> \<Down> v0" and hS:"v0 ∈ V (Data (DSum S1 S2))" using prems h by blast
  (* We distinguish between the cases InL and InR *)
  { fix v0'
    assume eqc:"v0 = InL v0'" and "v0' ∈ V' S1"
    then have inc: "v0' ∈ V (Data S1)" by auto
    have "valid Γ" using th typing_implies_valid by auto
    then moreover have "valid ((n1,Data S1)#Γ)" using f by auto
    then moreover have "(n1,v0')#ϑ Vcloses (n1,Data S1)#Γ" 
      using inc h monotonicity by blast
    moreover 
    have ih:"!!Γ ϑ T. [|ϑ Vcloses Γ; Γ \<turnstile> t1 : T|] ==> ∃v. ϑ<t1> \<Down> v ∧ v ∈ V T" by fact
    ultimately obtain v1 where ho: "((n1,v0')#ϑ)<t1> \<Down> v1 ∧ v1 ∈ V T" using hl by blast
    then have r:"ϑ<t1>[n1::=v0'] \<Down> v1 ∧ v1 ∈ V T" using psubst_subst_psubst f by simp
    then moreover have "n1\<sharp>(ϑ<t'>,ϑ<t2>,v1,n2)" 
      proof -
        have "n1\<sharp>v0" using ht' fresh_preserved fresh_psubst f by auto 
        then have "n1\<sharp>v0'" using eqc by auto
        then have "n1\<sharp>v1" using f r fresh_preserved fresh_subst_fresh by blast
        thus "n1\<sharp>(ϑ<t'>,ϑ<t2>,v1,n2)" using f by (simp add: fresh_atm fresh_psubst)
      qed
    moreover have "n2\<sharp>(ϑ<t'>,ϑ<t1>,v1,n1)"
      proof -
        have "n2\<sharp>v0" using ht' fresh_preserved fresh_psubst f by auto
        then have "n2\<sharp>v0'" using eqc by auto
        then have "n2\<sharp>((n1,v0')#ϑ)" using f by (simp add: fresh_list_cons fresh_prod fresh_atm)
        then have "n2\<sharp>((n1,v0')#ϑ)<t1>" using f fresh_psubst by auto
        moreover then have "n2 \<sharp> v1" using fresh_preserved ho by auto
        ultimately show  "n2\<sharp>(ϑ<t'>,ϑ<t1>,v1,n1)" using f by (simp add: fresh_psubst fresh_atm)
      qed
    ultimately have "Case ϑ<t'> of inl n1 -> ϑ<t1> | inr n2 -> ϑ<t2> \<Down> v1 ∧ v1 ∈ V T" using ht' eqc by auto
    moreover 
      have "Case ϑ<t'> of inl n1 -> ϑ<t1> | inr n2 -> ϑ<t2> = ϑ<Case t' of inl n1 -> t1 | inr n2 -> t2>" 
      using f by auto
    ultimately have "∃v. ϑ<Case t' of inl n1 -> t1 | inr n2 -> t2> \<Down> v ∧ v ∈ V T" by auto
  }
  moreover 
  { fix v0'
    assume eqc:"v0 = InR v0'" and "v0' ∈ V' S2"
    then have inc:"v0' ∈ V (Data S2)" by auto
    have "valid Γ" using th typing_implies_valid by auto
    then moreover have "valid ((n2,Data S2)#Γ)" using f by auto
    then moreover have "(n2,v0')#ϑ Vcloses (n2,Data S2)#Γ" 
      using inc h monotonicity by blast
    moreover have ih:"!!Γ ϑ T. [|ϑ Vcloses Γ; Γ \<turnstile> t2 : T|] ==> ∃v. ϑ<t2> \<Down> v ∧ v ∈ V T" by fact
    ultimately obtain v2 where ho:"((n2,v0')#ϑ)<t2> \<Down> v2 ∧ v2 ∈ V T" using hr by blast
    then have r:"ϑ<t2>[n2::=v0'] \<Down> v2 ∧ v2 ∈ V T" using psubst_subst_psubst f by simp
    moreover have "n1\<sharp>(ϑ<t'>,ϑ<t2>,v2,n2)"
    proof -
      have "n1\<sharp>ϑ<t'>" using fresh_psubst f by simp
      then have "n1\<sharp>v0" using ht' fresh_preserved by auto
      then have "n1\<sharp>v0'" using eqc by auto
      then have "n1\<sharp>((n2,v0')#ϑ)" using f by (simp add: fresh_list_cons fresh_atm fresh_prod)
      then have "n1\<sharp>((n2,v0')#ϑ)<t2>" using f fresh_psubst by auto
      moreover then have "n1\<sharp>v2" using fresh_preserved ho by auto
      ultimately show  "n1 \<sharp> (ϑ<t'>,ϑ<t2>,v2,n2)" using f by (simp add: fresh_psubst fresh_atm)
    qed
    moreover have "n2 \<sharp> (ϑ<t'>,ϑ<t1>,v2,n1)"
      proof -
        have "n2\<sharp>ϑ<t'>" using fresh_psubst f by simp
        then have "n2\<sharp>v0" using ht' fresh_preserved by auto
        then have "n2\<sharp>v0'" using eqc by auto
        then have "n2\<sharp>ϑ<t2>[n2::=v0']" using f fresh_subst_fresh by auto
        then have "n2\<sharp>v2" using f fresh_preserved r by blast
        then show "n2\<sharp>(ϑ<t'>,ϑ<t1>,v2,n1)" using f by (simp add: fresh_atm fresh_psubst)
      qed
    ultimately have "Case ϑ<t'> of inl n1 -> ϑ<t1> | inr n2 -> ϑ<t2> \<Down> v2 ∧ v2 ∈ V T" using ht' eqc by auto
    then have "∃v. ϑ<Case t' of inl n1 -> t1 | inr n2 -> t2> \<Down> v ∧ v ∈ V T" using f by auto
  }
  ultimately show "∃v. ϑ<Case t' of inl n1 -> t1 | inr n2 -> t2> \<Down> v ∧ v ∈ V T" using hS V_sum by blast
qed (force)+

theorem termination_of_evaluation:
  assumes a: "[] \<turnstile> e : T"
  shows "∃v. e \<Down> v ∧ val v"
proof -
  from a have "∃v. (([]::(name × trm) list)<e>) \<Down> v ∧ v ∈ V T" 
    by (rule termination_aux) (auto)
  thus "∃v. e \<Down> v ∧ val v" using V_are_values by auto
qed 

end

lemma in_eqvt:

  xX ==> pixpiX

lemma perm_data:

  piD = D

lemma perm_ty:

  piT = T

lemma fresh_ty:

  x \<sharp> T

lemma lookup_eqvt:

  pilookup ϑ X = lookup (piϑ) (piX)

lemma lookup_fresh:

  [| z \<sharp> ϑ; z \<sharp> x |] ==> z \<sharp> lookup ϑ x

lemma lookup_fresh':

  z \<sharp> ϑ ==> lookup ϑ z = Var z

lemma psubst_eqvt:

  piϑ<t> = (piϑ)<(pit)>

lemma fresh_psubst:

  [| z \<sharp> t; z \<sharp> ϑ |] ==> z \<sharp> ϑ<t>

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):

  Pr e1 e2[y::=t'] = Pr (e1[y::=t']) (e2[y::=t'])

and subst(6):

  Fst e[y::=t'] = Fst (e[y::=t'])

and subst(7):

  Snd e[y::=t'] = Snd (e[y::=t'])

and subst(8):

  InL e[y::=t'] = InL (e[y::=t'])

and subst(9):

  InR e[y::=t'] = InR (e[y::=t'])

and subst(10):

  [| z  x; x \<sharp> (y, e, e2, t'); z \<sharp> (y, e, e1, 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']

lemma subst_eqvt:

  pit[x::=t'] = (pit)[(pix)::=(pit')]

lemma fresh_subst:

  [| z \<sharp> t1; z \<sharp> t2 |] ==> z \<sharp> t1[y::=t2]

lemma fresh_subst':

  [| z \<sharp> [y].t1; z \<sharp> t2 |] ==> z \<sharp> t1[y::=t2]

lemma forget:

  x \<sharp> L ==> L[x::=P] = L

lemma psubst_empty:

  []<t> = t

lemma psubst_subst_psubst:

  x \<sharp> ϑ ==> ϑ<e>[x::=e'] = ((x, e') # ϑ)<e>

lemma fresh_subst_fresh:

  a \<sharp> e ==> a \<sharp> t[a::=e]

lemma type_unicity_in_context:

  [| (x, t2) ∈ set ((x, t1) # Γ); valid ((x, t1) # Γ) |] ==> t1 = t2

lemma case_distinction_on_context:

  [| valid ((m, t) # Γ); (n, U) ∈ set ((m, T) # Γ) |]
  ==> (n, U) = (m, T) ∨ (n, U) ∈ set Γn  m

lemma typing_eqvt':

  x1.0 \<turnstile> x2.0 : x3.0 ==> pix1.0 \<turnstile> pix2.0 : x3.0

lemma typing_implies_valid:

  Γ \<turnstile> t : T ==> valid Γ

lemma t_Lam_elim:

  [| Γ \<turnstile> Lam [x].t : T; x \<sharp> Γ;
     !!T1 T2. [| (x, T1) # Γ \<turnstile> t : T2; T = T1->T2 |] ==> thesis |]
  ==> thesis

lemma t_Case_elim:

  [| Γ \<turnstile> Case e of inl x1 -> e1 | inr x2 -> e2 : T; x1 \<sharp> Γ;
     x2 \<sharp> Γ;
     !!σ1 σ2.
        [| Γ \<turnstile> e : Data (DSum σ1 σ2);
           (x1, Data σ1) # Γ \<turnstile> e1 : T;
           (x2, Data σ2) # Γ \<turnstile> e2 : T |]
        ==> thesis |]
  ==> thesis

lemma weakening:

  [| Γ1 \<turnstile> e : T; valid Γ2; Γ1  Γ2 |] ==> Γ2 \<turnstile> e : T

lemma context_exchange:

  (x1, T1) # (x2, T2) # Γ \<turnstile> e : T
  ==> (x2, T2) # (x1, T1) # Γ \<turnstile> e : T

lemma typing_var_unicity:

  (x, t1) # Γ \<turnstile> Var x : t2 ==> t1 = t2

lemma typing_substitution:

  [| (x, T') # Γ \<turnstile> e : T; Γ \<turnstile> e' : T' |]
  ==> Γ \<turnstile> e[x::=e'] : T

lemma big_eqvt':

  pit \<Down> pit' ==> t \<Down> t'

lemma fresh_preserved:

  [| e \<Down> e'; x \<sharp> e |] ==> x \<sharp> e'

lemma b_App_elim:

  [| App e1 e2 \<Down> e'; x \<sharp> (e1, e2, e');
     !!f1 f2.
        [| e1 \<Down> Lam [x].f1; e2 \<Down> f2; f1[x::=f2] \<Down> e' |]
        ==> thesis |]
  ==> thesis

lemma b_CaseL_elim:

  [| Case e of inl x1 -> e1 | inr x2 -> e2 \<Down> e''; !!t. ¬ e \<Down> InR t;
     x1 \<sharp> e''; x1 \<sharp> e; x2 \<sharp> e''; x1 \<sharp> e;
     !!e'. [| e \<Down> InL e'; e1[x1::=e'] \<Down> e'' |] ==> thesis |]
  ==> thesis

lemma b_CaseR_elim:

  [| Case e of inl x1 -> e1 | inr x2 -> e2 \<Down> e''; !!t. ¬ e \<Down> InL t;
     x1 \<sharp> e''; x1 \<sharp> e; x2 \<sharp> e''; x2 \<sharp> e;
     !!e'. [| e \<Down> InR e'; e2[x2::=e'] \<Down> e'' |] ==> thesis |]
  ==> thesis

lemma subject_reduction:

  [| e \<Down> e'; Γ \<turnstile> e : T |] ==> Γ \<turnstile> e' : T

lemma unicity_of_evaluation:

  [| e \<Down> e1; e \<Down> e2 |] ==> e1 = e2

lemma not_val_App:

  ¬ val (App e1 e2)
  ¬ val (Fst e)
  ¬ val (Snd e)
  ¬ val (Var x)
  ¬ val (Case e of inl x1 -> e1 | inr x2 -> e2)

lemma reduces_evaluates_to_values:

  t \<Down> t' ==> val t'

lemma type_prod_evaluates_to_pairs:

  [| Γ \<turnstile> t : Data (DProd S1 S2); t \<Down> t';
     !!t1 t2. t' = Pr t1 t2 ==> thesis |]
  ==> thesis

lemma type_sum_evaluates_to_ins:

  [| Γ \<turnstile> t : Data (DSum σ1 σ2); t \<Down> t' |]
  ==> (∃t''. t' = InL t'') ∨ (∃t''. t' = InR t'')

lemma type_arrow_evaluates_to_lams:

  [| Γ \<turnstile> t : σ->τ; t \<Down> t';
     !!x t''. t' = Lam [x].t'' ==> thesis |]
  ==> thesis

lemma type_nat_evaluates_to_consts:

  [| Γ \<turnstile> t : Data DNat; t \<Down> t'; !!n. t' = Const n ==> thesis |]
  ==> thesis

lemma Vprimes_are_values:

  e ∈ V' S ==> val e

lemma V'_eqvt:

  v ∈ V' S ==> piv ∈ V' S

lemma V_eqvt:

  x ∈ V T ==> pix ∈ V T

lemma V_arrow_elim_weak:

  [| u ∈ V (T1->T2);
     !!a t. [| u = Lam [a].t; ∀v∈V T1. ∃v'. t[a::=v] \<Down> v'v' ∈ V T2 |]
            ==> thesis |]
  ==> thesis

lemma V_arrow_elim_strong:

  [| u ∈ V (T1->T2);
     !!a t. [| a \<sharp> c; u = Lam [a].t;
               ∀v∈V T1. ∃v'. t[a::=v] \<Down> v'v' ∈ V T2 |]
            ==> thesis |]
  ==> thesis

lemma V_are_values:

  e ∈ V T ==> val e

lemma values_reduce_to_themselves:

  val v ==> v \<Down> v

lemma Vs_reduce_to_themselves:

  v ∈ V T ==> v \<Down> v

lemma V_sum:

  x ∈ V (Data (DSum S1 S2))
  ==> (∃y. x = InL yy ∈ V' S1) ∨ (∃y. x = InR yy ∈ V' S2)

lemma monotonicity:

  [| ϑ Vcloses Γ; e ∈ V T; valid ((x, T) # Γ) |] ==> (x, e) # ϑ Vcloses (x, T) # Γ

lemma termination_aux:

  [| Γ \<turnstile> e : T; ϑ Vcloses Γ |] ==> ∃v. ϑ<e> \<Down> vv ∈ V T

theorem termination_of_evaluation:

  [] \<turnstile> e : T ==> ∃v. e \<Down> vval v