| author | haftmann |
| Thu, 20 Mar 2008 12:04:53 +0100 | |
| changeset 26357 | 19b153ebda0b |
| parent 24823 | bfb619994060 |
| child 38140 | 05691ad74079 |
| permissions | -rw-r--r-- |
| 15635 | 1 |
(* ID: $Id$ |
| 1476 | 2 |
Author: Martin Coen, Cambridge University Computer Laboratory |
| 968 | 3 |
Copyright 1993 University of Cambridge |
4 |
||
5 |
*) |
|
6 |
||
| 15635 | 7 |
header{*Association Lists*}
|
8 |
||
9 |
theory AList |
|
10 |
imports Main |
|
11 |
begin |
|
| 968 | 12 |
|
| 24823 | 13 |
consts alist_rec :: "[('a*'b)list, 'c, ['a, 'b, ('a*'b)list, 'c]=>'c] => 'c"
|
|
8874
3242637f668c
alist_rec and assoc are now defined using primrec and thus no longer
berghofe
parents:
3842
diff
changeset
|
14 |
primrec |
|
3242637f668c
alist_rec and assoc are now defined using primrec and thus no longer
berghofe
parents:
3842
diff
changeset
|
15 |
"alist_rec [] c d = c" |
|
3242637f668c
alist_rec and assoc are now defined using primrec and thus no longer
berghofe
parents:
3842
diff
changeset
|
16 |
"alist_rec (p # al) c d = d (fst p) (snd p) al (alist_rec al c d)" |
| 968 | 17 |
|
| 24823 | 18 |
consts assoc :: "['a,'b,('a*'b) list] => 'b"
|
|
8874
3242637f668c
alist_rec and assoc are now defined using primrec and thus no longer
berghofe
parents:
3842
diff
changeset
|
19 |
primrec |
|
3242637f668c
alist_rec and assoc are now defined using primrec and thus no longer
berghofe
parents:
3842
diff
changeset
|
20 |
"assoc v d [] = d" |
|
3242637f668c
alist_rec and assoc are now defined using primrec and thus no longer
berghofe
parents:
3842
diff
changeset
|
21 |
"assoc v d (p # al) = (if v = fst p then snd p else assoc v d al)" |
| 968 | 22 |
|
| 15635 | 23 |
lemma alist_induct: |
24 |
"[| P([]); |
|
25 |
!!x y xs. P(xs) ==> P((x,y)#xs) |] ==> P(l)" |
|
| 24823 | 26 |
by (induct l) auto |
| 15635 | 27 |
|
| 968 | 28 |
end |