author | paulson |
Tue, 01 Oct 1996 15:49:29 +0200 | |
changeset 2048 | bb54fbba0071 |
parent 1908 | 55d8e38262a8 |
child 2262 | c7ee913746fd |
permissions | -rw-r--r-- |
923 | 1 |
(* Title: HOL/List.thy |
2 |
ID: $Id$ |
|
3 |
Author: Tobias Nipkow |
|
4 |
Copyright 1994 TU Muenchen |
|
5 |
||
6 |
Definition of type 'a list as a datatype. This allows primrec to work. |
|
7 |
||
1812 | 8 |
TODO: delete list_all from this theory and from ex/Sorting, etc. |
1908 | 9 |
use set_of_list instead |
923 | 10 |
*) |
11 |
||
12 |
List = Arith + |
|
13 |
||
977
5d57287e5e1e
changed syntax of datatype declarations (curried types for constructor
clasohm
parents:
965
diff
changeset
|
14 |
datatype 'a list = "[]" ("[]") | "#" 'a ('a list) (infixr 65) |
923 | 15 |
|
16 |
consts |
|
17 |
||
1908 | 18 |
"@" :: ['a list, 'a list] => 'a list (infixr 65) |
19 |
drop :: [nat, 'a list] => 'a list |
|
20 |
filter :: ['a => bool, 'a list] => 'a list |
|
21 |
flat :: 'a list list => 'a list |
|
22 |
foldl :: [['b,'a] => 'b, 'b, 'a list] => 'b |
|
23 |
hd :: 'a list => 'a |
|
24 |
length :: 'a list => nat |
|
25 |
set_of_list :: ('a list => 'a set) |
|
26 |
list_all :: ('a => bool) => ('a list => bool) |
|
27 |
map :: ('a=>'b) => ('a list => 'b list) |
|
28 |
mem :: ['a, 'a list] => bool (infixl 55) |
|
29 |
nth :: [nat, 'a list] => 'a |
|
30 |
take :: [nat, 'a list] => 'a list |
|
31 |
tl,ttl :: 'a list => 'a list |
|
32 |
rev :: 'a list => 'a list |
|
923 | 33 |
|
34 |
syntax |
|
35 |
(* list Enumeration *) |
|
1908 | 36 |
"@list" :: args => 'a list ("[(_)]") |
923 | 37 |
|
38 |
(* Special syntax for list_all and filter *) |
|
1908 | 39 |
"@Alls" :: [idt, 'a list, bool] => bool ("(2Alls _:_./ _)" 10) |
40 |
"@filter" :: [idt, 'a list, bool] => 'a list ("(1[_:_ ./ _])") |
|
923 | 41 |
|
42 |
translations |
|
43 |
"[x, xs]" == "x#[xs]" |
|
44 |
"[x]" == "x#[]" |
|
45 |
||
1475 | 46 |
"[x:xs . P]" == "filter (%x.P) xs" |
47 |
"Alls x:xs.P" == "list_all (%x.P) xs" |
|
923 | 48 |
|
49 |
primrec hd list |
|
1898 | 50 |
"hd([]) = (@x.False)" |
51 |
"hd(x#xs) = x" |
|
923 | 52 |
primrec tl list |
1898 | 53 |
"tl([]) = (@x.False)" |
54 |
"tl(x#xs) = xs" |
|
923 | 55 |
primrec ttl list |
56 |
(* a "total" version of tl: *) |
|
1898 | 57 |
"ttl([]) = []" |
58 |
"ttl(x#xs) = xs" |
|
923 | 59 |
primrec "op mem" list |
1898 | 60 |
"x mem [] = False" |
61 |
"x mem (y#ys) = (if y=x then True else x mem ys)" |
|
1908 | 62 |
primrec set_of_list list |
63 |
"set_of_list [] = {}" |
|
64 |
"set_of_list (x#xs) = insert x (set_of_list xs)" |
|
923 | 65 |
primrec list_all list |
66 |
list_all_Nil "list_all P [] = True" |
|
67 |
list_all_Cons "list_all P (x#xs) = (P(x) & list_all P xs)" |
|
68 |
primrec map list |
|
1898 | 69 |
"map f [] = []" |
70 |
"map f (x#xs) = f(x)#map f xs" |
|
923 | 71 |
primrec "op @" list |
1898 | 72 |
"[] @ ys = ys" |
73 |
"(x#xs)@ys = x#(xs@ys)" |
|
1169 | 74 |
primrec rev list |
1898 | 75 |
"rev([]) = []" |
76 |
"rev(x#xs) = rev(xs) @ [x]" |
|
923 | 77 |
primrec filter list |
1898 | 78 |
"filter P [] = []" |
79 |
"filter P (x#xs) = (if P x then x#filter P xs else filter P xs)" |
|
923 | 80 |
primrec foldl list |
1898 | 81 |
"foldl f a [] = a" |
82 |
"foldl f a (x#xs) = foldl f (f a x) xs" |
|
923 | 83 |
primrec length list |
1898 | 84 |
"length([]) = 0" |
85 |
"length(x#xs) = Suc(length(xs))" |
|
923 | 86 |
primrec flat list |
1898 | 87 |
"flat([]) = []" |
88 |
"flat(x#xs) = x @ flat(xs)" |
|
1419
a6a034a47a71
defined take/drop by induction over list rather than nat.
nipkow
parents:
1370
diff
changeset
|
89 |
primrec drop list |
a6a034a47a71
defined take/drop by induction over list rather than nat.
nipkow
parents:
1370
diff
changeset
|
90 |
drop_Nil "drop n [] = []" |
a6a034a47a71
defined take/drop by induction over list rather than nat.
nipkow
parents:
1370
diff
changeset
|
91 |
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
|
92 |
primrec take list |
a6a034a47a71
defined take/drop by induction over list rather than nat.
nipkow
parents:
1370
diff
changeset
|
93 |
take_Nil "take n [] = []" |
a6a034a47a71
defined take/drop by induction over list rather than nat.
nipkow
parents:
1370
diff
changeset
|
94 |
take_Cons "take n (x#xs) = (case n of 0 => [] | Suc(m) => x # take m xs)" |
923 | 95 |
defs |
1824 | 96 |
nth_def "nth(n) == nat_rec hd (%m r xs. r(tl(xs))) n" |
923 | 97 |
end |
1898 | 98 |