diff -r 60acf3d3a4a0 -r 22715cab3995 LamEx.thy --- a/LamEx.thy Fri Oct 30 14:25:37 2009 +0100 +++ b/LamEx.thy Fri Oct 30 15:28:44 2009 +0100 @@ -4,6 +4,8 @@ atom_decl name +thm abs_fresh(1) + nominal_datatype rlam = rVar "name" | rApp "rlam" "rlam" @@ -16,6 +18,16 @@ | a2: "\t1 \ t2; s1 \ s2\ \ rApp t1 s1 \ rApp t2 s2" | a3: "\t \ ([(a,b)]\s); a\[b].s\ \ rLam a t \ rLam b s" +function + rfv :: "rlam \ name set" +where + rfv_var: "rfv (rVar a) = {a}" +| rfv_app: "rfv (rApp t1 t2) = (rfv t1) \ (rfv t2)" +| rfv_lam: "rfv (rLam a t) = (rfv t) - {a}" +sorry + +termination rfv sorry + quotient lam = rlam / alpha sorry @@ -40,6 +52,13 @@ thm App_def thm Lam_def +quotient_def (for lam) + fv :: "lam \ name set" +where + "fv \ rfv" + +thm fv_def + (* definition of overloaded permutation function *) (* for the lifted type lam *) overloading @@ -77,8 +96,20 @@ shows "(pi\rLam a t) \ rLam (pi\a) (pi\t)" sorry +lemma fv_var: + shows "fv (Var a) = {a}" +sorry + +lemma fv_app: + shows "fv (App t1 t2) = (fv t1) \ (fv t2)" +sorry + +lemma fv_lam: + shows "fv (Lam a t) = (fv t) - {a}" +sorry + lemma real_alpha: - assumes "t = ([(a,b)]\s)" "a\s" + assumes "t = ([(a,b)]\s)" "a\[b].s" shows "Lam a t = Lam b s" sorry