diff -r 7ee9a2fefc77 -r 1bddffddc03f Nominal/Equivp.thy --- a/Nominal/Equivp.thy Sat May 01 09:15:46 2010 +0100 +++ b/Nominal/Equivp.thy Sun May 02 14:06:26 2010 +0100 @@ -1,5 +1,5 @@ theory Equivp -imports "Fv" +imports "NewFv" "Tacs" "Rsp" "NewFv" begin ML {* @@ -188,7 +188,7 @@ val rhs = list_comb (cnstr, frees) fun mk_supp_arg (x, ty) = - if is_atom thy ty then mk_supp @{typ atom} (mk_atom ty $ x) else + if is_atom thy ty then mk_supp @{typ atom} (mk_atom_ty ty x) else if is_atom_set thy ty then mk_supp @{typ "atom set"} (mk_atom_set x) else if is_atom_fset thy ty then mk_supp @{typ "atom set"} (mk_atom_fset x) else mk_supp ty x