| author | paulson | 
| Fri, 06 Dec 1996 10:47:10 +0100 | |
| changeset 2330 | 3eea6b72bb4f | 
| parent 2262 | c7ee913746fd | 
| child 2369 | 8100f00e8950 | 
| 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: 
965diff
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 *) | |
| 2262 | 36 |   "@list"     :: args => 'a list                          ("[(_)]")
 | 
| 923 | 37 | |
| 38 | (* Special syntax for list_all and filter *) | |
| 2262 | 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#[]" | |
| 1475 | 45 | "[x:xs . P]" == "filter (%x.P) xs" | 
| 46 | "Alls x:xs.P" == "list_all (%x.P) xs" | |
| 923 | 47 | |
| 2262 | 48 | syntax (symbols) | 
| 49 | "op #" :: ['a, 'a list] => 'a list (infixr "\\<bullet>" 65) | |
| 50 | "op @" :: ['a list, 'a list] => 'a list (infixr "\\<circ>" 65) | |
| 51 | "op mem" :: ['a, 'a list] => bool (infixl "\\<in>" 55) | |
| 52 |   "@Alls"     :: [idt, 'a list, bool] => bool             ("(2\\<forall> _\\<in>_./ _)" 10)
 | |
| 53 |   "@filter"   :: [idt, 'a list, bool] => 'a list          ("(1[_\\<in>_ ./ _])")
 | |
| 54 | ||
| 55 | ||
| 923 | 56 | primrec hd list | 
| 1898 | 57 | "hd([]) = (@x.False)" | 
| 58 | "hd(x#xs) = x" | |
| 923 | 59 | primrec tl list | 
| 1898 | 60 | "tl([]) = (@x.False)" | 
| 61 | "tl(x#xs) = xs" | |
| 923 | 62 | primrec ttl list | 
| 63 | (* a "total" version of tl: *) | |
| 1898 | 64 | "ttl([]) = []" | 
| 65 | "ttl(x#xs) = xs" | |
| 923 | 66 | primrec "op mem" list | 
| 1898 | 67 | "x mem [] = False" | 
| 68 | "x mem (y#ys) = (if y=x then True else x mem ys)" | |
| 1908 | 69 | primrec set_of_list list | 
| 70 |   "set_of_list [] = {}"
 | |
| 71 | "set_of_list (x#xs) = insert x (set_of_list xs)" | |
| 923 | 72 | primrec list_all list | 
| 73 | list_all_Nil "list_all P [] = True" | |
| 74 | list_all_Cons "list_all P (x#xs) = (P(x) & list_all P xs)" | |
| 75 | primrec map list | |
| 1898 | 76 | "map f [] = []" | 
| 77 | "map f (x#xs) = f(x)#map f xs" | |
| 923 | 78 | primrec "op @" list | 
| 1898 | 79 | "[] @ ys = ys" | 
| 80 | "(x#xs)@ys = x#(xs@ys)" | |
| 1169 | 81 | primrec rev list | 
| 1898 | 82 | "rev([]) = []" | 
| 83 | "rev(x#xs) = rev(xs) @ [x]" | |
| 923 | 84 | primrec filter list | 
| 1898 | 85 | "filter P [] = []" | 
| 86 | "filter P (x#xs) = (if P x then x#filter P xs else filter P xs)" | |
| 923 | 87 | primrec foldl list | 
| 1898 | 88 | "foldl f a [] = a" | 
| 89 | "foldl f a (x#xs) = foldl f (f a x) xs" | |
| 923 | 90 | primrec length list | 
| 1898 | 91 | "length([]) = 0" | 
| 92 | "length(x#xs) = Suc(length(xs))" | |
| 923 | 93 | primrec flat list | 
| 1898 | 94 | "flat([]) = []" | 
| 95 | "flat(x#xs) = x @ flat(xs)" | |
| 1419 
a6a034a47a71
defined take/drop by induction over list rather than nat.
 nipkow parents: 
1370diff
changeset | 96 | primrec drop list | 
| 
a6a034a47a71
defined take/drop by induction over list rather than nat.
 nipkow parents: 
1370diff
changeset | 97 | drop_Nil "drop n [] = []" | 
| 
a6a034a47a71
defined take/drop by induction over list rather than nat.
 nipkow parents: 
1370diff
changeset | 98 | 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: 
1370diff
changeset | 99 | primrec take list | 
| 
a6a034a47a71
defined take/drop by induction over list rather than nat.
 nipkow parents: 
1370diff
changeset | 100 | take_Nil "take n [] = []" | 
| 
a6a034a47a71
defined take/drop by induction over list rather than nat.
 nipkow parents: 
1370diff
changeset | 101 | take_Cons "take n (x#xs) = (case n of 0 => [] | Suc(m) => x # take m xs)" | 
| 923 | 102 | defs | 
| 1824 | 103 | nth_def "nth(n) == nat_rec hd (%m r xs. r(tl(xs))) n" | 
| 923 | 104 | end |