author | wenzelm |
Thu, 27 Jul 1995 13:13:32 +0200 | |
changeset 1201 | de2fc8cf9b6a |
parent 1169 | 5873833cf37f |
child 1327 | 6c29cfab679c |
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 |
||
8 |
*) |
|
9 |
||
10 |
List = Arith + |
|
11 |
||
977
5d57287e5e1e
changed syntax of datatype declarations (curried types for constructor
clasohm
parents:
965
diff
changeset
|
12 |
datatype 'a list = "[]" ("[]") | "#" 'a ('a list) (infixr 65) |
923 | 13 |
|
14 |
consts |
|
15 |
||
16 |
null :: "'a list => bool" |
|
17 |
hd :: "'a list => 'a" |
|
18 |
tl,ttl :: "'a list => 'a list" |
|
19 |
mem :: "['a, 'a list] => bool" (infixl 55) |
|
20 |
list_all :: "('a => bool) => ('a list => bool)" |
|
21 |
map :: "('a=>'b) => ('a list => 'b list)" |
|
22 |
"@" :: "['a list, 'a list] => 'a list" (infixr 65) |
|
1169 | 23 |
rev :: "'a list => 'a list" |
923 | 24 |
filter :: "['a => bool, 'a list] => 'a list" |
25 |
foldl :: "[['b,'a] => 'b, 'b, 'a list] => 'b" |
|
26 |
length :: "'a list => nat" |
|
27 |
flat :: "'a list list => 'a list" |
|
28 |
nth :: "[nat, 'a list] => 'a" |
|
29 |
||
30 |
syntax |
|
31 |
(* list Enumeration *) |
|
32 |
"@list" :: "args => 'a list" ("[(_)]") |
|
33 |
||
34 |
(* Special syntax for list_all and filter *) |
|
35 |
"@Alls" :: "[idt, 'a list, bool] => bool" ("(2Alls _:_./ _)" 10) |
|
36 |
"@filter" :: "[idt, 'a list, bool] => 'a list" ("(1[_:_ ./ _])") |
|
37 |
||
38 |
translations |
|
39 |
"[x, xs]" == "x#[xs]" |
|
40 |
"[x]" == "x#[]" |
|
41 |
||
42 |
"[x:xs . P]" == "filter (%x.P) xs" |
|
43 |
"Alls x:xs.P" == "list_all (%x.P) xs" |
|
44 |
||
45 |
primrec null list |
|
46 |
null_Nil "null([]) = True" |
|
47 |
null_Cons "null(x#xs) = False" |
|
48 |
primrec hd list |
|
49 |
hd_Nil "hd([]) = (@x.False)" |
|
50 |
hd_Cons "hd(x#xs) = x" |
|
51 |
primrec tl list |
|
52 |
tl_Nil "tl([]) = (@x.False)" |
|
53 |
tl_Cons "tl(x#xs) = xs" |
|
54 |
primrec ttl list |
|
55 |
(* a "total" version of tl: *) |
|
56 |
ttl_Nil "ttl([]) = []" |
|
57 |
ttl_Cons "ttl(x#xs) = xs" |
|
58 |
primrec "op mem" list |
|
59 |
mem_Nil "x mem [] = False" |
|
965 | 60 |
mem_Cons "x mem (y#ys) = (if y=x then True else x mem ys)" |
923 | 61 |
primrec list_all list |
62 |
list_all_Nil "list_all P [] = True" |
|
63 |
list_all_Cons "list_all P (x#xs) = (P(x) & list_all P xs)" |
|
64 |
primrec map list |
|
65 |
map_Nil "map f [] = []" |
|
66 |
map_Cons "map f (x#xs) = f(x)#map f xs" |
|
67 |
primrec "op @" list |
|
68 |
append_Nil "[] @ ys = ys" |
|
69 |
append_Cons "(x#xs)@ys = x#(xs@ys)" |
|
1169 | 70 |
primrec rev list |
71 |
rev_Nil "rev([]) = []" |
|
72 |
rev_Cons "rev(x#xs) = rev(xs) @ [x]" |
|
923 | 73 |
primrec filter list |
74 |
filter_Nil "filter P [] = []" |
|
965 | 75 |
filter_Cons "filter P (x#xs) = (if P x then x#filter P xs else filter P xs)" |
923 | 76 |
primrec foldl list |
77 |
foldl_Nil "foldl f a [] = a" |
|
78 |
foldl_Cons "foldl f a (x#xs) = foldl f (f a x) xs" |
|
79 |
primrec length list |
|
80 |
length_Nil "length([]) = 0" |
|
81 |
length_Cons "length(x#xs) = Suc(length(xs))" |
|
82 |
primrec flat list |
|
83 |
flat_Nil "flat([]) = []" |
|
84 |
flat_Cons "flat(x#xs) = x @ flat(xs)" |
|
85 |
defs |
|
86 |
nth_def "nth(n) == nat_rec n hd (%m r xs. r(tl(xs)))" |
|
87 |
end |