| author | wenzelm | 
| Mon, 09 Nov 1998 15:33:12 +0100 | |
| changeset 5826 | 977f789566b7 | 
| parent 5518 | 654ead0ba4f7 | 
| child 6141 | a6922171b396 | 
| 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  | 
||
| 5427 | 9  | 
List = Datatype + WF_Rel +  | 
| 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  | 
| 
3896
 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 
nipkow 
parents: 
3842 
diff
changeset
 | 
18  | 
hd, last :: 'a list => 'a  | 
| 3465 | 19  | 
set :: 'a list => 'a set  | 
| 5518 | 20  | 
  list_all    :: ('a => bool) => ('a list => bool)
 | 
| 1908 | 21  | 
  map         :: ('a=>'b) => ('a list => 'b list)
 | 
| 5518 | 22  | 
mem :: ['a, 'a list] => bool (infixl 55)  | 
| 4502 | 23  | 
nth :: ['a list, nat] => 'a (infixl "!" 100)  | 
| 
5077
 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 
nipkow 
parents: 
4643 
diff
changeset
 | 
24  | 
list_update :: 'a list => nat => 'a => 'a list  | 
| 2608 | 25  | 
take, drop :: [nat, 'a list] => 'a list  | 
26  | 
takeWhile,  | 
|
27  | 
  dropWhile   :: ('a => bool) => 'a list => 'a list
 | 
|
| 
3896
 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 
nipkow 
parents: 
3842 
diff
changeset
 | 
28  | 
tl, butlast :: 'a list => 'a list  | 
| 1908 | 29  | 
rev :: 'a list => 'a list  | 
| 4132 | 30  | 
  zip	      :: "'a list => 'b list => ('a * 'b) list"
 | 
| 5427 | 31  | 
  upt         :: nat => nat => nat list                   ("(1[_../_'(])")
 | 
| 4605 | 32  | 
remdups :: 'a list => 'a list  | 
33  | 
nodups :: "'a list => bool"  | 
|
| 
3589
 
244daa75f890
Added function `replicate' and lemmas map_cong and set_replicate.
 
nipkow 
parents: 
3584 
diff
changeset
 | 
34  | 
replicate :: nat => 'a => 'a list  | 
| 923 | 35  | 
|
| 
5077
 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 
nipkow 
parents: 
4643 
diff
changeset
 | 
36  | 
nonterminals  | 
| 
 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 
nipkow 
parents: 
4643 
diff
changeset
 | 
37  | 
lupdbinds lupdbind  | 
| 
 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 
nipkow 
parents: 
4643 
diff
changeset
 | 
38  | 
|
| 923 | 39  | 
syntax  | 
40  | 
(* list Enumeration *)  | 
|
| 2262 | 41  | 
  "@list"     :: args => 'a list                          ("[(_)]")
 | 
| 923 | 42  | 
|
| 2512 | 43  | 
(* Special syntax for filter *)  | 
| 5295 | 44  | 
  "@filter"   :: [pttrn, 'a list, bool] => 'a list        ("(1[_:_ ./ _])")
 | 
| 923 | 45  | 
|
| 
5077
 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 
nipkow 
parents: 
4643 
diff
changeset
 | 
46  | 
(* list update *)  | 
| 
 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 
nipkow 
parents: 
4643 
diff
changeset
 | 
47  | 
  "_lupdbind"      :: ['a, 'a] => lupdbind            ("(2_ :=/ _)")
 | 
| 
 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 
nipkow 
parents: 
4643 
diff
changeset
 | 
48  | 
  ""               :: lupdbind => lupdbinds           ("_")
 | 
| 
 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 
nipkow 
parents: 
4643 
diff
changeset
 | 
49  | 
  "_lupdbinds"     :: [lupdbind, lupdbinds] => lupdbinds ("_,/ _")
 | 
| 
 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 
nipkow 
parents: 
4643 
diff
changeset
 | 
50  | 
  "_LUpdate"       :: ['a, lupdbinds] => 'a           ("_/[(_)]" [900,0] 900)
 | 
| 
 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 
nipkow 
parents: 
4643 
diff
changeset
 | 
51  | 
|
| 5427 | 52  | 
  upto        :: nat => nat => nat list                   ("(1[_../_])")
 | 
53  | 
||
| 923 | 54  | 
translations  | 
55  | 
"[x, xs]" == "x#[xs]"  | 
|
56  | 
"[x]" == "x#[]"  | 
|
| 3842 | 57  | 
"[x:xs . P]" == "filter (%x. P) xs"  | 
| 923 | 58  | 
|
| 
5077
 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 
nipkow 
parents: 
4643 
diff
changeset
 | 
59  | 
"_LUpdate xs (_lupdbinds b bs)" == "_LUpdate (_LUpdate xs b) bs"  | 
| 
 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 
nipkow 
parents: 
4643 
diff
changeset
 | 
60  | 
"xs[i:=x]" == "list_update xs i x"  | 
| 
 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 
nipkow 
parents: 
4643 
diff
changeset
 | 
61  | 
|
| 5427 | 62  | 
"[i..j]" == "[i..(Suc j)(]"  | 
63  | 
||
64  | 
||
| 2262 | 65  | 
syntax (symbols)  | 
| 5295 | 66  | 
  "@filter"   :: [pttrn, 'a list, bool] => 'a list        ("(1[_\\<in>_ ./ _])")
 | 
| 2262 | 67  | 
|
68  | 
||
| 
3342
 
ec3b55fcb165
New operator "lists" for formalizing sets of lists
 
paulson 
parents: 
3320 
diff
changeset
 | 
69  | 
consts  | 
| 
 
ec3b55fcb165
New operator "lists" for formalizing sets of lists
 
paulson 
parents: 
3320 
diff
changeset
 | 
70  | 
lists :: 'a set => 'a list set  | 
| 
 
ec3b55fcb165
New operator "lists" for formalizing sets of lists
 
paulson 
parents: 
3320 
diff
changeset
 | 
71  | 
|
| 
 
ec3b55fcb165
New operator "lists" for formalizing sets of lists
 
paulson 
parents: 
3320 
diff
changeset
 | 
72  | 
inductive "lists A"  | 
| 
 
ec3b55fcb165
New operator "lists" for formalizing sets of lists
 
paulson 
parents: 
3320 
diff
changeset
 | 
73  | 
intrs  | 
| 
 
ec3b55fcb165
New operator "lists" for formalizing sets of lists
 
paulson 
parents: 
3320 
diff
changeset
 | 
74  | 
Nil "[]: lists A"  | 
| 
 
ec3b55fcb165
New operator "lists" for formalizing sets of lists
 
paulson 
parents: 
3320 
diff
changeset
 | 
75  | 
Cons "[| a: A; l: lists A |] ==> a#l : lists A"  | 
| 
 
ec3b55fcb165
New operator "lists" for formalizing sets of lists
 
paulson 
parents: 
3320 
diff
changeset
 | 
76  | 
|
| 
 
ec3b55fcb165
New operator "lists" for formalizing sets of lists
 
paulson 
parents: 
3320 
diff
changeset
 | 
77  | 
|
| 
3437
 
bea2faf1641d
Replacing the primrec definition of "length" by a translation to the built-in
 
paulson 
parents: 
3401 
diff
changeset
 | 
78  | 
(*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
 | 
79  | 
list version as "length".*)  | 
| 
 
bea2faf1641d
Replacing the primrec definition of "length" by a translation to the built-in
 
paulson 
parents: 
3401 
diff
changeset
 | 
80  | 
syntax length :: 'a list => nat  | 
| 3507 | 81  | 
translations "length" => "size:: _ list => nat"  | 
| 
3437
 
bea2faf1641d
Replacing the primrec definition of "length" by a translation to the built-in
 
paulson 
parents: 
3401 
diff
changeset
 | 
82  | 
|
| 5183 | 83  | 
primrec  | 
| 3320 | 84  | 
"hd([]) = arbitrary"  | 
| 1898 | 85  | 
"hd(x#xs) = x"  | 
| 5183 | 86  | 
primrec  | 
| 
3896
 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 
nipkow 
parents: 
3842 
diff
changeset
 | 
87  | 
"tl([]) = []"  | 
| 1898 | 88  | 
"tl(x#xs) = xs"  | 
| 5183 | 89  | 
primrec  | 
| 
3896
 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 
nipkow 
parents: 
3842 
diff
changeset
 | 
90  | 
"last [] = arbitrary"  | 
| 
 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 
nipkow 
parents: 
3842 
diff
changeset
 | 
91  | 
"last(x#xs) = (if xs=[] then x else last xs)"  | 
| 5183 | 92  | 
primrec  | 
| 
3896
 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 
nipkow 
parents: 
3842 
diff
changeset
 | 
93  | 
"butlast [] = []"  | 
| 
 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 
nipkow 
parents: 
3842 
diff
changeset
 | 
94  | 
"butlast(x#xs) = (if xs=[] then [] else x#butlast xs)"  | 
| 5183 | 95  | 
primrec  | 
| 5518 | 96  | 
"x mem [] = False"  | 
97  | 
"x mem (y#ys) = (if y=x then True else x mem ys)"  | 
|
98  | 
primrec  | 
|
| 3465 | 99  | 
  "set [] = {}"
 | 
100  | 
"set (x#xs) = insert x (set xs)"  | 
|
| 5183 | 101  | 
primrec  | 
| 5518 | 102  | 
list_all_Nil "list_all P [] = True"  | 
103  | 
list_all_Cons "list_all P (x#xs) = (P(x) & list_all P xs)"  | 
|
104  | 
primrec  | 
|
| 1898 | 105  | 
"map f [] = []"  | 
106  | 
"map f (x#xs) = f(x)#map f xs"  | 
|
| 5183 | 107  | 
primrec  | 
108  | 
append_Nil "[] @ys = ys"  | 
|
109  | 
append_Cons "(x#xs)@ys = x#(xs@ys)"  | 
|
110  | 
primrec  | 
|
| 1898 | 111  | 
"rev([]) = []"  | 
112  | 
"rev(x#xs) = rev(xs) @ [x]"  | 
|
| 5183 | 113  | 
primrec  | 
| 1898 | 114  | 
"filter P [] = []"  | 
115  | 
"filter P (x#xs) = (if P x then x#filter P xs else filter P xs)"  | 
|
| 5183 | 116  | 
primrec  | 
| 1898 | 117  | 
"foldl f a [] = a"  | 
118  | 
"foldl f a (x#xs) = foldl f (f a x) xs"  | 
|
| 5183 | 119  | 
primrec  | 
| 2608 | 120  | 
"concat([]) = []"  | 
121  | 
"concat(x#xs) = x @ concat(xs)"  | 
|
| 5183 | 122  | 
primrec  | 
| 
1419
 
a6a034a47a71
defined take/drop by induction over list rather than nat.
 
nipkow 
parents: 
1370 
diff
changeset
 | 
123  | 
drop_Nil "drop n [] = []"  | 
| 
 
a6a034a47a71
defined take/drop by induction over list rather than nat.
 
nipkow 
parents: 
1370 
diff
changeset
 | 
124  | 
drop_Cons "drop n (x#xs) = (case n of 0 => x#xs | Suc(m) => drop m xs)"  | 
| 5183 | 125  | 
primrec  | 
| 
1419
 
a6a034a47a71
defined take/drop by induction over list rather than nat.
 
nipkow 
parents: 
1370 
diff
changeset
 | 
126  | 
take_Nil "take n [] = []"  | 
| 
 
a6a034a47a71
defined take/drop by induction over list rather than nat.
 
nipkow 
parents: 
1370 
diff
changeset
 | 
127  | 
take_Cons "take n (x#xs) = (case n of 0 => [] | Suc(m) => x # take m xs)"  | 
| 5183 | 128  | 
primrec  | 
| 4502 | 129  | 
"xs!0 = hd xs"  | 
130  | 
"xs!(Suc n) = (tl xs)!n"  | 
|
| 5183 | 131  | 
primrec  | 
| 
5077
 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 
nipkow 
parents: 
4643 
diff
changeset
 | 
132  | 
" [][i:=v] = []"  | 
| 
 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 
nipkow 
parents: 
4643 
diff
changeset
 | 
133  | 
"(x#xs)[i:=v] = (case i of 0 => v # xs  | 
| 
 
71043526295f
* HOL/List: new function list_update written xs[i:=v] that updates the i-th
 
nipkow 
parents: 
4643 
diff
changeset
 | 
134  | 
| Suc j => x # xs[j:=v])"  | 
| 5183 | 135  | 
primrec  | 
| 2608 | 136  | 
"takeWhile P [] = []"  | 
137  | 
"takeWhile P (x#xs) = (if P x then x#takeWhile P xs else [])"  | 
|
| 5183 | 138  | 
primrec  | 
| 2608 | 139  | 
"dropWhile P [] = []"  | 
| 
3584
 
8f9ee0f79d9a
Corected bug in def of dropWhile (also present in Haskell lib!)
 
nipkow 
parents: 
3507 
diff
changeset
 | 
140  | 
"dropWhile P (x#xs) = (if P x then dropWhile P xs else x#xs)"  | 
| 5183 | 141  | 
primrec  | 
| 4132 | 142  | 
"zip xs [] = []"  | 
143  | 
"zip xs (y#ys) = (hd xs,y)#zip (tl xs) ys"  | 
|
| 5427 | 144  | 
primrec  | 
145  | 
"[i..0(] = []"  | 
|
146  | 
"[i..(Suc j)(] = (if i <= j then [i..j(] @ [j] else [])"  | 
|
| 5183 | 147  | 
primrec  | 
| 4605 | 148  | 
"nodups [] = True"  | 
149  | 
"nodups (x#xs) = (x ~: set xs & nodups xs)"  | 
|
| 5183 | 150  | 
primrec  | 
| 4605 | 151  | 
"remdups [] = []"  | 
152  | 
"remdups (x#xs) = (if x : set xs then remdups xs else x # remdups xs)"  | 
|
| 5183 | 153  | 
primrec  | 
| 
5443
 
e2459d18ff47
changed constants mem and list_all to mere translations
 
oheimb 
parents: 
5427 
diff
changeset
 | 
154  | 
replicate_0 "replicate 0 x = []"  | 
| 5183 | 155  | 
replicate_Suc "replicate (Suc n) x = x # replicate n x"  | 
| 3196 | 156  | 
|
| 5281 | 157  | 
(** Lexcicographic orderings on lists **)  | 
158  | 
||
159  | 
consts  | 
|
160  | 
 lexn :: "('a * 'a)set => nat => ('a list * 'a list)set"
 | 
|
161  | 
primrec  | 
|
162  | 
"lexn r 0       = {}"
 | 
|
163  | 
"lexn r (Suc n) = (prod_fun (split op#) (split op#) `` (r ** lexn r n)) Int  | 
|
164  | 
                  {(xs,ys). length xs = Suc n & length ys = Suc n}"
 | 
|
165  | 
||
166  | 
constdefs  | 
|
167  | 
 lex :: "('a * 'a)set => ('a list * 'a list)set"
 | 
|
168  | 
"lex r == UN n. lexn r n"  | 
|
169  | 
||
170  | 
 lexico :: "('a * 'a)set => ('a list * 'a list)set"
 | 
|
171  | 
"lexico r == inv_image (less_than ** lex r) (%xs. (length xs, xs))"  | 
|
172  | 
||
| 923 | 173  | 
end  | 
| 3507 | 174  | 
|
175  | 
ML  | 
|
176  | 
||
177  | 
local  | 
|
178  | 
||
179  | 
(* translating size::list -> length *)  | 
|
180  | 
||
| 4151 | 181  | 
fun size_tr' _ (Type ("fun", (Type ("list", _) :: _))) [t] =
 | 
| 3507 | 182  | 
Syntax.const "length" $ t  | 
| 4151 | 183  | 
| size_tr' _ _ _ = raise Match;  | 
| 3507 | 184  | 
|
185  | 
in  | 
|
186  | 
||
187  | 
val typed_print_translation = [("size", size_tr')];
 | 
|
188  | 
||
189  | 
end;  |