author | nipkow |
Wed, 08 Feb 1995 11:34:11 +0100 | |
changeset 210 | 1a3d3b5b5d15 |
parent 203 | d465d3be2744 |
child 212 | 2740293cc458 |
permissions | -rw-r--r-- |
196
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
1 |
(* Title: HOL/List.thy |
0 | 2 |
ID: $Id$ |
196
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
3 |
Author: Tobias Nipkow |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
4 |
Copyright 1994 TU Muenchen |
0 | 5 |
|
196
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
6 |
Definition of type 'a list as a datatype. This allows primrec to work. |
0 | 7 |
|
8 |
*) |
|
9 |
||
196
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
10 |
List = Arith + |
0 | 11 |
|
196
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
12 |
datatype 'a list = "[]" ("[]") | "#"('a,'a list) (infixr 65) |
0 | 13 |
|
14 |
consts |
|
15 |
||
113 | 16 |
null :: "'a list => bool" |
17 |
hd :: "'a list => 'a" |
|
18 |
tl,ttl :: "'a list => 'a list" |
|
196
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
19 |
mem :: "['a, 'a list] => bool" (infixl 55) |
113 | 20 |
list_all :: "('a => bool) => ('a list => bool)" |
21 |
map :: "('a=>'b) => ('a list => 'b list)" |
|
196
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
22 |
"@" :: "['a list, 'a list] => 'a list" (infixr 65) |
113 | 23 |
filter :: "['a => bool, 'a list] => 'a list" |
196
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
24 |
foldl :: "[['b,'a] => 'b, 'b, 'a list] => 'b" |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
25 |
length :: "'a list => nat" |
203 | 26 |
nth :: "[nat, 'a list] => 'a" |
0 | 27 |
|
196
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
28 |
syntax |
128 | 29 |
(* list Enumeration *) |
196
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
30 |
"@list" :: "args => 'a list" ("[(_)]") |
0 | 31 |
|
34 | 32 |
(* Special syntax for list_all and filter *) |
39 | 33 |
"@Alls" :: "[idt, 'a list, bool] => bool" ("(2Alls _:_./ _)" 10) |
34 |
"@filter" :: "[idt, 'a list, bool] => 'a list" ("(1[_:_ ./ _])") |
|
0 | 35 |
|
36 |
translations |
|
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
42
diff
changeset
|
37 |
"[x, xs]" == "x#[xs]" |
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
42
diff
changeset
|
38 |
"[x]" == "x#[]" |
42 | 39 |
|
34 | 40 |
"[x:xs . P]" == "filter(%x.P,xs)" |
41 |
"Alls x:xs.P" == "list_all(%x.P,xs)" |
|
0 | 42 |
|
196
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
43 |
primrec null list |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
44 |
null_Nil "null([]) = True" |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
45 |
null_Cons "null(x#xs) = False" |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
46 |
primrec hd list |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
47 |
hd_Nil "hd([]) = (@x.False)" |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
48 |
hd_Cons "hd(x#xs) = x" |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
49 |
primrec tl list |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
50 |
tl_Nil "tl([]) = (@x.False)" |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
51 |
tl_Cons "tl(x#xs) = xs" |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
52 |
primrec ttl list |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
53 |
(* a "total" version of tl: *) |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
54 |
ttl_Nil "ttl([]) = []" |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
55 |
ttl_Cons "ttl(x#xs) = xs" |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
56 |
primrec "op mem" list |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
57 |
mem_Nil "x mem [] = False" |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
58 |
mem_Cons "x mem (y#ys) = if(y=x, True, x mem ys)" |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
59 |
primrec list_all list |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
60 |
list_all_Nil "list_all(P,[]) = True" |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
61 |
list_all_Cons "list_all(P,x#xs) = (P(x) & list_all(P,xs))" |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
62 |
primrec map list |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
63 |
map_Nil "map(f,[]) = []" |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
64 |
map_Cons "map(f,x#xs) = f(x)#map(f,xs)" |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
65 |
primrec "op @" list |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
66 |
append_Nil "[] @ ys = ys" |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
67 |
append_Cons "(x#xs)@ys = x#(xs@ys)" |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
68 |
primrec filter list |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
69 |
filter_Nil "filter(P,[]) = []" |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
70 |
filter_Cons "filter(P,x#xs) = if(P(x), x#filter(P,xs), filter(P,xs))" |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
71 |
primrec foldl list |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
72 |
foldl_Nil "foldl(f,a,[]) = a" |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
73 |
foldl_Cons "foldl(f,a,x#xs) = foldl(f, f(a,x), xs)" |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
74 |
primrec length list |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
75 |
length_Nil "length([]) = 0" |
61620d959717
Moved the old List to ex and replaced it by one defined via
nipkow
parents:
128
diff
changeset
|
76 |
length_Cons "length(x#xs) = Suc(length(xs))" |
203 | 77 |
defs |
78 |
nth_def "nth(n) == nat_rec(n, hd, %m r xs. r(tl(xs)))" |
|
0 | 79 |
end |