# HG changeset patch # User Christian Urban # Date 1271690767 -7200 # Node ID f8c8e2afcc18f2e39e00d438481569105efc626c # Parent 3e92714ce76aabac82cfbcea8ec1e034b402b508 deleting function perm_arg in favour of the library function mk_perm diff -r 3e92714ce76a -r f8c8e2afcc18 Nominal/Equivp.thy --- a/Nominal/Equivp.thy Mon Apr 19 16:55:57 2010 +0200 +++ b/Nominal/Equivp.thy Mon Apr 19 17:26:07 2010 +0200 @@ -346,7 +346,7 @@ val typ = domain_type (fastype_of fnctn); val arg = the (AList.lookup (op=) frees typ); in - ([HOLogic.mk_eq ((perm_arg (fnctn $ arg) $ pi $ (fnctn $ arg)), (fnctn $ (perm_arg arg $ pi $ arg)))], ctxt) + ([HOLogic.mk_eq (mk_perm pi (fnctn $ arg), fnctn $ (mk_perm pi arg))], ctxt) end *} diff -r 3e92714ce76a -r f8c8e2afcc18 Nominal/Rsp.thy --- a/Nominal/Rsp.thy Mon Apr 19 16:55:57 2010 +0200 +++ b/Nominal/Rsp.thy Mon Apr 19 17:26:07 2010 +0200 @@ -87,14 +87,6 @@ )) *} -ML {* -fun perm_arg arg = -let - val ty = fastype_of arg -in - Const (@{const_name permute}, @{typ perm} --> ty --> ty) -end -*} lemma exi: "\(pi :: perm). P pi \ (\(p :: perm). P p \ Q (pi \ p)) \ \pi. Q pi" apply (erule exE) @@ -123,7 +115,7 @@ val (tys, _) = strip_type (fastype_of alpha) val indnames = Name.variant_list names (Datatype_Prop.make_tnames (map body_type tys)); val args = map Free (indnames ~~ tys); - val perm_args = map (fn x => perm_arg x $ pi $ x) args + val perm_args = map (fn x => mk_perm pi x) args in (HOLogic.mk_imp (list_comb (alpha, args), list_comb (alpha, perm_args)), indnames @ names) end