src/HOL/Nominal/Examples/SN.thy
author urbanc
Wed, 28 Mar 2007 17:27:44 +0200
changeset 22540 e4817fa0f6a1
parent 22538 3ccb92dfb5e9
child 22541 c33b542394f3
permissions -rw-r--r--
adapted to new nominal_inductive

(* $Id$ *)

theory SN
imports Lam_Funs
begin

text {* Strong Normalisation proof from the Proofs and Types book *}

section {* Beta Reduction *}

lemma subst_rename: 
  assumes a: "c\<sharp>t1"
  shows "t1[a::=t2] = ([(c,a)]\<bullet>t1)[c::=t2]"
using a
by (nominal_induct t1 avoiding: a c t2 rule: lam.induct)
   (auto simp add: calc_atm fresh_atm abs_fresh)

lemma forget: 
  assumes a: "a\<sharp>t1"
  shows "t1[a::=t2] = t1"
  using a
by (nominal_induct t1 avoiding: a t2 rule: lam.induct)
   (auto simp add: abs_fresh fresh_atm)

lemma fresh_fact: 
  fixes a::"name"
  assumes a: "a\<sharp>t1"
  and     b: "a\<sharp>t2"
  shows "a\<sharp>t1[b::=t2]"
using a b
by (nominal_induct t1 avoiding: a b t2 rule: lam.induct)
   (auto simp add: abs_fresh fresh_atm)

lemma fresh_fact': 
  fixes a::"name"
  assumes a: "a\<sharp>t2"
  shows "a\<sharp>t1[a::=t2]"
using a 
by (nominal_induct t1 avoiding: a t2 rule: lam.induct)
   (auto simp add: abs_fresh fresh_atm)

lemma subst_lemma:  
  assumes a: "x\<noteq>y"
  and     b: "x\<sharp>L"
  shows "M[x::=N][y::=L] = M[y::=L][x::=N[y::=L]]"
using a b
by (nominal_induct M avoiding: x y N L rule: lam.induct)
   (auto simp add: fresh_fact forget)

lemma id_subs: 
  shows "t[x::=Var x] = t"
  by (nominal_induct t avoiding: x rule: lam.induct)
     (simp_all add: fresh_atm)

lemma subst_eqvt[eqvt]:
  fixes pi::"name prm" 
  and   t::"lam"
  shows "pi\<bullet>(t[x::=t']) = (pi\<bullet>t)[(pi\<bullet>x)::=(pi\<bullet>t')]"
