Nominal/Ex/Lambda.thy
author Christian Urban <urbanc@in.tum.de>
Tue, 18 Jan 2011 11:02:57 +0100
changeset 2669 1d1772a89026
parent 2667 e3f8673085b1
child 2675 68ccf847507d
permissions -rw-r--r--
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions

theory Lambda
imports "../Nominal2" 
begin

atom_decl name

nominal_datatype lam =
  Var "name"
| App "lam" "lam"
| Lam x::"name" l::"lam"  bind x in l

thm lam.distinct
thm lam.induct
thm lam.exhaust lam.strong_exhaust
thm lam.fv_defs
thm lam.bn_defs
thm lam.perm_simps
thm lam.eq_iff
thm lam.fv_bn_eqvt
thm lam.size_eqvt

nominal_primrec
  depth :: "lam \<Rightarrow> nat"
where
  "depth (Var x) = 1"
| "depth (App t1 t2) = (max (depth t1) (depth t2)) + 1"
| "depth (Lam x t) = (depth t) + 1"
apply(rule_tac y="x" in lam.exhaust)
apply(simp_all)[3]
apply(simp_all only: lam.distinct)
apply(simp add: lam.eq_iff)
apply(simp add: lam.eq_iff)
apply(subst (asm) Abs_eq_iff)
apply(erule exE)
apply(simp add: alphas)
apply(clarify)
apply(rule trans)
apply(rule_tac p="p" in supp_perm_eq[symmetric])
apply(simp add: pure_supp)
apply(simp add: fresh_star_def)
apply(simp add: eqvt_at_def)
done

termination
  apply(relation "measure size")
  apply(simp_all add: lam.size)
  done
  
lemma removeAll_eqvt[eqvt]:
  shows "p \<bullet> (removeAll x xs) = removeAll (p \<bullet> x) (p \<bullet> xs)"
by (induct xs) (auto)

lemma supp_removeAll:
  fixes x::"atom"
  shows "supp (removeAll x xs) = (supp xs - {x})"
apply(induct xs)
apply(simp_all add: supp_Nil supp_Cons)
apply(rule conjI)
apply(rule impI)
apply(simp add: supp_atom)
apply(rule impI)
apply(simp add: supp_atom)
apply(blast)
done

nominal_primrec 
  frees_lst :: "lam \<Rightarrow> atom list"
where
  "frees_lst (Var x) = [atom x]"
| "frees_lst (App t1 t2) = (frees_lst t1) @ (frees_lst t2)"
| "frees_lst (Lam x t) = removeAll (atom x) (frees_lst t)"
apply(rule_tac y="x" in lam.exhaust)
apply(simp_all)[3]
apply(simp_all only: lam.distinct)
apply(simp add: lam.eq_iff)
apply(simp add: lam.eq_iff)
apply(simp add: lam.eq_iff)
apply(simp add: Abs_eq_iff)
apply(erule exE)
apply(simp add: alphas)
apply(simp add: atom_eqvt)
apply(clarify)
apply(rule trans)
apply(rule_tac p="p" in supp_perm_eq[symmetric])
apply(simp (no_asm) add: supp_removeAll)
apply(drule supp_eqvt_at)
apply(simp add: finite_supp)
apply(auto simp add: fresh_star_def)[1]
unfolding eqvt_at_def
apply(simp only: removeAll_eqvt atom_eqvt)
done

termination
  apply(relation "measure size")
  apply(simp_all add: lam.size)
  done

(* a small lemma *)
lemma
  "supp t = set (frees_lst t)"
apply(induct t rule: lam.induct)
apply(simp_all add: lam.supp supp_at_base)
done

nominal_datatype ln = 
  LNBnd nat
| LNVar name
| LNApp ln ln
| LNLam ln

fun
  lookup :: "name list \<Rightarrow> nat \<Rightarrow> name \<Rightarrow> ln" 
where
  "lookup [] n x = LNVar x"
| "lookup (y # ys) n x = (if x = y then LNBnd n else (lookup ys (n + 1) x))"

lemma [eqvt]:
  shows "(p \<bullet> lookup xs n x) = lookup (p \<bullet> xs) (p \<bullet> n) (p \<bullet> x)"
apply(induct xs arbitrary: n)
apply(simp_all add: permute_pure)
done

nominal_primrec
  trans :: "lam \<Rightarrow> name list \<Rightarrow> ln"
where
  "trans (Var x) xs = lookup xs 0 x"
| "trans (App t1 t2) xs = LNApp (trans t1 xs) (trans t2 xs)"
| "atom x \<sharp> xs \<Longrightarrow> trans (Lam x t) xs = LNLam (trans t (x # xs))"
apply(case_tac x)
apply(simp)
apply(rule_tac y="a" and c="b" in lam.strong_exhaust)
apply(simp_all)[3]
apply(blast)
apply(blast)
apply(simp add: fresh_star_def)
apply(simp_all add: lam.distinct)
apply(simp add: lam.eq_iff)
apply(simp add: lam.eq_iff)
apply(simp add: lam.eq_iff)
apply(simp add: Abs_eq_iff)
apply(erule conjE)
apply(erule exE)
apply(simp add: alphas)
apply(simp add: atom_eqvt)
apply(clarify)
apply(rule trans)
apply(rule_tac p="p" in supp_perm_eq[symmetric])
apply(simp (no_asm) add: ln.supp)
apply(drule supp_eqvt_at)
apply(simp add: finite_supp)
oops





