diff -r 55b156ac4a40 -r 6da7d538e5e0 LamEx.thy --- a/LamEx.thy Fri Oct 30 15:32:04 2009 +0100 +++ b/LamEx.thy Fri Oct 30 15:35:42 2009 +0100 @@ -11,12 +11,6 @@ | rApp "rlam" "rlam" | rLam "name" "rlam" -inductive - alpha :: "rlam \ rlam \ bool" ("_ \ _" [100, 100] 100) -where - a1: "a = b \ (rVar a) \ (rVar b)" -| 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" @@ -28,6 +22,13 @@ termination rfv sorry +inductive + alpha :: "rlam \ rlam \ bool" ("_ \ _" [100, 100] 100) +where + a1: "a = b \ (rVar a) \ (rVar b)" +| a2: "\t1 \ t2; s1 \ s2\ \ rApp t1 s1 \ rApp t2 s2" +| a3: "\t \ ([(a,b)]\s); a \ rfv (rLam b t)\ \ rLam a t \ rLam b s" + quotient lam = rlam / alpha sorry