by (nominal_induct t avoiding: x t' rule: lam.induct)
   (perm_simp add: fresh_bij)+

inductive2 Beta :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^isub>\<beta> _" [80,80] 80)
where
  b1[intro!]: "s1\<longrightarrow>\<^isub>\<beta>s2 \<Longrightarrow> (App s1 t)\<longrightarrow>\<^isub>\<beta>(App s2 t)"
| b2[intro!]: "s1\<longrightarrow>\<^isub>\<beta>s2 \<Longrightarrow> (App t s1)\<longrightarrow>\<^isub>\<beta>(App t s2)"
| b3[intro!]: "s1\<longrightarrow>\<^isub>\<beta>s2 \<Longrightarrow> (Lam [a].s1)\<longrightarrow>\<^isub>\<beta> (Lam [a].s2)"
| b4[intro!]: "a\<sharp>s2 \<Longrightarrow>(App (Lam [a].s1) s2)\<longrightarrow>\<^isub>\<beta>(s1[a::=s2])"

nominal_inductive Beta
  by (simp_all add: abs_fresh fresh_fact')

abbreviation "Beta_star"  :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^isub>\<beta>\<^sup>* _" [80,80] 80) where
  "t1 \<longrightarrow>\<^isub>\<beta>\<^sup>* t2 \<equiv> Beta\<^sup>*\<^sup>* t1 t2"

lemma supp_beta: 
  assumes a: "t\<longrightarrow>\<^isub>\<beta> s"
  shows "(supp s)\<subseteq>((supp t)::name set)"
using a
by (induct)
   (auto intro!: simp add: abs_supp lam.supp subst_supp)

lemma beta_abs: "Lam [a].t\<longrightarrow>\<^isub>\<beta> t'\<Longrightarrow>\<exists>t''. t'=Lam [a].t'' \<and> t\<longrightarrow>\<^isub>\<beta> t''"
apply(ind_cases2 "Lam [a].t  \<longrightarrow>\<^isub>\<beta> t'")
apply(auto simp add: lam.distinct lam.inject)
apply(auto simp add: alpha)
apply(rule_tac x="[(a,aa)]\<bullet>s2" in exI)
apply(rule conjI)
apply(rule sym)
apply(rule pt_bij2[OF pt_name_inst, OF at_name_inst])
apply(simp)
apply(rule pt_name3)
apply(simp add: at_ds5[OF at_name_inst])
apply(rule conjI)
apply(simp add: pt_fresh_left[OF pt_name_inst, OF at_name_inst] calc_atm)
apply(force dest!: supp_beta simp add: fresh_def)
apply(force intro!: eqvt)
done

lemma beta_subst: 
  assumes a: "M \<longrightarrow>\<^isub>\<beta> M'"
  shows "M[x::=N]\<longrightarrow>\<^isub>\<beta> M'[x::=N]" 
using a
apply(nominal_induct M M' avoiding: x N rule: Beta.strong_induct)
apply(auto simp add: fresh_atm subst_lemma fresh_fact)
done

section {* types *}

nominal_datatype ty =
    TVar "nat"
  | TArr "ty" "ty" (infix "\<rightarrow>" 200)

lemma perm_ty[simp]:
  fixes pi ::"name prm"
  and   \<tau>  ::"ty"
  shows "pi\<bullet>\<tau> = \<tau>"
by (nominal_induct \<tau> rule: ty.induct) 
   (simp_all add: perm_nat_def)

lemma fresh_ty:
  fixes a ::"name"
  and   \<tau>  ::"ty"
  shows "a\<sharp>\<tau>"
  by (simp add: fresh_def supp_def)

(* valid contexts *)

consts
  "dom_ty" :: "(name\<times>ty) list \<Rightarrow> (name list)"
primrec
  "dom_ty []    = []"
  "dom_ty (x#\<Gamma>) = (fst x)#(dom_ty \<Gamma>)" 

inductive2 valid :: "(name\<times>ty) list \<Rightarrow> bool"
where
  v1[intro]: "valid []"
| v2[intro]: "\<lbrakk>valid \<Gamma>;a\<sharp>\<Gamma>\<rbrakk>\<Longrightarrow> valid ((a,\<sigma>)#\<Gamma>)"

equivariance valid 

(* typing judgements *)

lemma fresh_context: 
  fixes  \<Gamma> :: "(name\<times>ty)list"
  and    a :: "name"
  assumes a: "a\<sharp>\<Gamma>"
  shows "\<not>(\<exists>\<tau>::ty. (a,\<tau>)\<in>set \<Gamma>)"
using a
apply(induct \<Gamma>)
apply(auto simp add: fresh_prod fresh_list_cons fresh_atm)
done

inductive_cases2 valid_elim[elim]: "valid ((a,\<tau>)#\<Gamma>)"

lemma valid_unicity: 
  assumes a: "valid \<Gamma>"
  and     b: "(c,\<sigma>)\<in>set \<Gamma>"
  and     c: "(c,\<tau>)\<in>set \<Gamma>"
  shows "\<sigma>=\<tau>" 
using a b c
by (induct \<Gamma>) (auto dest!: fresh_context)

inductive2 typing :: "(name\<times>ty) list\<Rightarrow>lam\<Rightarrow>ty\<Rightarrow>bool" (" _ \<turnstile> _ : _ " [80,80,80] 80)
where
  t1[intro]: "\<lbrakk>valid \<Gamma>; (a,\<tau>)\<in>set \<Gamma>\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> Var a : \<tau>"
| t2[intro]: "\<lbrakk>\<Gamma> \<turnstile> t1 : \<tau>\<rightarrow>\<sigma>; \<Gamma> \<turnstile> t2 : \<tau>\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> App t1 t2 : \<sigma>"
| t3[intro]: "\<lbrakk>a\<sharp>\<Gamma>;((a,\<tau>)#\<Gamma>) \<turnstile> t : \<sigma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [a].t : \<tau>\<rightarrow>\<sigma>"

nominal_inductive typing
  by (simp_all add: abs_fresh fresh_ty)

abbreviation
  "sub" :: "(name\<times>ty) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool" (" _ \<lless> _ " [80,80] 80) where
  "\<Gamma>1 \<lless> \<Gamma>2 \<equiv> \<forall>a \<sigma>. (a,\<sigma>)\<in>set \<Gamma>1 \<longrightarrow>  (a,\<sigma>)\<in>set \<Gamma>2"

lemma weakening: 
  assumes a: "\<Gamma>1 \<turnstile> t : \<sigma>" 
  and     b: "valid \<Gamma>2" 
  and     c: "\<Gamma>1 \<lless> \<Gamma>2"
  shows "\<Gamma>2 \<turnstile> t:\<sigma>"
using a b c
apply(nominal_induct \<Gamma>1 t \<sigma> avoiding: \<Gamma>2 rule: typing.strong_induct)
apply(auto | atomize)+
(* FIXME: before using meta-connectives and the new induction *)
(* method, this was completely automatic *)
done

lemma in_ctxt: 
  assumes a: "(a,\<tau>)\<in>set \<Gamma>"
  shows "a\<in>set(dom_ty \<Gamma>)"
using a
by (induct \<Gamma>) (auto)

lemma free_vars: 
  assumes a: "\<Gamma> \<turnstile> t : \<tau>"
  shows " (supp t)\<subseteq>set(dom_ty \<Gamma>)"
using a
by (nominal_induct \<Gamma> t \<tau> rule: typing.strong_induct)
   (auto simp add: lam.supp abs_supp supp_atm in_ctxt)

lemma t1_elim: "\<Gamma> \<turnstile> Var a : \<tau> \<Longrightarrow> valid \<Gamma> \<and> (a,\<tau>) \<in> set \<Gamma>"
apply(ind_cases2 "\<Gamma> \<turnstile> Var a : \<tau>")
apply(auto simp add: lam.inject lam.distinct)
done

lemma t2_elim: "\<Gamma> \<turnstile> App t1 t2 : \<sigma> \<Longrightarrow> \<exists>\<tau>. (\<Gamma> \<turnstile> t1 : \<tau>\<rightarrow>\<sigma> \<and> \<Gamma> \<turnstile> t2 : \<tau>)"
apply(ind_cases2 "\<Gamma> \<turnstile> App t1 t2 : \<sigma>")
apply(auto simp add: lam.inject lam.distinct)
done

lemma t3_elim: "\<lbrakk>\<Gamma> \<turnstile> Lam [a].t : \<sigma>;a\<sharp>\<Gamma>\<rbrakk>\<Longrightarrow> \<exists>\<tau> \<tau>'. \<sigma>=\<tau>\<rightarrow>\<tau>' \<and> ((a,\<tau>)#\<Gamma>) \<turnstile> t : \<tau>'"
apply(ind_cases2 "\<Gamma> \<turnstile> Lam [a].t : \<sigma>")
apply(auto simp add: lam.distinct lam.inject alpha) 
apply(drule_tac pi="[(a,aa)]::name prm" in typing_eqvt)
apply(simp)
apply(subgoal_tac "([(a,aa)]::name prm)\<bullet>\<Gamma> = \<Gamma>")(*A*)
apply(force simp add: calc_atm)
(*A*)
apply(force intro!: perm_fresh_fresh)
done

lemma typing_valid: 
  assumes a: "\<Gamma> \<turnstile> t : \<tau>" 
  shows "valid \<Gamma>"
using a by (induct, auto)

lemma ty_subs:
  assumes a: "((c,\<sigma>)#\<Gamma>) \<turnstile> t1:\<tau>"
  and     b: "\<Gamma>\<turnstile> t2:\<sigma>"
  shows  "\<Gamma> \<turnstile> t1[c::=t2]:\<tau>"
using a b
proof(nominal_induct t1 avoiding: \<Gamma> \<sigma> \<tau> c t2 rule: lam.induct)
  case (Var a) 
  have a1: "\<Gamma> \<turnstile>t2:\<sigma>" by fact
  have a2: "((c,\<sigma>)#\<Gamma>) \<turnstile> Var a:\<tau>" by fact
  hence a21: "(a,\<tau>)\<in>set((c,\<sigma>)#\<Gamma>)" and a22: "valid((c,\<sigma>)#\<Gamma>)" by (auto dest: t1_elim)
  from a22 have a23: "valid \<Gamma>" and a24: "c\<sharp>\<Gamma>" by (auto dest: valid_elim) 
  from a24 have a25: "\<not>(\<exists>\<tau>. (c,\<tau>)\<in>set \<Gamma>)" by (rule fresh_context)
  show "\<Gamma>\<turnstile>(Var a)[c::=t2] : \<tau>"
  proof (cases "a=c", simp_all)
    assume case1: "a=c"
    show "\<Gamma> \<turnstile> t2:\<tau>" using a1
    proof (cases "\<sigma>=\<tau>")
      assume "\<sigma>=\<tau>" thus ?thesis using a1 by simp 
    next
      assume a3: "\<sigma>\<noteq>\<tau>"
      show ?thesis
      proof (rule ccontr)
	from a3 a21 have "(a,\<tau>)\<in>set \<Gamma>" by force
	with case1 a25 show False by force 
      qed
    qed
  next
    assume case2: "a\<noteq>c"
    with a21 have a26: "(a,\<tau>)\<in>set \<Gamma>" by force 
    from a23 a26 show "\<Gamma> \<turnstile> Var a:\<tau>" by force
  qed
next
  case (App s1 s2)
  have ih_s1: "\<And>c \<sigma> \<tau> t2 \<Gamma>. ((c,\<sigma>)#\<Gamma>) \<turnstile> s1:\<tau> \<Longrightarrow> \<Gamma>\<turnstile> t2: \<sigma> \<Longrightarrow> \<Gamma> \<turnstile> s1[c::=t2]:\<tau>" by fact
  have ih_s2: "\<And>c \<sigma> \<tau> t2 \<Gamma>. ((c,\<sigma>)#\<Gamma>) \<turnstile> s2:\<tau> \<Longrightarrow> \<Gamma>\<turnstile> t2: \<sigma> \<Longrightarrow> \<Gamma> \<turnstile> s2[c::=t2]:\<tau>" by fact
  have "((c,\<sigma>)#\<Gamma>)\<turnstile>App s1 s2 : \<tau>" by fact
  hence "\<exists>\<tau>'. ((c,\<sigma>)#\<Gamma>)\<turnstile>s1:\<tau>'\<rightarrow>\<tau> \<and> ((c,\<sigma>)#\<Gamma>)\<turnstile>s2:\<tau>'" by (rule t2_elim) 
  then obtain \<tau>' where "((c,\<sigma>)#\<Gamma>)\<turnstile>s1:\<tau>'\<rightarrow>\<tau>" and "((c,\<sigma>)#\<Gamma>)\<turnstile>s2:\<tau>'" by blast
  moreover
  have "\<Gamma> \<turnstile>t2:\<sigma>" by fact
  ultimately show "\<Gamma> \<turnstile>  (App s1 s2)[c::=t2] : \<tau>" using ih_s1 ih_s2 by (simp, blast)
next
  case (Lam a s)
  have "a\<sharp>\<Gamma>" "a\<sharp>\<sigma>" "a\<sharp>\<tau>" "a\<sharp>c" "a\<sharp>t2" by fact 
  hence f1: "a\<sharp>\<Gamma>" and f2: "a\<noteq>c" and f2': "c\<sharp>a" and f3: "a\<sharp>t2" and f4: "a\<sharp>((c,\<sigma>)#\<Gamma>)"
    by (auto simp add: fresh_atm fresh_prod fresh_list_cons)
  have c1: "((c,\<sigma>)#\<Gamma>)\<turnstile>Lam [a].s : \<tau>" by fact
  hence "\<exists>\<tau>1 \<tau>2. \<tau>=\<tau>1\<rightarrow>\<tau>2 \<and> ((a,\<tau>1)#(c,\<sigma>)#\<Gamma>) \<turnstile> s : \<tau>2" using f4 by (auto dest: t3_elim) 
  then obtain \<tau>1 \<tau>2 where c11: "\<tau>=\<tau>1\<rightarrow>\<tau>2" and c12: "((a,\<tau>1)#(c,\<sigma>)#\<Gamma>) \<turnstile> s : \<tau>2" by force
  from c12 have "valid ((a,\<tau>1)#(c,\<sigma>)#\<Gamma>)" by (rule typing_valid)
  hence ca: "valid \<Gamma>" and cb: "a\<sharp>\<Gamma>" and cc: "c\<sharp>\<Gamma>" 
    by (auto simp add: fresh_list_cons) 
  from c12 have c14: "((c,\<sigma>)#(a,\<tau>1)#\<Gamma>) \<turnstile> s : \<tau>2"
  proof -
    have c2: "((a,\<tau>1)#(c,\<sigma>)#\<Gamma>) \<lless> ((c,\<sigma>)#(a,\<tau>1)#\<Gamma>)" by force
    have c3: "valid ((c,\<sigma>)#(a,\<tau>1)#\<Gamma>)"
      by (rule v2, rule v2, auto simp add: fresh_list_cons fresh_prod ca cb cc f2' fresh_ty)
    from c12 c2 c3 show ?thesis by (force intro: weakening)
  qed
  assume c8: "\<Gamma> \<turnstile> t2 : \<sigma>"
  have c81: "((a,\<tau>1)#\<Gamma>)\<turnstile>t2 :\<sigma>"
  proof -
    have c82: "\<Gamma> \<lless> ((a,\<tau>1)#\<Gamma>)" by force
    have c83: "valid ((a,\<tau>1)#\<Gamma>)" using f1 ca by force
    with c8 c82 c83 show ?thesis by (force intro: weakening)
  qed
  show "\<Gamma> \<turnstile> (Lam [a].s)[c::=t2] : \<tau>"
    using c11 prems c14 c81 f1 by force
qed

lemma subject: 
  assumes a: "t1\<longrightarrow>\<^isub>\<beta>t2"
  and     b: "\<Gamma> \<turnstile> t1:\<tau>"
  shows "\<Gamma> \<turnstile> t2:\<tau>"
using a b
proof (nominal_induct t1 t2 avoiding: \<Gamma> \<tau> rule: Beta.strong_induct)
  case (b1 s1 s2 t) --"App-case left"
  have ih: "\<And>\<Gamma> \<tau>. \<Gamma> \<turnstile> s1:\<tau> \<Longrightarrow> \<Gamma> \<turnstile> s2 : \<tau>" by fact
  have "\<Gamma> \<turnstile> App s1 t : \<tau>" by fact 
  hence "\<exists>\<sigma>. \<Gamma> \<turnstile> s1 : \<sigma>\<rightarrow>\<tau> \<and> \<Gamma> \<turnstile> t : \<sigma>" by (rule t2_elim)
  then obtain \<sigma> where "\<Gamma> \<turnstile> s1 : \<sigma>\<rightarrow>\<tau>" and "\<Gamma> \<turnstile> t : \<sigma>" by blast
  with ih show "\<Gamma> \<turnstile> App s2 t : \<tau>" by blast
next
  case (b2 s1 s2 t) --"App-case right"
  have ih: "\<And>\<Gamma> \<tau>. \<Gamma> \<turnstile> s1 : \<tau> \<Longrightarrow> \<Gamma> \<turnstile> s2 : \<tau>" by fact 
  have "\<Gamma> \<turnstile> App t s1 : \<tau>" by fact
  hence "\<exists>\<sigma>. \<Gamma> \<turnstile> t : \<sigma>\<rightarrow>\<tau> \<and> \<Gamma> \<turnstile> s1 : \<sigma>" by (rule t2_elim)
  then obtain \<sigma> where "\<Gamma> \<turnstile> t : \<sigma>\<rightarrow>\<tau>" and "\<Gamma> \<turnstile> s1 : \<sigma>" by blast
  with ih show "\<Gamma> \<turnstile> App t s2 : \<tau>" by blast
next
  case (b3 s1 s2 a) --"Lam-case"
  have fr: "a\<sharp>\<Gamma>" "a\<sharp>\<tau>" by fact
  have ih: "\<And>\<Gamma> \<tau>. \<Gamma> \<turnstile> s1 : \<tau> \<Longrightarrow> \<Gamma> \<turnstile> s2 : \<tau>" by fact
  have "\<Gamma> \<turnstile> Lam [a].s1 : \<tau>" by fact
  with fr have "\<exists>\<tau>1 \<tau>2. \<tau>=\<tau>1\<rightarrow>\<tau>2 \<and> ((a,\<tau>1)#\<Gamma>) \<turnstile> s1 : \<tau>2" by (simp add: t3_elim)
  then obtain \<tau>1 \<tau>2 where "\<tau>=\<tau>1\<rightarrow>\<tau>2" and "((a,\<tau>1)#\<Gamma>) \<turnstile> s1 : \<tau>2" by blast
  with ih show "\<Gamma> \<turnstile> Lam [a].s2 : \<tau>" using fr by blast
next
  case (b4 a s2 s1) --"Beta-redex"
  have fr: "a\<sharp>\<Gamma>" by fact
  have "\<Gamma> \<turnstile> App (Lam [a].s1) s2 : \<tau>" by fact
  hence "\<exists>\<sigma>. (\<Gamma> \<turnstile> (Lam [a].s1) : \<sigma>\<rightarrow>\<tau> \<and> \<Gamma> \<turnstile> s2 : \<sigma>)" by (simp add: t2_elim)
  then obtain \<sigma> where a1: "\<Gamma> \<turnstile> (Lam [a].s1) : \<sigma>\<rightarrow>\<tau>" and a2: "\<Gamma> \<turnstile> s2 : \<sigma>" by blast
  from a1 have "((a,\<sigma>)#\<Gamma>) \<turnstile> s1 : \<tau>" using fr by (auto dest!: t3_elim simp add: ty.inject) 
  with a2 show "\<Gamma> \<turnstile> s1[a::=s2] : \<tau>" by (simp add: ty_subs)
qed

lemma subject_automatic: 
  assumes a: "t1\<longrightarrow>\<^isub>\<beta>t2"
  and     b: "\<Gamma> \<turnstile> t1:\<tau>"
  shows "\<Gamma> \<turnstile> t2:\<tau>"
using a b
apply(nominal_induct t1 t2 avoiding: \<Gamma> \<tau> rule: Beta.strong_induct)
apply(auto dest!: t2_elim t3_elim intro: ty_subs simp add: ty.inject)
done

subsection {* some facts about beta *}

constdefs
  "NORMAL" :: "lam \<Rightarrow> bool"
  "NORMAL t \<equiv> \<not>(\<exists>t'. t\<longrightarrow>\<^isub>\<beta> t')"

lemma NORMAL_Var:
  shows "NORMAL (Var a)"
proof -
  { assume "\<exists>t'. (Var a) \<longrightarrow>\<^isub>\<beta> t'"
    then obtain t' where "(Var a) \<longrightarrow>\<^isub>\<beta> t'" by blast
    hence False by (cases, auto) 
  }
  thus "NORMAL (Var a)" by (force simp add: NORMAL_def)
qed

constdefs
  "SN" :: "lam \<Rightarrow> bool"
  "SN t \<equiv> termi Beta t"

lemma SN_preserved: "\<lbrakk>SN(t1);t1\<longrightarrow>\<^isub>\<beta> t2\<rbrakk>\<Longrightarrow>SN(t2)"
apply(simp add: SN_def)
apply(drule_tac a="t2" in acc_downward)
apply(auto)
done

lemma SN_intro: "(\<forall>t2. t1\<longrightarrow>\<^isub>\<beta>t2 \<longrightarrow> SN(t2))\<Longrightarrow>SN(t1)"
apply(simp add: SN_def)
apply(rule accI)
apply(auto)
done

section {* Candidates *}

consts
  RED :: "ty \<Rightarrow> lam set"

nominal_primrec
  "RED (TVar X) = {t. SN(t)}"
  "RED (\<tau>\<rightarrow>\<sigma>) =   {t. \<forall>u. (u\<in>RED \<tau> \<longrightarrow> (App t u)\<in>RED \<sigma>)}"
apply(rule TrueI)+
done

constdefs
  NEUT :: "lam \<Rightarrow> bool"
  "NEUT t \<equiv> (\<exists>a. t=Var a)\<or>(\<exists>t1 t2. t=App t1 t2)" 

(* a slight hack to get the first element of applications *)
inductive2 FST :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<guillemotright> _" [80,80] 80)
where
  fst[intro!]:  "(App t s) \<guillemotright> t"

lemma fst_elim[elim!]: 
  shows "(App t s) \<guillemotright> t' \<Longrightarrow> t=t'"
apply(ind_cases2 "App t s \<guillemotright> t'")
apply(simp add: lam.inject)
done

lemma qq3: "SN(App t s)\<Longrightarrow>SN(t)"
apply(simp add: SN_def)
apply(subgoal_tac "\<forall>z. (App t s \<guillemotright> z) \<longrightarrow> termi Beta z")(*A*)
apply(force)
(*A*)
apply(erule acc_induct)
apply(clarify)
apply(ind_cases2 "x \<guillemotright> z" for x z)
apply(clarify)
apply(rule accI)
apply(auto intro: b1)
done

section {* Candidates *}

constdefs
  "CR1" :: "ty \<Rightarrow> bool"
  "CR1 \<tau> \<equiv> \<forall> t. (t\<in>RED \<tau> \<longrightarrow> SN(t))"

  "CR2" :: "ty \<Rightarrow> bool"
  "CR2 \<tau> \<equiv> \<forall>t t'. (t\<in>RED \<tau> \<and> t \<longrightarrow>\<^isub>\<beta> t') \<longrightarrow> t'\<in>RED \<tau>"

  "CR3_RED" :: "lam \<Rightarrow> ty \<Rightarrow> bool"
  "CR3_RED t \<tau> \<equiv> \<forall>t'. t\<longrightarrow>\<^isub>\<beta> t' \<longrightarrow>  t'\<in>RED \<tau>" 

  "CR3" :: "ty \<Rightarrow> bool"
  "CR3 \<tau> \<equiv> \<forall>t. (NEUT t \<and> CR3_RED t \<tau>) \<longrightarrow> t\<in>RED \<tau>"
   
  "CR4" :: "ty \<Rightarrow> bool"
  "CR4 \<tau> \<equiv> \<forall>t. (NEUT t \<and> NORMAL t) \<longrightarrow>t\<in>RED \<tau>"

lemma CR3_CR4: "CR3 \<tau> \<Longrightarrow> CR4 \<tau>"
apply(simp (no_asm_use) add: CR3_def CR3_RED_def CR4_def NORMAL_def)
apply(blast)
done

lemma sub_ind: 
  "SN(u)\<Longrightarrow>(u\<in>RED \<tau>\<longrightarrow>(\<forall>t. (NEUT t\<and>CR2 \<tau>\<and>CR3 \<sigma>\<and>CR3_RED t (\<tau>\<rightarrow>\<sigma>))\<longrightarrow>(App t u)\<in>RED \<sigma>))"
apply(simp add: SN_def)
apply(erule acc_induct)
apply(auto)
apply(simp add: CR3_def)
apply(rotate_tac 5)
apply(drule_tac x="App t x" in spec)
apply(drule mp)
apply(rule conjI)
apply(force simp only: NEUT_def)
apply(simp (no_asm) add: CR3_RED_def)
apply(clarify)
apply(ind_cases2 "App t x \<longrightarrow>\<^isub>\<beta> t'" for x t t')
apply(simp_all add: lam.inject)
apply(simp only:  CR3_RED_def)
apply(drule_tac x="s2" in spec)
apply(simp)
apply(drule_tac x="s2" in spec)
apply(simp)
apply(drule mp)
apply(simp (no_asm_use) add: CR2_def)
apply(blast)
apply(drule_tac x="ta" in spec)
apply(force)
apply(auto simp only: NEUT_def lam.inject lam.distinct)
done

lemma RED_props: 
  shows "CR1 \<tau>" and "CR2 \<tau>" and "CR3 \<tau>"
proof (nominal_induct \<tau> rule: ty.induct)
  case (TVar a)
  { case 1 show "CR1 (TVar a)" by (simp add: CR1_def)
  next
    case 2 show "CR2 (TVar a)" by (force intro: SN_preserved simp add: CR2_def)
  next
    case 3 show "CR3 (TVar a)" by (force intro: SN_intro simp add: CR3_def CR3_RED_def)
  }
next
  case (TArr \<tau>1 \<tau>2)
  { case 1
    have ih_CR3_\<tau>1: "CR3 \<tau>1" by fact
    have ih_CR1_\<tau>2: "CR1 \<tau>2" by fact
    show "CR1 (\<tau>1 \<rightarrow> \<tau>2)"
    proof (simp add: CR1_def, intro strip)
      fix t
      assume a: "\<forall>u. u \<in> RED \<tau>1 \<longrightarrow> App t u \<in> RED \<tau>2"
      from ih_CR3_\<tau>1 have "CR4 \<tau>1" by (simp add: CR3_CR4) 
      moreover
      have "NEUT (Var a)" by (force simp add: NEUT_def)
      moreover
      have "NORMAL (Var a)" by (rule NORMAL_Var)
      ultimately have "(Var a)\<in> RED \<tau>1" by (simp add: CR4_def)
      with a have "App t (Var a) \<in> RED \<tau>2" by simp
      hence "SN (App t (Var a))" using ih_CR1_\<tau>2 by (simp add: CR1_def)
      thus "SN(t)" by (rule qq3)
    qed
  next
    case 2
    have ih_CR1_\<tau>1: "CR1 \<tau>1" by fact
    have ih_CR2_\<tau>2: "CR2 \<tau>2" by fact
    show "CR2 (\<tau>1 \<rightarrow> \<tau>2)"
    proof (simp add: CR2_def, intro strip)
      fix t1 t2 u
      assume "(\<forall>u. u \<in> RED \<tau>1 \<longrightarrow> App t1 u \<in> RED \<tau>2) \<and>  t1 \<longrightarrow>\<^isub>\<beta> t2" 
	and  "u \<in> RED \<tau>1"
      hence "t1 \<longrightarrow>\<^isub>\<beta> t2" and "App t1 u \<in> RED \<tau>2" by simp_all
      thus "App t2 u \<in> RED \<tau>2" using ih_CR2_\<tau>2 by (force simp add: CR2_def)
    qed
  next
    case 3
    have ih_CR1_\<tau>1: "CR1 \<tau>1" by fact
    have ih_CR2_\<tau>1: "CR2 \<tau>1" by fact
    have ih_CR3_\<tau>2: "CR3 \<tau>2" by fact
    show "CR3 (\<tau>1 \<rightarrow> \<tau>2)"
    proof (simp add: CR3_def, intro strip)
      fix t u
      assume a1: "u \<in> RED \<tau>1"
      assume a2: "NEUT t \<and> CR3_RED t (\<tau>1 \<rightarrow> \<tau>2)"
      from a1 have "SN(u)" using ih_CR1_\<tau>1 by (simp add: CR1_def)
      hence "u\<in>RED \<tau>1\<longrightarrow>(\<forall>t. (NEUT t\<and>CR2 \<tau>1\<and>CR3 \<tau>2\<and>CR3_RED t (\<tau>1\<rightarrow>\<tau>2))\<longrightarrow>(App t u)\<in>RED \<tau>2)" 
	by (rule sub_ind)
      with a1 a2 show "(App t u)\<in>RED \<tau>2" using ih_CR2_\<tau>1 ih_CR3_\<tau>2 by simp
    qed
  }
qed
    
lemma double_acc_aux:
  assumes a_acc: "acc r a"
  and b_acc: "acc r b"
  and hyp: "\<And>x z.
    (\<And>y. r y x \<Longrightarrow> acc r y) \<Longrightarrow>
    (\<And>y. r y x \<Longrightarrow> P y z) \<Longrightarrow>
    (\<And>u. r u z \<Longrightarrow> acc r u) \<Longrightarrow>
    (\<And>u. r u z \<Longrightarrow> P x u) \<Longrightarrow> P x z"
  shows "P a b"
proof -
  from a_acc
  have r: "\<And>b. acc r b \<Longrightarrow> P a b"
  proof (induct a rule: acc.induct)
    case (accI x)
    note accI' = accI
    have "acc r b" .
    thus ?case
    proof (induct b rule: acc.induct)
      case (accI y)
      show ?case
	apply (rule hyp)
	apply (erule accI')
	apply (erule accI')
	apply (rule acc.accI)
	apply (erule accI)
	apply (erule accI)
	apply (erule accI)
	done
    qed
  qed
  from b_acc show ?thesis by (rule r)
qed

lemma double_acc:
  "\<lbrakk>acc r a; acc r b; \<forall>x z. ((\<forall>y. r y x \<longrightarrow> P y z) \<and> (\<forall>u. r u z \<longrightarrow> P x u)) \<longrightarrow> P x z\<rbrakk> \<Longrightarrow> P a b"
apply(rule_tac r="r" in double_acc_aux)
apply(assumption)+
apply(blast)
done

lemma abs_RED: "(\<forall>s\<in>RED \<tau>. t[x::=s]\<in>RED \<sigma>)\<longrightarrow>Lam [x].t\<in>RED (\<tau>\<rightarrow>\<sigma>)"
apply(simp)
apply(clarify)
apply(subgoal_tac "termi Beta t")(*1*)
apply(erule rev_mp)
apply(subgoal_tac "u \<in> RED \<tau>")(*A*)
apply(erule rev_mp)
apply(rule_tac a="t" and b="u" in double_acc)
apply(assumption)
apply(subgoal_tac "CR1 \<tau>")(*A*)
apply(simp add: CR1_def SN_def)
(*A*)
apply(force simp add: RED_props)
apply(simp)
apply(clarify)
apply(subgoal_tac "CR3 \<sigma>")(*B*)
apply(simp add: CR3_def)
apply(rotate_tac 6)
apply(drule_tac x="App(Lam[x].xa ) z" in spec)
apply(drule mp)
apply(rule conjI)
apply(force simp add: NEUT_def)
apply(simp add: CR3_RED_def)
apply(clarify)
apply(ind_cases2 "App(Lam[x].xa) z \<longrightarrow>\<^isub>\<beta> t'" for xa z t')
apply(auto simp add: lam.inject lam.distinct)
apply(drule beta_abs)
apply(auto)
apply(drule_tac x="t''" in spec)
apply(simp)
apply(drule mp)
apply(clarify)
apply(drule_tac x="s" in bspec)
apply(assumption)
apply(subgoal_tac "xa [ x ::= s ] \<longrightarrow>\<^isub>\<beta>  t'' [ x ::= s ]")(*B*)
apply(subgoal_tac "CR2 \<sigma>")(*C*)
apply(simp (no_asm_use) add: CR2_def)
apply(blast)
(*C*)
apply(force simp add: RED_props)
(*B*)
apply(force intro!: beta_subst)
apply(assumption)
apply(rotate_tac 3)
apply(drule_tac x="s2" in spec)
apply(subgoal_tac "s2\<in>RED \<tau>")(*D*)
apply(simp)
(*D*)
apply(subgoal_tac "CR2 \<tau>")(*E*)
apply(simp (no_asm_use) add: CR2_def)
apply(blast)
(*E*)
apply(force simp add: RED_props)
apply(simp add: alpha)
apply(erule disjE)
apply(force)
apply(auto)
apply(simp add: subst_rename)
apply(drule_tac x="z" in bspec)
apply(assumption)
(*B*)
apply(force simp add: RED_props)
(*1*)
apply(drule_tac x="Var x" in bspec)
apply(subgoal_tac "CR3 \<tau>")(*2*) 
apply(drule CR3_CR4)
apply(simp add: CR4_def)
apply(drule_tac x="Var x" in spec)
apply(drule mp)
apply(rule conjI)
apply(force simp add: NEUT_def)
apply(simp add: NORMAL_def)
apply(clarify)
apply(ind_cases2 "Var x \<longrightarrow>\<^isub>\<beta> t'" for t')
apply(auto simp add: lam.inject lam.distinct)
apply(force simp add: RED_props)
apply(simp add: id_subs)
apply(subgoal_tac "CR1 \<sigma>")(*3*)
apply(simp add: CR1_def SN_def)
(*3*)
apply(force simp add: RED_props)
done

lemma psubst_subst:
  assumes h:"c\<sharp>\<theta>"
  shows "(\<theta><t>)[c::=s] = ((c,s)#\<theta>)<t>"
  using h
by (nominal_induct t avoiding: \<theta> c s rule: lam.induct)
   (auto simp add: fresh_list_cons fresh_atm forget lookup_fresh lookup_fresh')
 

abbreviation 
 mapsto :: "(name\<times>lam) list \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> bool" ("_ maps _ to _" [55,55,55] 55) 
where
 "\<theta> maps x to e\<equiv> (lookup \<theta> x) = e"

abbreviation 
  closes :: "(name\<times>lam) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool" ("_ closes _" [55,55] 55) 
where
  "\<theta> closes \<Gamma> \<equiv> \<forall>x T. ((x,T) \<in> set \<Gamma> \<longrightarrow> (\<exists>t. \<theta> maps x to t \<and> t \<in> RED T))"

lemma all_RED: 
  assumes a: "\<Gamma> \<turnstile> t : \<tau>"
  and     b: "\<theta> closes \<Gamma>"
  shows "\<theta><t> \<in> RED \<tau>"
using a b
proof(nominal_induct t avoiding: \<Gamma> \<tau> \<theta> rule: lam.induct)
  case (Lam a t) --"lambda case"
  have ih: "\<And>\<Gamma> \<tau> \<theta>. \<Gamma> \<turnstile> t : \<tau> \<Longrightarrow> \<theta> closes \<Gamma> \<Longrightarrow> \<theta><t> \<in> RED \<tau>" 
  and  \<theta>_cond: "\<theta> closes \<Gamma>" 
  and fresh: "a\<sharp>\<Gamma>" "a\<sharp>\<theta>" 
  and "\<Gamma> \<turnstile> Lam [a].t:\<tau>" by fact
  hence "\<exists>\<tau>1 \<tau>2. \<tau>=\<tau>1\<rightarrow>\<tau>2 \<and> ((a,\<tau>1)#\<Gamma>)\<turnstile>t:\<tau>2" using t3_elim fresh by simp
  then obtain \<tau>1 \<tau>2 where \<tau>_inst: "\<tau>=\<tau>1\<rightarrow>\<tau>2" and typing: "((a,\<tau>1)#\<Gamma>)\<turnstile>t:\<tau>2" by blast
  from ih have "\<forall>s\<in>RED \<tau>1. \<theta><t>[a::=s] \<in> RED \<tau>2" using fresh typing \<theta>_cond
    by (force dest: fresh_context simp add: psubst_subst)
  hence "(Lam [a].(\<theta><t>)) \<in> RED (\<tau>1 \<rightarrow> \<tau>2)" by (simp only: abs_RED)
  thus "\<theta><(Lam [a].t)> \<in> RED \<tau>" using fresh \<tau>_inst by simp
qed (force dest!: t1_elim t2_elim)+

(* identity substitution generated from a context \<Gamma> *)
consts
  "id" :: "(name\<times>ty) list \<Rightarrow> (name\<times>lam) list"
primrec
  "id []    = []"
  "id (x#\<Gamma>) = ((fst x),Var (fst x))#(id \<Gamma>)"

lemma id_var:
  shows "(id \<Gamma>) maps a to Var a"
apply(induct \<Gamma>, auto)
done

lemma id_fresh:
  fixes a::"name"
  assumes a: "a\<sharp>\<Gamma>"
  shows "a\<sharp>(id \<Gamma>)"
using a
apply(induct \<Gamma>)
apply(auto simp add: fresh_list_nil fresh_list_cons fresh_prod)
done

lemma id_apply:  
  shows "(id \<Gamma>)<t> = t"
apply(nominal_induct t avoiding: \<Gamma> rule: lam.induct)
apply(auto)
apply(rule id_var)
apply(drule id_fresh)+
apply(simp)
done

lemma id_mem:
  assumes a: "(a,\<tau>)\<in>set \<Gamma>"
  shows "lookup (id \<Gamma>) a = Var a"
using a
apply(induct \<Gamma>, auto)
done

lemma id_prop:
  shows "(id \<Gamma>) closes \<Gamma>"
apply(auto)
apply(simp add: id_mem)
apply(subgoal_tac "CR3 T") --"A"
apply(drule CR3_CR4)
apply(simp add: CR4_def)
apply(drule_tac x="Var x" in spec)
apply(force simp add: NEUT_def NORMAL_Var)
--"A"
apply(rule RED_props)
done

lemma typing_implies_RED:  
  assumes a: "\<Gamma>\<turnstile>t:\<tau>"
  shows "t \<in> RED \<tau>"
proof -
  have "(id \<Gamma>)<t>\<in>RED \<tau>" 
  proof -
    have "(id \<Gamma>) closes \<Gamma>" by (rule id_prop)
    with a show ?thesis by (rule all_RED)
  qed
  thus"t \<in> RED \<tau>" by (simp add: id_apply)
qed

lemma typing_implies_SN: 
  assumes a: "\<Gamma>\<turnstile>t:\<tau>"
  shows "SN(t)"
proof -
  from a have "t \<in> RED \<tau>" by (rule typing_implies_RED)
  moreover
  have "CR1 \<tau>" by (rule RED_props)
  ultimately show "SN(t)" by (simp add: CR1_def)
qed

end