diff -r 5ebd327ffb96 -r e2da10806a34 Nominal/Ex/CPS/CPS2_DanvyNielsen.thy --- a/Nominal/Ex/CPS/CPS2_DanvyNielsen.thy Mon May 19 12:45:26 2014 +0100 +++ b/Nominal/Ex/CPS/CPS2_DanvyNielsen.thy Mon May 19 16:45:46 2014 +0100 @@ -32,7 +32,7 @@ apply (simp add: flip_fresh_fresh) done -termination (eqvt) by lexicographic_order +nominal_termination (eqvt) by lexicographic_order nominal_function ccomp :: "cpsctxt => cpsctxt => cpsctxt" @@ -57,7 +57,7 @@ apply (simp add: flip_fresh_fresh) done -termination (eqvt) by lexicographic_order +nominal_termination (eqvt) by lexicographic_order nominal_function CPSv :: "lt => lt"