author | clasohm |
Wed, 02 Nov 1994 11:50:09 +0100 | |
changeset 156 | fd1be45b64bf |
parent 106 | d27056ec0a5a |
child 194 | b93cc55cb7ab |
permissions | -rw-r--r-- |
(* Title: Substitutions/alist.thy Author: Martin Coen, Cambridge University Computer Laboratory Copyright 1993 University of Cambridge Association lists. *) AList = List + consts alist_rec :: "[('a*'b)list, 'c, ['a, 'b, ('a*'b)list, 'c]=>'c] => 'c" assoc :: "['a,'b,('a*'b) list] => 'b" rules alist_rec_def "alist_rec(al,b,c) == list_rec(al, b, split(c))" assoc_def "assoc(v,d,al) == alist_rec(al,d,%x y xs g.if(v=x,y,g))" end