changeset 3235 | 5ebd327ffb96 |
parent 3231 | 188826f1ccdb |
--- a/Nominal/Ex/Let.thy Mon May 19 11:19:48 2014 +0100 +++ b/Nominal/Ex/Let.thy Mon May 19 12:45:26 2014 +0100 @@ -107,7 +107,7 @@ apply simp_all done -nominal_primrec +nominal_function height_trm :: "trm \<Rightarrow> nat" where "height_trm (Var x) = 1" @@ -160,7 +160,7 @@ apply (simp_all add: trm_assn.bn_defs) done -nominal_primrec +nominal_function subst :: "name \<Rightarrow> trm \<Rightarrow> trm \<Rightarrow> trm" where "subst s t (Var x) = (if (s = x) then t else (Var x))"