author | nipkow |
Thu, 26 Jun 1997 13:20:50 +0200 | |
changeset 3465 | e85c24717cad |
parent 3437 | bea2faf1641d |
child 3507 | 157be29ad5ba |
permissions | -rw-r--r-- |
923 | 1 |
(* Title: HOL/List.thy |
2 |
ID: $Id$ |
|
3 |
Author: Tobias Nipkow |
|
4 |
Copyright 1994 TU Muenchen |
|
5 |
||
2512 | 6 |
The datatype of finite lists. |
923 | 7 |
*) |
8 |
||
3367 | 9 |
List = Divides + |
923 | 10 |
|
977
5d57287e5e1e
changed syntax of datatype declarations (curried types for constructor
clasohm
parents:
965
diff
changeset
|
11 |
datatype 'a list = "[]" ("[]") | "#" 'a ('a list) (infixr 65) |
923 | 12 |
|
13 |
consts |
|
1908 | 14 |
"@" :: ['a list, 'a list] => 'a list (infixr 65) |
15 |
filter :: ['a => bool, 'a list] => 'a list |
|
2608 | 16 |
concat :: 'a list list => 'a list |
1908 | 17 |
foldl :: [['b,'a] => 'b, 'b, 'a list] => 'b |
18 |
hd :: 'a list => 'a |
|
3465 | 19 |
set :: 'a list => 'a set |
1908 | 20 |
list_all :: ('a => bool) => ('a list => bool) |
21 |
map :: ('a=>'b) => ('a list => 'b list) |
|
22 |
mem :: ['a, 'a list] => bool (infixl 55) |
|
23 |
nth :: [nat, 'a list] => 'a |
|
2608 | 24 |
take, drop :: [nat, 'a list] => 'a list |
25 |
takeWhile, |
|
26 |
dropWhile :: ('a => bool) => 'a list => 'a list |
|
1908 | 27 |
tl,ttl :: 'a list => 'a list |
28 |
rev :: 'a list => 'a list |
|
923 | 29 |
|
30 |
syntax |
|
31 |
(* list Enumeration *) |
|
2262 | 32 |
"@list" :: args => 'a list ("[(_)]") |
923 | 33 |
|
2512 | 34 |
(* Special syntax for filter *) |
2262 | 35 |
"@filter" :: [idt, 'a list, bool] => 'a list ("(1[_:_ ./ _])") |
923 | 36 |
|
37 |
translations |
|
38 |
"[x, xs]" == "x#[xs]" |
|
39 |
"[x]" == "x#[]" |
|
1475 | 40 |
"[x:xs . P]" == "filter (%x.P) xs" |
923 | 41 |
|
2262 | 42 |
syntax (symbols) |
43 |
"@filter" :: [idt, 'a list, bool] => 'a list ("(1[_\\<in>_ ./ _])") |
|
44 |
||
45 |
||
3342
ec3b55fcb165
New operator "lists" for formalizing sets of lists
paulson
parents:
3320
diff
changeset
|
46 |
consts |
ec3b55fcb165
New operator "lists" for formalizing sets of lists
paulson
parents:
3320
diff
changeset
|
47 |
lists :: 'a set => 'a list set |
ec3b55fcb165
New operator "lists" for formalizing sets of lists
paulson
parents:
3320
diff
changeset
|
48 |
|
ec3b55fcb165
New operator "lists" for formalizing sets of lists
paulson
parents:
3320
diff
changeset
|
49 |
inductive "lists A" |
ec3b55fcb165
New operator "lists" for formalizing sets of lists
paulson
parents:
3320
diff
changeset
|
50 |
intrs |
ec3b55fcb165
New operator "lists" for formalizing sets of lists
paulson
parents:
3320
diff
changeset
|
51 |
Nil "[]: lists A" |
ec3b55fcb165
New operator "lists" for formalizing sets of lists
paulson
parents:
3320
diff
changeset
|
52 |
Cons "[| a: A; l: lists A |] ==> a#l : lists A" |
ec3b55fcb165
New operator "lists" for formalizing sets of lists
paulson
parents:
3320
diff
changeset
|
53 |
|
ec3b55fcb165
New operator "lists" for formalizing sets of lists
paulson
parents:
3320
diff
changeset
|
54 |
|
3437
bea2faf1641d
Replacing the primrec definition of "length" by a translation to the built-in
paulson
parents:
3401
diff
changeset
|
55 |
(*Function "size" is overloaded for all datatypes. Users may refer to the |
bea2faf1641d
Replacing the primrec definition of "length" by a translation to the built-in
paulson
parents:
3401
diff
changeset
|
56 |
list version as "length".*) |
bea2faf1641d
Replacing the primrec definition of "length" by a translation to the built-in
paulson
parents:
3401
diff
changeset
|
57 |
syntax length :: 'a list => nat |
bea2faf1641d
Replacing the primrec definition of "length" by a translation to the built-in
paulson
parents:
3401
diff
changeset
|
58 |
translations "length" => "size" |
bea2faf1641d
Replacing the primrec definition of "length" by a translation to the built-in
paulson
parents:
3401
diff
changeset
|
59 |
|
923 | 60 |
primrec hd list |
3320 | 61 |
"hd([]) = arbitrary" |
1898 | 62 |
"hd(x#xs) = x" |
923 | 63 |
primrec tl list |
3320 | 64 |
"tl([]) = arbitrary" |
1898 | 65 |
"tl(x#xs) = xs" |
923 | 66 |
primrec ttl list |
67 |
(* a "total" version of tl: *) |
|
1898 | 68 |
"ttl([]) = []" |
69 |
"ttl(x#xs) = xs" |
|
923 | 70 |
primrec "op mem" list |
1898 | 71 |
"x mem [] = False" |
72 |
"x mem (y#ys) = (if y=x then True else x mem ys)" |
|
3465 | 73 |
primrec set list |
74 |
"set [] = {}" |
|
75 |
"set (x#xs) = insert x (set xs)" |
|
923 | 76 |
primrec list_all list |
77 |
list_all_Nil "list_all P [] = True" |
|
78 |
list_all_Cons "list_all P (x#xs) = (P(x) & list_all P xs)" |
|
79 |
primrec map list |
|
1898 | 80 |
"map f [] = []" |
81 |
"map f (x#xs) = f(x)#map f xs" |
|
923 | 82 |
primrec "op @" list |
1898 | 83 |
"[] @ ys = ys" |
84 |
"(x#xs)@ys = x#(xs@ys)" |
|
1169 | 85 |
primrec rev list |
1898 | 86 |
"rev([]) = []" |
87 |
"rev(x#xs) = rev(xs) @ [x]" |
|
923 | 88 |
primrec filter list |
1898 | 89 |
"filter P [] = []" |
90 |
"filter P (x#xs) = (if P x then x#filter P xs else filter P xs)" |
|
923 | 91 |
primrec foldl list |
1898 | 92 |
"foldl f a [] = a" |
93 |
"foldl f a (x#xs) = foldl f (f a x) xs" |
|
2608 | 94 |
primrec concat list |
95 |
"concat([]) = []" |
|
96 |
"concat(x#xs) = x @ concat(xs)" |
|
1419
a6a034a47a71
defined take/drop by induction over list rather than nat.
nipkow
parents:
1370
diff
changeset
|
97 |
primrec drop list |
a6a034a47a71
defined take/drop by induction over list rather than nat.
nipkow
parents:
1370
diff
changeset
|
98 |
drop_Nil "drop n [] = []" |
a6a034a47a71
defined take/drop by induction over list rather than nat.
nipkow
parents:
1370
diff
changeset
|
99 |
drop_Cons "drop n (x#xs) = (case n of 0 => x#xs | Suc(m) => drop m xs)" |
a6a034a47a71
defined take/drop by induction over list rather than nat.
nipkow
parents:
1370
diff
changeset
|
100 |
primrec take list |
a6a034a47a71
defined take/drop by induction over list rather than nat.
nipkow
parents:
1370
diff
changeset
|
101 |
take_Nil "take n [] = []" |
a6a034a47a71
defined take/drop by induction over list rather than nat.
nipkow
parents:
1370
diff
changeset
|
102 |
take_Cons "take n (x#xs) = (case n of 0 => [] | Suc(m) => x # take m xs)" |
2738 | 103 |
primrec nth nat |
104 |
"nth 0 xs = hd xs" |
|
105 |
"nth (Suc n) xs = nth n (tl xs)" |
|
2608 | 106 |
primrec takeWhile list |
107 |
"takeWhile P [] = []" |
|
108 |
"takeWhile P (x#xs) = (if P x then x#takeWhile P xs else [])" |
|
109 |
primrec dropWhile list |
|
110 |
"dropWhile P [] = []" |
|
111 |
"dropWhile P (x#xs) = (if P x then dropWhile P xs else xs)" |
|
3196 | 112 |
|
923 | 113 |
end |