src/HOL/List.thy
author oheimb
Fri, 14 Jul 2000 16:28:49 +0200
changeset 9341 40bab6613000
parent 9336 9ae89b9ce206
child 9355 1b2cd40579c6
permissions -rw-r--r--
tuned syntax
clasohm@923
     1
(*  Title:      HOL/List.thy
clasohm@923
     2
    ID:         $Id$
clasohm@923
     3
    Author:     Tobias Nipkow
clasohm@923
     4
    Copyright   1994 TU Muenchen
clasohm@923
     5
nipkow@2512
     6
The datatype of finite lists.
clasohm@923
     7
*)
clasohm@923
     8
wenzelm@8490
     9
List = PreList +
clasohm@923
    10
wenzelm@7224
    11
datatype 'a list = Nil ("[]") | Cons 'a ('a list) (infixr "#" 65)
clasohm@923
    12
clasohm@923
    13
consts
paulson@1908
    14
  "@"         :: ['a list, 'a list] => 'a list            (infixr 65)
paulson@1908
    15
  filter      :: ['a => bool, 'a list] => 'a list
nipkow@2608
    16
  concat      :: 'a list list => 'a list
paulson@1908
    17
  foldl       :: [['b,'a] => 'b, 'b, 'a list] => 'b
paulson@8000
    18
  foldr       :: [['a,'b] => 'b, 'a list, 'b] => 'b
nipkow@3896
    19
  hd, last    :: 'a list => 'a
nipkow@3465
    20
  set         :: 'a list => 'a set
oheimb@5518
    21
  list_all    :: ('a => bool) => ('a list => bool)
nipkow@8115
    22
  list_all2   :: ('a => 'b => bool) => 'a list => 'b list => bool
paulson@1908
    23
  map         :: ('a=>'b) => ('a list => 'b list)
oheimb@5518
    24
  mem         :: ['a, 'a list] => bool                    (infixl 55)
nipkow@4502
    25
  nth         :: ['a list, nat] => 'a			  (infixl "!" 100)
nipkow@5077
    26
  list_update :: 'a list => nat => 'a => 'a list
nipkow@2608
    27
  take, drop  :: [nat, 'a list] => 'a list
nipkow@2608
    28
  takeWhile,
nipkow@2608
    29
  dropWhile   :: ('a => bool) => 'a list => 'a list
nipkow@3896
    30
  tl, butlast :: 'a list => 'a list
paulson@1908
    31
  rev         :: 'a list => 'a list
oheimb@4132
    32
  zip	      :: "'a list => 'b list => ('a * 'b) list"
nipkow@5427
    33
  upt         :: nat => nat => nat list                   ("(1[_../_'(])")
nipkow@4605
    34
  remdups     :: 'a list => 'a list
paulson@8972
    35
  null, nodups :: "'a list => bool"
nipkow@3589
    36
  replicate   :: nat => 'a => 'a list
clasohm@923
    37
nipkow@5077
    38
nonterminals
nipkow@5077
    39
  lupdbinds  lupdbind
nipkow@5077
    40
clasohm@923
    41
syntax
clasohm@923
    42
  (* list Enumeration *)
wenzelm@2262
    43
  "@list"     :: args => 'a list                          ("[(_)]")
clasohm@923
    44
nipkow@2512
    45
  (* Special syntax for filter *)
oheimb@9341
    46
  "@filter"   :: [pttrn, 'a list, bool] => 'a list        ("(1[_:_./ _])")
clasohm@923
    47
nipkow@5077
    48
  (* list update *)
nipkow@5077
    49
  "_lupdbind"      :: ['a, 'a] => lupdbind            ("(2_ :=/ _)")
nipkow@5077
    50
  ""               :: lupdbind => lupdbinds           ("_")
nipkow@5077
    51
  "_lupdbinds"     :: [lupdbind, lupdbinds] => lupdbinds ("_,/ _")
nipkow@5077
    52
  "_LUpdate"       :: ['a, lupdbinds] => 'a           ("_/[(_)]" [900,0] 900)
nipkow@5077
    53
nipkow@5427
    54
  upto        :: nat => nat => nat list                   ("(1[_../_])")
nipkow@5427
    55
clasohm@923
    56
translations
clasohm@923
    57
  "[x, xs]"     == "x#[xs]"
clasohm@923
    58
  "[x]"         == "x#[]"
wenzelm@3842
    59
  "[x:xs . P]"  == "filter (%x. P) xs"
clasohm@923
    60
nipkow@5077
    61
  "_LUpdate xs (_lupdbinds b bs)"  == "_LUpdate (_LUpdate xs b) bs"
nipkow@5077
    62
  "xs[i:=x]"                       == "list_update xs i x"
nipkow@5077
    63
nipkow@5427
    64
  "[i..j]" == "[i..(Suc j)(]"
nipkow@5427
    65
nipkow@5427
    66
wenzelm@2262
    67
syntax (symbols)
oheimb@5295
    68
  "@filter"   :: [pttrn, 'a list, bool] => 'a list        ("(1[_\\<in>_ ./ _])")
wenzelm@2262
    69
wenzelm@2262
    70
paulson@3342
    71
consts
paulson@3342
    72
  lists        :: 'a set => 'a list set
paulson@3342
    73
paulson@3342
    74
  inductive "lists A"
paulson@3342
    75
  intrs
paulson@3342
    76
    Nil  "[]: lists A"
paulson@3342
    77
    Cons "[| a: A;  l: lists A |] ==> a#l : lists A"
paulson@3342
    78
paulson@3342
    79
paulson@3437
    80
(*Function "size" is overloaded for all datatypes.  Users may refer to the
paulson@3437
    81
  list version as "length".*)
paulson@3437
    82
syntax   length :: 'a list => nat
nipkow@3507
    83
translations  "length"  =>  "size:: _ list => nat"
paulson@3437
    84
berghofe@5183
    85
primrec
berghofe@1898
    86
  "hd(x#xs) = x"
berghofe@5183
    87
primrec
paulson@8972
    88
  "tl([])   = []"
berghofe@1898
    89
  "tl(x#xs) = xs"
berghofe@5183
    90
primrec
paulson@8972
    91
  "null([])   = True"
paulson@8972
    92
  "null(x#xs) = False"
paulson@8972
    93
primrec
nipkow@3896
    94
  "last(x#xs) = (if xs=[] then x else last xs)"
berghofe@5183
    95
primrec
paulson@8972
    96
  "butlast []    = []"
nipkow@3896
    97
  "butlast(x#xs) = (if xs=[] then [] else x#butlast xs)"
berghofe@5183
    98
primrec
paulson@8972
    99
  "x mem []     = False"
oheimb@5518
   100
  "x mem (y#ys) = (if y=x then True else x mem ys)"
oheimb@5518
   101
primrec
nipkow@3465
   102
  "set [] = {}"
nipkow@3465
   103
  "set (x#xs) = insert x (set xs)"
berghofe@5183
   104
primrec
oheimb@5518
   105
  list_all_Nil  "list_all P [] = True"
oheimb@5518
   106
  list_all_Cons "list_all P (x#xs) = (P(x) & list_all P xs)"
oheimb@5518
   107
primrec
paulson@8972
   108
  "map f []     = []"
berghofe@1898
   109
  "map f (x#xs) = f(x)#map f xs"
berghofe@5183
   110
primrec
berghofe@5183
   111
  append_Nil  "[]    @ys = ys"
berghofe@5183
   112
  append_Cons "(x#xs)@ys = x#(xs@ys)"
berghofe@5183
   113
primrec
paulson@8972
   114
  "rev([])   = []"
berghofe@1898
   115
  "rev(x#xs) = rev(xs) @ [x]"
berghofe@5183
   116
primrec
paulson@8972
   117
  "filter P []     = []"
berghofe@1898
   118
  "filter P (x#xs) = (if P x then x#filter P xs else filter P xs)"
berghofe@5183
   119
primrec
paulson@6141
   120
  foldl_Nil  "foldl f a [] = a"
paulson@6141
   121
  foldl_Cons "foldl f a (x#xs) = foldl f (f a x) xs"
berghofe@5183
   122
primrec
paulson@8972
   123
  "foldr f [] a     = a"
paulson@8000
   124
  "foldr f (x#xs) a = f x (foldr f xs a)"
paulson@8000
   125
primrec
paulson@8972
   126
  "concat([])   = []"
nipkow@2608
   127
  "concat(x#xs) = x @ concat(xs)"
berghofe@5183
   128
primrec
nipkow@1419
   129
  drop_Nil  "drop n [] = []"
nipkow@1419
   130
  drop_Cons "drop n (x#xs) = (case n of 0 => x#xs | Suc(m) => drop m xs)"
pusch@6408
   131
  (* Warning: simpset does not contain this definition but separate theorems 
pusch@6408
   132
     for n=0 / n=Suc k*)
berghofe@5183
   133
primrec
nipkow@1419
   134
  take_Nil  "take n [] = []"
nipkow@1419
   135
  take_Cons "take n (x#xs) = (case n of 0 => [] | Suc(m) => x # take m xs)"
pusch@6408
   136
  (* Warning: simpset does not contain this definition but separate theorems 
pusch@6408
   137
     for n=0 / n=Suc k*)
pusch@6408
   138
primrec 
pusch@6408
   139
  nth_Cons  "(x#xs)!n = (case n of 0 => x | (Suc k) => xs!k)"
pusch@6408
   140
  (* Warning: simpset does not contain this definition but separate theorems 
pusch@6408
   141
     for n=0 / n=Suc k*)
berghofe@5183
   142
primrec
nipkow@5077
   143
 "    [][i:=v] = []"
nipkow@5077
   144
 "(x#xs)[i:=v] = (case i of 0     => v # xs 
nipkow@5077
   145
			  | Suc j => x # xs[j:=v])"
berghofe@5183
   146
primrec
paulson@8972
   147
  "takeWhile P []     = []"
nipkow@2608
   148
  "takeWhile P (x#xs) = (if P x then x#takeWhile P xs else [])"
berghofe@5183
   149
primrec
paulson@8972
   150
  "dropWhile P []     = []"
nipkow@3584
   151
  "dropWhile P (x#xs) = (if P x then dropWhile P xs else x#xs)"
berghofe@5183
   152
primrec
oheimb@4132
   153
  "zip xs []     = []"
nipkow@6306
   154
  "zip xs (y#ys) = (case xs of [] => [] | z#zs => (z,y)#zip zs ys)"
pusch@6408
   155
  (* Warning: simpset does not contain this definition but separate theorems 
pusch@6408
   156
     for xs=[] / xs=z#zs *)
nipkow@5427
   157
primrec
paulson@8983
   158
  upt_0   "[i..0(] = []"
paulson@8983
   159
  upt_Suc "[i..(Suc j)(] = (if i <= j then [i..j(] @ [j] else [])"
berghofe@5183
   160
primrec
nipkow@4605
   161
  "nodups []     = True"
nipkow@4605
   162
  "nodups (x#xs) = (x ~: set xs & nodups xs)"
berghofe@5183
   163
primrec
nipkow@4605
   164
  "remdups [] = []"
nipkow@4605
   165
  "remdups (x#xs) = (if x : set xs then remdups xs else x # remdups xs)"
berghofe@5183
   166
primrec
oheimb@5443
   167
  replicate_0   "replicate  0      x = []"
berghofe@5183
   168
  replicate_Suc "replicate (Suc n) x = x # replicate n x"
nipkow@8115
   169
defs
nipkow@8115
   170
 list_all2_def
nipkow@8115
   171
 "list_all2 P xs ys == length xs = length ys & (!(x,y):set(zip xs ys). P x y)"
nipkow@8115
   172
paulson@3196
   173
pusch@6408
   174
(** Lexicographic orderings on lists **)
nipkow@5281
   175
nipkow@5281
   176
consts
nipkow@5281
   177
 lexn :: "('a * 'a)set => nat => ('a list * 'a list)set"
nipkow@5281
   178
primrec
nipkow@5281
   179
"lexn r 0       = {}"
nipkow@8703
   180
"lexn r (Suc n) = (prod_fun (split op#) (split op#) `` (r <*lex*> lexn r n)) Int
nipkow@5281
   181
                  {(xs,ys). length xs = Suc n & length ys = Suc n}"
nipkow@5281
   182
nipkow@5281
   183
constdefs
paulson@9336
   184
  lex :: "('a * 'a)set => ('a list * 'a list)set"
paulson@9336
   185
    "lex r == UN n. lexn r n"
nipkow@5281
   186
paulson@9336
   187
  lexico :: "('a * 'a)set => ('a list * 'a list)set"
paulson@9336
   188
    "lexico r == inv_image (less_than <*lex*> lex r) (%xs. (length xs, xs))"
paulson@9336
   189
paulson@9336
   190
  sublist :: "['a list, nat set] => 'a list"
paulson@9336
   191
    "sublist xs A == map fst (filter (%p. snd p : A) (zip xs [0..size xs(]))"
nipkow@5281
   192
clasohm@923
   193
end
nipkow@3507
   194
nipkow@3507
   195
ML
nipkow@3507
   196
nipkow@3507
   197
local
nipkow@3507
   198
nipkow@3507
   199
(* translating size::list -> length *)
nipkow@3507
   200
wenzelm@4151
   201
fun size_tr' _ (Type ("fun", (Type ("list", _) :: _))) [t] =
nipkow@3507
   202
      Syntax.const "length" $ t
wenzelm@4151
   203
  | size_tr' _ _ _ = raise Match;
nipkow@3507
   204
nipkow@3507
   205
in
nipkow@3507
   206
nipkow@3507
   207
val typed_print_translation = [("size", size_tr')];
nipkow@3507
   208
nipkow@3507
   209
end;