| author | chaieb | 
| Wed, 04 Aug 2004 17:43:55 +0200 | |
| changeset 15107 | f233706d9fce | 
| parent 12406 | c9775847ed66 | 
| child 15635 | 8408a06590a6 | 
| permissions | -rw-r--r-- | 
| 3268 
012c43174664
Mostly cosmetic changes: updated headers, ID lines, etc.
 paulson parents: 
3192diff
changeset | 1 | (* Title: Subst/AList.thy | 
| 
012c43174664
Mostly cosmetic changes: updated headers, ID lines, etc.
 paulson parents: 
3192diff
changeset | 2 | ID: $Id$ | 
| 1476 | 3 | Author: Martin Coen, Cambridge University Computer Laboratory | 
| 968 | 4 | Copyright 1993 University of Cambridge | 
| 5 | ||
| 6 | Association lists. | |
| 7 | *) | |
| 8 | ||
| 12406 | 9 | AList = Main + | 
| 968 | 10 | |
| 11 | consts | |
| 12 |   alist_rec  :: "[('a*'b)list, 'c, ['a, 'b, ('a*'b)list, 'c]=>'c] => 'c"
 | |
| 13 |   assoc      :: "['a,'b,('a*'b) list] => 'b"
 | |
| 14 | ||
| 8874 
3242637f668c
alist_rec and assoc are now defined using primrec and thus no longer
 berghofe parents: 
3842diff
changeset | 15 | primrec | 
| 
3242637f668c
alist_rec and assoc are now defined using primrec and thus no longer
 berghofe parents: 
3842diff
changeset | 16 | "alist_rec [] c d = c" | 
| 
3242637f668c
alist_rec and assoc are now defined using primrec and thus no longer
 berghofe parents: 
3842diff
changeset | 17 | "alist_rec (p # al) c d = d (fst p) (snd p) al (alist_rec al c d)" | 
| 968 | 18 | |
| 8874 
3242637f668c
alist_rec and assoc are now defined using primrec and thus no longer
 berghofe parents: 
3842diff
changeset | 19 | primrec | 
| 
3242637f668c
alist_rec and assoc are now defined using primrec and thus no longer
 berghofe parents: 
3842diff
changeset | 20 | "assoc v d [] = d" | 
| 
3242637f668c
alist_rec and assoc are now defined using primrec and thus no longer
 berghofe parents: 
3842diff
changeset | 21 | "assoc v d (p # al) = (if v = fst p then snd p else assoc v d al)" | 
| 968 | 22 | |
| 23 | end |