nominal_datatype db = 
  DBVar nat
| DBApp db db
| DBLam db

abbreviation
  mbind :: "'a option => ('a => 'b option) => 'b option"  ("_ \<guillemotright>= _" [65,65] 65) 
where  
  "c \<guillemotright>= f \<equiv> case c of None => None | (Some v) => f v"

lemma mbind_eqvt:
  fixes c::"'a::pt option"
  shows "(p \<bullet> (c \<guillemotright>= f)) = ((p \<bullet> c) \<guillemotright>= (p \<bullet> f))"
apply(cases c)
apply(simp_all)
apply(perm_simp)
apply(rule refl)
done

lemma mbind_eqvt_raw[eqvt_raw]:
  shows "(p \<bullet> option_case) \<equiv> option_case"
apply(rule eq_reflection)
apply(rule ext)+
apply(case_tac xb)
apply(simp_all)
apply(rule_tac p="-p" in permute_boolE)
apply(perm_simp add: permute_minus_cancel)
apply(simp)
apply(rule_tac p="-p" in permute_boolE)
apply(perm_simp add: permute_minus_cancel)
apply(simp)
done

fun
  index :: "atom list \<Rightarrow> nat \<Rightarrow> atom \<Rightarrow> nat option" 
where
  "index [] n x = None"
| "index (y # ys) n x = (if x = y then (Some n) else (index ys (n + 1) x))"

lemma [eqvt]:
  shows "(p \<bullet> index xs n x) = index (p \<bullet> xs) (p \<bullet> n) (p \<bullet> x)"
apply(induct xs arbitrary: n)
apply(simp_all add: permute_pure)
done

ML {*
Nominal_Function_Core.trace := true
*}

(*
inductive
  trans_graph
where
  "trans_graph (Var x, xs) (index xs 0 (atom x) \<guillemotright>= (\<lambda>v. Some (DBVar v)))" 
| "\<lbrakk>trans_graph (t1, xs) (trans_sum (t1, xs)); 
    \<And>a. trans_sum (t1, xs) = Some a \<Longrightarrow> trans_graph (t2, xs) (trans_sum (t2, xs))\<rbrakk>
    \<Longrightarrow> trans_graph (App t1 t2, xs) 
       (trans_sum (t1, xs) \<guillemotright>= (\<lambda>v. trans_sum (t2, xs) \<guillemotright>= (\<lambda>va. Some (DBApp v va))))"
| "trans_graph (t, atom x # xs) (trans_sum (t, atom x # xs)) \<Longrightarrow>
    trans_graph (Lam x t, xs) (trans_sum (t, atom x # xs) \<guillemotright>= (\<lambda>v. Some (DBLam v)))"

lemma
  assumes a: "trans_graph x t"
  shows "trans_graph (p \<bullet> x) (p \<bullet> t)"
using a
apply(induct)
apply(perm_simp)
apply(rule trans_graph.intros)
apply(perm_simp)
apply(rule trans_graph.intros)
apply(simp)
apply(simp)
defer
apply(perm_simp)
apply(rule trans_graph.intros)
apply(simp)
apply(rotate_tac 3)
apply(drule_tac x="FOO" in meta_spec)
apply(drule meta_mp)
prefer 2
apply(simp)

equivariance trans_graph
*)

(* equivariance fails at the moment
nominal_primrec
  trans :: "lam \<Rightarrow> atom list \<Rightarrow> db option"
where
  "trans (Var x) xs = (index xs 0 (atom x) \<guillemotright>= (\<lambda>n. Some (DBVar n)))"
| "trans (App t1 t2) xs = (trans t1 xs \<guillemotright>= (\<lambda>db1. trans t2 xs \<guillemotright>= (\<lambda>db2. Some (DBApp db1 db2))))"
| "trans (Lam x t) xs = (trans t (atom x # xs) \<guillemotright>= (\<lambda>db. Some (DBLam db)))"
*)

nominal_primrec
  subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam"  ("_ [_ ::= _]" [100,100,100] 100)
where
  "(Var x)[y ::= s] = (if x=y then s else (Var x))"
| "(App t\<^isub>1 t\<^isub>2)[y ::= s] = App (t\<^isub>1[y ::= s]) (t\<^isub>2[y ::= s])"
| "atom x \<sharp> (y, s) \<Longrightarrow> (Lam x t)[y ::= s] = Lam x (t[y ::= s])"
apply(case_tac x)
apply(simp)
apply(rule_tac y="a" and c="(b, c)" in lam.strong_exhaust)
apply(simp add: lam.eq_iff lam.distinct)
apply(auto)[1]
apply(simp add: lam.eq_iff lam.distinct)
apply(auto)[1]
apply(simp add: fresh_star_def lam.eq_iff lam.distinct)
apply(simp_all add: lam.distinct)[5]
apply(simp add: lam.eq_iff)
apply(simp add: lam.eq_iff)
apply(simp add: lam.eq_iff)
apply(erule conjE)+
oops


end