| author | paulson | 
| Fri, 18 Apr 1997 11:47:11 +0200 | |
| changeset 2981 | aa5aeb6467c6 | 
| parent 2738 | e28a0668dbfe | 
| child 3196 | c522bc46aea7 | 
| 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  | 
||
9  | 
List = Arith +  | 
|
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  | 
|
14  | 
||
| 1908 | 15  | 
"@" :: ['a list, 'a list] => 'a list (infixr 65)  | 
16  | 
filter :: ['a => bool, 'a list] => 'a list  | 
|
| 2608 | 17  | 
concat :: 'a list list => 'a list  | 
| 1908 | 18  | 
foldl :: [['b,'a] => 'b, 'b, 'a list] => 'b  | 
19  | 
hd :: 'a list => 'a  | 
|
20  | 
length :: 'a list => nat  | 
|
| 2608 | 21  | 
set_of_list :: 'a list => 'a set  | 
| 1908 | 22  | 
  list_all    :: ('a => bool) => ('a list => bool)
 | 
23  | 
  map         :: ('a=>'b) => ('a list => 'b list)
 | 
|
24  | 
mem :: ['a, 'a list] => bool (infixl 55)  | 
|
25  | 
nth :: [nat, 'a list] => 'a  | 
|
| 2608 | 26  | 
take, drop :: [nat, 'a list] => 'a list  | 
27  | 
takeWhile,  | 
|
28  | 
  dropWhile   :: ('a => bool) => 'a list => 'a list
 | 
|
| 1908 | 29  | 
tl,ttl :: 'a list => 'a list  | 
30  | 
rev :: 'a list => 'a list  | 
|
| 923 | 31  | 
|
32  | 
syntax  | 
|
33  | 
(* list Enumeration *)  | 
|
| 2262 | 34  | 
  "@list"     :: args => 'a list                          ("[(_)]")
 | 
| 923 | 35  | 
|
| 2512 | 36  | 
(* Special syntax for filter *)  | 
| 2262 | 37  | 
  "@filter"   :: [idt, 'a list, bool] => 'a list          ("(1[_:_ ./ _])")
 | 
| 923 | 38  | 
|
39  | 
translations  | 
|
40  | 
"[x, xs]" == "x#[xs]"  | 
|
41  | 
"[x]" == "x#[]"  | 
|
| 1475 | 42  | 
"[x:xs . P]" == "filter (%x.P) xs"  | 
| 923 | 43  | 
|
| 2262 | 44  | 
syntax (symbols)  | 
45  | 
  "@filter"   :: [idt, 'a list, bool] => 'a list          ("(1[_\\<in>_ ./ _])")
 | 
|
46  | 
||
47  | 
||
| 923 | 48  | 
primrec hd list  | 
| 1898 | 49  | 
"hd([]) = (@x.False)"  | 
50  | 
"hd(x#xs) = x"  | 
|
| 923 | 51  | 
primrec tl list  | 
| 1898 | 52  | 
"tl([]) = (@x.False)"  | 
53  | 
"tl(x#xs) = xs"  | 
|
| 923 | 54  | 
primrec ttl list  | 
55  | 
(* a "total" version of tl: *)  | 
|
| 1898 | 56  | 
"ttl([]) = []"  | 
57  | 
"ttl(x#xs) = xs"  | 
|
| 923 | 58  | 
primrec "op mem" list  | 
| 1898 | 59  | 
"x mem [] = False"  | 
60  | 
"x mem (y#ys) = (if y=x then True else x mem ys)"  | 
|
| 1908 | 61  | 
primrec set_of_list list  | 
62  | 
  "set_of_list [] = {}"
 | 
|
63  | 
"set_of_list (x#xs) = insert x (set_of_list xs)"  | 
|
| 923 | 64  | 
primrec list_all list  | 
65  | 
list_all_Nil "list_all P [] = True"  | 
|
66  | 
list_all_Cons "list_all P (x#xs) = (P(x) & list_all P xs)"  | 
|
67  | 
primrec map list  | 
|
| 1898 | 68  | 
"map f [] = []"  | 
69  | 
"map f (x#xs) = f(x)#map f xs"  | 
|
| 923 | 70  | 
primrec "op @" list  | 
| 1898 | 71  | 
"[] @ ys = ys"  | 
72  | 
"(x#xs)@ys = x#(xs@ys)"  | 
|
| 1169 | 73  | 
primrec rev list  | 
| 1898 | 74  | 
"rev([]) = []"  | 
75  | 
"rev(x#xs) = rev(xs) @ [x]"  | 
|
| 923 | 76  | 
primrec filter list  | 
| 1898 | 77  | 
"filter P [] = []"  | 
78  | 
"filter P (x#xs) = (if P x then x#filter P xs else filter P xs)"  | 
|
| 923 | 79  | 
primrec foldl list  | 
| 1898 | 80  | 
"foldl f a [] = a"  | 
81  | 
"foldl f a (x#xs) = foldl f (f a x) xs"  | 
|
| 923 | 82  | 
primrec length list  | 
| 1898 | 83  | 
"length([]) = 0"  | 
84  | 
"length(x#xs) = Suc(length(xs))"  | 
|
| 2608 | 85  | 
primrec concat list  | 
86  | 
"concat([]) = []"  | 
|
87  | 
"concat(x#xs) = x @ concat(xs)"  | 
|
| 
1419
 
a6a034a47a71
defined take/drop by induction over list rather than nat.
 
nipkow 
parents: 
1370 
diff
changeset
 | 
88  | 
primrec drop list  | 
| 
 
a6a034a47a71
defined take/drop by induction over list rather than nat.
 
nipkow 
parents: 
1370 
diff
changeset
 | 
89  | 
drop_Nil "drop n [] = []"  | 
| 
 
a6a034a47a71
defined take/drop by induction over list rather than nat.
 
nipkow 
parents: 
1370 
diff
changeset
 | 
90  | 
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
 | 
91  | 
primrec take list  | 
| 
 
a6a034a47a71
defined take/drop by induction over list rather than nat.
 
nipkow 
parents: 
1370 
diff
changeset
 | 
92  | 
take_Nil "take n [] = []"  | 
| 
 
a6a034a47a71
defined take/drop by induction over list rather than nat.
 
nipkow 
parents: 
1370 
diff
changeset
 | 
93  | 
take_Cons "take n (x#xs) = (case n of 0 => [] | Suc(m) => x # take m xs)"  | 
| 2738 | 94  | 
primrec nth nat  | 
95  | 
"nth 0 xs = hd xs"  | 
|
96  | 
"nth (Suc n) xs = nth n (tl xs)"  | 
|
| 2608 | 97  | 
primrec takeWhile list  | 
98  | 
"takeWhile P [] = []"  | 
|
99  | 
"takeWhile P (x#xs) = (if P x then x#takeWhile P xs else [])"  | 
|
100  | 
primrec dropWhile list  | 
|
101  | 
"dropWhile P [] = []"  | 
|
102  | 
"dropWhile P (x#xs) = (if P x then dropWhile P xs else xs)"  | 
|
| 923 | 103  | 
end  |