src/HOL/Subst/AList.thy
author krauss
Sun, 21 Aug 2011 22:13:04 +0200
changeset 44367 74c08021ab2e
parent 38140 05691ad74079
permissions -rw-r--r--
changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
38140
05691ad74079 modernized specifications;
wenzelm
parents: 24823
diff changeset
     1
(*  Title:      HOL/Subst/AList.thy
1476
608483c2122a expanded tabs; incorporated Konrad's changes
clasohm
parents: 968
diff changeset
     2
    Author:     Martin Coen, Cambridge University Computer Laboratory
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
     3
    Copyright   1993  University of Cambridge
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
     4
*)
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
     5
38140
05691ad74079 modernized specifications;
wenzelm
parents: 24823
diff changeset
     6
header {* Association Lists *}
15635
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
     7
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
     8
theory AList
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
     9
imports Main
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    10
begin
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    11
38140
05691ad74079 modernized specifications;
wenzelm
parents: 24823
diff changeset
    12
primrec alist_rec :: "[('a*'b)list, 'c, ['a, 'b, ('a*'b)list, 'c]=>'c] => 'c"
05691ad74079 modernized specifications;
wenzelm
parents: 24823
diff changeset
    13
where
8874
3242637f668c alist_rec and assoc are now defined using primrec and thus no longer
berghofe
parents: 3842
diff changeset
    14
  "alist_rec [] c d = c"
38140
05691ad74079 modernized specifications;
wenzelm
parents: 24823
diff changeset
    15
| "alist_rec (p # al) c d = d (fst p) (snd p) al (alist_rec al c d)"
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    16
38140
05691ad74079 modernized specifications;
wenzelm
parents: 24823
diff changeset
    17
primrec assoc :: "['a,'b,('a*'b) list] => 'b"
05691ad74079 modernized specifications;
wenzelm
parents: 24823
diff changeset
    18
where
8874
3242637f668c alist_rec and assoc are now defined using primrec and thus no longer
berghofe
parents: 3842
diff changeset
    19
  "assoc v d [] = d"
38140
05691ad74079 modernized specifications;
wenzelm
parents: 24823
diff changeset
    20
| "assoc v d (p # al) = (if v = fst p then snd p else assoc v d al)"
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    21
15635
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    22
lemma alist_induct:
38140
05691ad74079 modernized specifications;
wenzelm
parents: 24823
diff changeset
    23
    "P [] \<Longrightarrow> (!!x y xs. P xs \<Longrightarrow> P ((x,y) # xs)) \<Longrightarrow> P l"
24823
bfb619994060 modernized specifications;
wenzelm
parents: 15635
diff changeset
    24
  by (induct l) auto
15635
8408a06590a6 converted HOL-Subst to tactic scripts
paulson
parents: 12406
diff changeset
    25
44367
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 38140
diff changeset
    26
74c08021ab2e changed constant names and notation to match HOL/Subst/*.thy, from which this theory is a clone.
krauss
parents: 38140
diff changeset
    27
968
3cdaa8724175 converted Subst with curried function application
clasohm
parents:
diff changeset
    28
end