diff -r 7bc38b93a1fc -r 9ae285ce814e Nominal/Nominal2_Base.thy --- a/Nominal/Nominal2_Base.thy Sun Apr 06 13:07:24 2014 +0100 +++ b/Nominal/Nominal2_Base.thy Fri May 16 08:38:23 2014 +0100 @@ -2951,13 +2951,13 @@ simproc_setup fresh_ineq ("x \ (y::'a::at_base)") = {* fn _ => fn ctxt => fn ctrm => - case term_of ctrm of @{term "HOL.Not"} $ (Const ("HOL.eq", _) $ lhs $ rhs) => + case term_of ctrm of @{term "HOL.Not"} $ (Const (@{const_name HOL.eq}, _) $ lhs $ rhs) => let fun first_is_neg lhs rhs [] = NONE | first_is_neg lhs rhs (thm::thms) = (case Thm.prop_of thm of - _ $ (@{term "HOL.Not"} $ (Const ("HOL.eq", _) $ l $ r)) => + _ $ (@{term "HOL.Not"} $ (Const (@{const_name HOL.eq}, _) $ l $ r)) => (if l = lhs andalso r = rhs then SOME(thm) else if r = lhs andalso l = rhs then SOME(thm RS @{thm not_sym}) else first_is_neg lhs rhs thms)