Theory SOS

(*                                                        *)
(* Formalisation of some typical SOS-proofs.              *)
(*                                                        *) 
(* This work was inspired by challenge suggested by Adam  *)
(* Chlipala 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 "HOL-Nominal.Nominal"
begin

atom_decl name

text ‹types and terms›
nominal_datatype ty = 
    TVar "nat"
  | Arrow "ty" "ty" ("__" [100,100] 100)

nominal_datatype trm = 
    Var "name"
  | Lam "«name»trm" ("Lam [_]._" [100,100] 100)
  | App "trm" "trm"

lemma fresh_ty:
  fixes x::"name" 
  and   T::"ty"
  shows "xT"
by (induct T rule: ty.induct)
   (auto simp add: fresh_nat)

text ‹Parallel and single 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[eqvt]:
  fixes pi::"name prm"
  shows "pi(lookup θ X) = lookup (piθ) (piX)"
by (induct θ) (auto simp add: eqvts)

lemma lookup_fresh:
  fixes z::"name"
  assumes a: "zθ" and b: "zx"
  shows "z lookup θ x"
using a b
by (induct rule: lookup.induct) (auto simp add: fresh_list_cons)

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

(* parallel substitution *)

nominal_primrec
  psubst :: "(name×trm) list  trm  trm"  ("_<_>" [95,95] 105)
where
  "θ<(Var x)> = (lookup θ x)"
| "θ<(App e1 e2)> = App (θ<e1>) (θ<e2>)"
| "xθ  θ<(Lam [x].e)> = Lam [x].(θ<e>)"
apply(finite_guess)+
apply(rule TrueI)+
apply(simp add: abs_fresh)+
apply(fresh_guess)+
done

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

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

lemma psubst_empty[simp]:
  shows "[]<t> = t"
  by (nominal_induct t rule: trm.strong_induct) 
     (auto simp add: fresh_list_nil)

text ‹Single substitution›
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(y,t')  (Lam [x].t)[y::=t'] = Lam [x].(t[y::=t'])"
by (simp_all add: fresh_list_cons fresh_list_nil)

lemma fresh_subst:
  fixes z::"name"
  shows "zs; (z=y  zt)  zt[y::=s]"
by (nominal_induct t avoiding: z y s rule: trm.strong_induct)
   (auto simp add: abs_fresh fresh_prod fresh_atm)

lemma forget: 
  assumes a: "xe" 
  shows "e[x::=e'] = e"
using a
by (nominal_induct e avoiding: x e' rule: trm.strong_induct)
   (auto simp add: fresh_atm abs_fresh)

lemma psubst_subst_psubst:
  assumes h: "xθ"
  shows "θ<e>[x::=e'] = ((x,e')#θ)<e>"
  using h
by (nominal_induct e avoiding: θ x e' rule: trm.strong_induct)
   (auto simp add: fresh_list_cons fresh_atm forget lookup_fresh lookup_fresh')

text ‹Typing Judgements›

inductive
  valid :: "(name×ty) list  bool"
where
  v_nil[intro]:  "valid []"
| v_cons[intro]: "valid Γ;xΓ  valid ((x,T)#Γ)"

equivariance valid 

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

lemma valid_insert:
  assumes a: "valid (Δ@[(x,T)]@Γ)"
  shows "valid (Δ @ Γ)" 
using a
by (induct Δ)
   (auto simp add:  fresh_list_append fresh_list_cons elim!: valid_elim)

lemma fresh_set: 
  shows "yxs = (xset xs. yx)"
by (induct xs) (simp_all add: fresh_list_nil fresh_list_cons)

lemma context_unique:
  assumes a1: "valid Γ"
  and     a2: "(x,T)  set Γ"
  and     a3: "(x,U)  set Γ"
  shows "T = U" 
using a1 a2 a3
by (induct) (auto simp add: fresh_set fresh_prod fresh_atm)

text ‹Typing Relation›

inductive
  typing :: "(name×ty) listtrmtybool" ("_  _ : _" [60,60,60] 60) 
where
  t_Var[intro]:   "valid Γ; (x,T)set Γ  Γ  Var x : T"
| t_App[intro]:   "Γ  e1 : T1T2; Γ  e2 : T1  Γ  App e1 e2 : T2"
| t_Lam[intro]:   "xΓ; (x,T1)#Γ  e : T2  Γ  Lam [x].e : T1T2"

equivariance typing

nominal_inductive typing
  by (simp_all add: abs_fresh fresh_ty)

lemma typing_implies_valid:
  assumes a: "Γ  t : T"
  shows "valid Γ"
using a by (induct) (auto)


lemma t_App_elim:
  assumes a: "Γ  App t1 t2 : T"
  obtains T' where "Γ  t1 : T'  T" and "Γ  t2 : T'"
using a
by (cases) (auto simp add: trm.inject)

lemma t_Lam_elim: 
  assumes a: "Γ  Lam [x].t : T" "xΓ"
  obtains T1 and T2 where "(x,T1)#Γ  t : T2" and "T=T1T2"
using a
by (cases rule: typing.strong_cases [where x="x"])
   (auto simp add: abs_fresh fresh_ty alpha trm.inject)

abbreviation
  "sub_context" :: "(name×ty) list  (name×ty) list  bool" ("_  _" [55,55] 55)
where
  "Γ1  Γ2  x T. (x,T)set Γ1  (x,T)set Γ2"

lemma weakening: 
  fixes Γ1 Γ2::"(name×ty) list"
  assumes "Γ1  e: T" and "valid Γ2" and "Γ1  Γ2"
  shows "Γ2  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 vc: "xΓ2" by fact
  have ih: "valid ((x,T1)#Γ2); (x,T1)#Γ1  (x,T1)#Γ2  (x,T1)#Γ2  t : T2" by fact
  have "valid Γ2" by fact
  then have "valid ((x,T1)#Γ2)" using vc by auto
  moreover
  have "Γ1  Γ2" by fact
  then have "(x,T1)#Γ1  (x,T1)#Γ2" by simp
  ultimately have "(x,T1)#Γ2  t : T2" using ih by simp 
  with vc show "Γ2  Lam [x].t : T1T2" by auto
qed (auto)

lemma type_substitutivity_aux:
  assumes a: "(Δ@[(x,T')]@Γ)  e : T"
  and     b: "Γ  e' : T'"
  shows "(Δ@Γ)  e[x::=e'] : T" 
using a b 
proof (nominal_induct Γ"Δ@[(x,T')]@Γ" e T avoiding: e' Δ rule: typing.strong_induct)
  case (t_Var y T e' Δ)
  then have a1: "valid (Δ@[(x,T')]@Γ)" 
       and  a2: "(y,T)  set (Δ@[(x,T')]@Γ)" 
       and  a3: "Γ  e' : T'" .
  from a1 have a4: "valid (Δ@Γ)" by (rule valid_insert)
  { assume eq: "x=y"
    from a1 a2 have "T=T'" using eq by (auto intro: context_unique)
    with a3 have "Δ@Γ  Var y[x::=e'] : T" using eq a4 by (auto intro: weakening)
  }
  moreover
  { assume ineq: "xy"
    from a2 have "(y,T)  set (Δ@Γ)" using ineq by simp
    then have "Δ@Γ  Var y[x::=e'] : T" using ineq a4 by auto
  }
  ultimately show "Δ@Γ  Var y[x::=e'] : T" by blast
qed (force simp add: fresh_list_append fresh_list_cons)+

corollary type_substitutivity:
  assumes a: "(x,T')#Γ  e : T"
  and     b: "Γ  e' : T'"
  shows "Γ  e[x::=e'] : T"
using a b type_substitutivity_aux[where Δ="[]"]
by (auto)

text ‹Values›
inductive
  val :: "trmbool" 
where
  v_Lam[intro]:   "val (Lam [x].e)"

equivariance val 

lemma not_val_App[simp]:
  shows 
  "¬ val (App e1 e2)" 
  "¬ val (Var x)"
  by (auto elim: val.cases)

text ‹Big-Step Evaluation›

inductive
  big :: "trmtrmbool" ("_  _" [80,80] 80) 
where
  b_Lam[intro]:   "Lam [x].e  Lam [x].e"
| b_App[intro]:   "x(e1,e2,e'); e1Lam [x].e; e2e2'; e[x::=e2']e'  App e1 e2  e'"

equivariance big

nominal_inductive big
  by (simp_all add: abs_fresh)

lemma big_preserves_fresh:
  fixes x::"name"
  assumes a: "e  e'" "xe" 
  shows "xe'"
  using a by (induct) (auto simp add: abs_fresh fresh_subst)

lemma b_App_elim:
  assumes a: "App e1 e2  e'" "x(e1,e2,e')"
  obtains f1 and f2 where "e1  Lam [x]. f1" "e2  f2" "f1[x::=f2]  e'"
  using a
by (cases rule: big.strong_cases[where x="x" and xa="x"])
   (auto simp add: trm.inject)

lemma subject_reduction:
  assumes a: "e  e'" and b: "Γ  e : T"
  shows "Γ  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Γ" by fact
  have "Γ  App e1 e2 : T" by fact
  then obtain T' where a1: "Γ  e1 : T'T" and a2: "Γ  e2 : T'" 
    by (cases) (auto simp add: trm.inject)
  have ih1: "Γ  e1 : T'  T  Γ  Lam [x].e : T'  T" by fact
  have ih2: "Γ  e2 : T'  Γ  e2' : T'" by fact 
  have ih3: "Γ  e[x::=e2'] : T  Γ  e' : T" by fact
  have "Γ  Lam [x].e : T'T" using ih1 a1 by simp 
  then have "((x,T')#Γ)  e : T" using vc  
    by (auto elim: t_Lam_elim simp add: ty.inject)
  moreover
  have "Γ  e2': T'" using ih2 a2 by simp
  ultimately have "Γ  e[x::=e2'] : T" by (simp add: type_substitutivity)
  thus "Γ  e' : T" using ih3 by simp
qed (blast)

lemma subject_reduction2:
  assumes a: "e  e'" and b: "Γ  e : T"
  shows "Γ  e' : T"
  using a b
by (nominal_induct avoiding: Γ T rule: big.strong_induct)
   (force elim: t_App_elim t_Lam_elim simp add: ty.inject type_substitutivity)+

lemma unicity_of_evaluation:
  assumes a: "e  e1" 
  and     b: "e  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  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  t  Lam [x].e1' = t" by fact
  have ih2:"t. e2  t  e2' = t" by fact
  have ih3: "t. e1'[x::=e2']  t  e' = t" by fact
  have app: "App e1 e2  t2" by fact
  have vc: "xe1" "xe2" "xt2" by fact+
  then have "xApp e1 e2" by auto
  from app vc obtain f1 f2 where x1: "e1  Lam [x]. f1" and x2: "e2  f2" and x3: "f1[x::=f2]  t2"
    by (auto elim!: b_App_elim)
  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']  t2" using x3 by simp
  thus "e' = t2" using ih3 by simp
qed

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

(* Valuation *)

nominal_primrec
  V :: "ty  trm set" 
where
  "V (TVar x) = {e. val e}"
| "V (T1  T2) = {Lam [x].e | x e.  v  (V T1).  v'. e[x::=v]  v'  v'  V T2}"
  by (rule TrueI)+ 

lemma V_eqvt:
  fixes pi::"name prm"
  assumes a: "xV T"
  shows "(pix)V T"
using a
apply(nominal_induct T arbitrary: pi x rule: ty.strong_induct)
apply(auto simp add: trm.inject)
apply(simp add: eqvts)
apply(rule_tac x="pixa" in exI)
apply(rule_tac x="pie" in exI)
apply(simp)
apply(auto)
apply(drule_tac x="(rev pi)v" in bspec)
apply(force)
apply(auto)
apply(rule_tac x="piv'" in exI)
apply(auto)
apply(drule_tac pi="pi" in big.eqvt)
apply(perm_simp add: eqvts)
done

lemma V_arrow_elim_weak:
  assumes h:"u  V (T1  T2)"
  obtains a t where "u = Lam [a].t" and " v  (V T1).  v'. t[a::=v]  v'  v'  V T2"
using h by (auto)

lemma V_arrow_elim_strong:
  fixes c::"'a::fs_name"
  assumes h: "u  V (T1  T2)"
  obtains a t where "ac" "u = Lam [a].t" "v  (V T1).  v'. t[a::=v]  v'  v'  V T2"
using h
apply -
apply(erule V_arrow_elim_weak)
apply(subgoal_tac "a'::name. a'(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 Vs_are_values:
  assumes a: "e  V T"
  shows "val e"
using a by (nominal_induct T arbitrary: e rule: ty.strong_induct) (auto)

lemma values_reduce_to_themselves:
  assumes a: "val v"
  shows "v  v"
using a by (induct) (auto)

lemma Vs_reduce_to_themselves:
  assumes a: "v  V T"
  shows "v  v"
using a by (simp add: values_reduce_to_themselves Vs_are_values)

text ‹'θ maps x to e' asserts that θ substitutes x with e›
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 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Γ" 
      by (induct Γ) (auto simp add: fresh_list_cons fresh_prod fresh_atm)
    moreover have "mΓ" using asm1 by auto
    ultimately have False using eq by auto
  }
  ultimately show ?thesis by auto
qed

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:
  assumes h1: "Γ  e : T"
  and     h2: "θ Vcloses Γ"
  shows "v. θ<e>  v  v  V T" 
using h2 h1
proof(nominal_induct e avoiding: Γ θ arbitrary: T rule: trm.strong_induct)
  case (App e1 e2 Γ θ T)
  have ih1: "θ Γ T. θ Vcloses Γ; Γ  e1 : T  v. θ<e1>  v  v  V T" by fact
  have ih2: "θ Γ T. θ Vcloses Γ; Γ  e2 : T  v. θ<e2>  v  v  V T" by fact
  have as1: "θ Vcloses Γ" by fact 
  have as2: "Γ  App e1 e2 : T" by fact
  then obtain T' where "Γ  e1 : T'  T" and "Γ  e2 : T'" by (auto elim: t_App_elim)
  then obtain v1 v2 where "(i)": "θ<e1>  v1" "v1  V (T'  T)"
                      and "(ii)": "θ<e2>  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]  v'  v'  V T)"
            and "(iv)":  "θ<e1>  Lam [x].e'" 
            and fr: "x(θ,e1,e2)" by (blast elim: V_arrow_elim_strong)
  from fr have fr1: "xθ<e1>" and fr2: "xθ<e2>" by (simp_all add: fresh_psubst)
  from "(ii)" "(iii)" obtain v3 where "(v)": "e'[x::=v2]  v3" "v3  V T" by auto
  from fr2 "(ii)" have "xv2" by (simp add: big_preserves_fresh)
  then have "xe'[x::=v2]" by (simp add: fresh_subst)
  then have fr3: "xv3" using "(v)" by (simp add: big_preserves_fresh)
  from fr1 fr2 fr3 have "x(θ<e1>,θ<e2>,v3)" by simp
  with "(iv)" "(ii)" "(v)" have "App (θ<e1>) (θ<e2>)  v3" by auto
  then show "v. θ<App e1 e2>  v  v  V T" using "(v)" by auto
next
  case (Lam x e Γ θ T)
  have ih:"θ Γ T. θ Vcloses Γ; Γ  e : T  v. θ<e>  v  v  V T" by fact
  have as1: "θ Vcloses Γ" by fact
  have as2: "Γ  Lam [x].e : T" by fact
  have fs: "xΓ" "xθ" by fact+
  from as2 fs obtain T1 T2 
    where "(i)": "(x,T1)#Γ  e:T2" and "(ii)": "T = T1  T2" using fs
    by (auto elim: t_Lam_elim)
  from "(i)" have "(iii)": "valid ((x,T1)#Γ)" by (simp add: typing_implies_valid)
  have "v  (V T1). v'. (θ<e>)[x::=v]  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>  v'  v'  V T2" by blast
    then have "θ<e>[x::=v]  v'  v'  V T2" using fs by (simp add: psubst_subst_psubst)
    then show "v'. θ<e>[x::=v]  v'  v'  V T2" by auto
  qed
  then have "Lam[x].θ<e>  V (T1  T2)" by auto
  then have "θ<Lam [x].e>  Lam [x].θ<e>  Lam [x].θ<e>  V (T1T2)" using fs by auto
  thus "v. θ<Lam [x].e>  v  v  V T" using "(ii)" by auto
next
  case (Var x Γ θ T)
  have "Γ  (Var x) : T" by fact
  then have "(x,T)set Γ" by (cases) (auto simp add: trm.inject)
  with Var have "θ<Var x>  θ<Var x>  θ<Var x> V T" 
    by (auto intro!: Vs_reduce_to_themselves)
  then show "v. θ<Var x>  v  v  V T" by auto
qed 

theorem termination_of_evaluation:
  assumes a: "[]  e : T"
  shows "v. e  v  val v"
proof -
  from a have "v. []<e>  v  v  V T" by (rule termination_aux) (auto)
  thus "v. e  v  val v" using Vs_are_values by auto
qed 

end