diff -r d7183105e304 -r 3c769bf10e63 Nominal/Ex/Lambda.thy --- a/Nominal/Ex/Lambda.thy Tue May 03 13:25:02 2011 +0100 +++ b/Nominal/Ex/Lambda.thy Tue May 03 15:39:30 2011 +0100 @@ -2,6 +2,23 @@ imports "../Nominal2" begin +inductive + even :: "nat \ bool" and + odd :: "nat \ bool" +where + "even 0" +| "odd n \ even (Suc n)" +| "even n \ odd (Suc n)" + +thm even_odd_def + +lemma + shows "p \ (even n) = even (p \ n)" +unfolding even_def +unfolding even_odd_def +apply(perm_simp) +apply(rule refl) +done atom_decl name @@ -10,18 +27,27 @@ | App "lam" "lam" | Lam x::"name" l::"lam" bind x in l ("Lam [_]. _" [100, 100] 100) + inductive triv :: "lam \ nat \ bool" where Var: "triv (Var x) n" +| App: "\triv t1 n; triv t2 n\ \ triv (App t1 t2) n" lemma "p \ (triv t x) = triv (p \ t) (p \ x)" unfolding triv_def apply(perm_simp) +apply(rule refl) oops (*apply(perm_simp)*) +ML {* + Inductive.the_inductive @{context} "Lambda.triv" +*} + +thm triv_def + equivariance triv nominal_inductive triv avoids Var: "{}::name set" apply(auto simp add: fresh_star_def)