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