src/HOL/List.thy
author nipkow
Wed, 09 Jul 1997 12:57:04 +0200
changeset 3507 157be29ad5ba
parent 3465 e85c24717cad
child 3584 8f9ee0f79d9a
permissions -rw-r--r--
Improved length = size translation.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     1
(*  Title:      HOL/List.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     3
    Author:     Tobias Nipkow
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     4
    Copyright   1994 TU Muenchen
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     5
2512
0231e4f467f2 Got rid of Alls in List.
nipkow
parents: 2369
diff changeset
     6
The datatype of finite lists.
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     7
*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     8
3367
832c245d967c Now Divides must be the parent
paulson
parents: 3342
diff changeset
     9
List = Divides +
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    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
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    12
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    13
consts
1908
55d8e38262a8 Renamed setOfList to set_of_list
paulson
parents: 1898
diff changeset
    14
  "@"         :: ['a list, 'a list] => 'a list            (infixr 65)
55d8e38262a8 Renamed setOfList to set_of_list
paulson
parents: 1898
diff changeset
    15
  filter      :: ['a => bool, 'a list] => 'a list
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents: 2512
diff changeset
    16
  concat      :: 'a list list => 'a list
1908
55d8e38262a8 Renamed setOfList to set_of_list
paulson
parents: 1898
diff changeset
    17
  foldl       :: [['b,'a] => 'b, 'b, 'a list] => 'b
55d8e38262a8 Renamed setOfList to set_of_list
paulson
parents: 1898
diff changeset
    18
  hd          :: 'a list => 'a
3465
e85c24717cad set_of_list -> set
nipkow
parents: 3437
diff changeset
    19
  set         :: 'a list => 'a set
1908
55d8e38262a8 Renamed setOfList to set_of_list
paulson
parents: 1898
diff changeset
    20
  list_all    :: ('a => bool) => ('a list => bool)
55d8e38262a8 Renamed setOfList to set_of_list
paulson
parents: 1898
diff changeset
    21
  map         :: ('a=>'b) => ('a list => 'b list)
55d8e38262a8 Renamed setOfList to set_of_list
paulson
parents: 1898
diff changeset
    22
  mem         :: ['a, 'a list] => bool                    (infixl 55)
55d8e38262a8 Renamed setOfList to set_of_list
paulson
parents: 1898
diff changeset
    23
  nth         :: [nat, 'a list] => 'a
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents: 2512
diff changeset
    24
  take, drop  :: [nat, 'a list] => 'a list
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents: 2512
diff changeset
    25
  takeWhile,
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents: 2512
diff changeset
    26
  dropWhile   :: ('a => bool) => 'a list => 'a list
1908
55d8e38262a8 Renamed setOfList to set_of_list
paulson
parents: 1898
diff changeset
    27
  tl,ttl      :: 'a list => 'a list
55d8e38262a8 Renamed setOfList to set_of_list
paulson
parents: 1898
diff changeset
    28
  rev         :: 'a list => 'a list
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    29
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    30
syntax
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    31
  (* list Enumeration *)
2262
c7ee913746fd added symbols syntax;
wenzelm
parents: 1908
diff changeset
    32
  "@list"     :: args => 'a list                          ("[(_)]")
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    33
2512
0231e4f467f2 Got rid of Alls in List.
nipkow
parents: 2369
diff changeset
    34
  (* Special syntax for filter *)
2262
c7ee913746fd added symbols syntax;
wenzelm
parents: 1908
diff changeset
    35
  "@filter"   :: [idt, 'a list, bool] => 'a list          ("(1[_:_ ./ _])")
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    36
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    37
translations
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    38
  "[x, xs]"     == "x#[xs]"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    39
  "[x]"         == "x#[]"
1475
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1419
diff changeset
    40
  "[x:xs . P]"  == "filter (%x.P) xs"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    41
2262
c7ee913746fd added symbols syntax;
wenzelm
parents: 1908
diff changeset
    42
syntax (symbols)
c7ee913746fd added symbols syntax;
wenzelm
parents: 1908
diff changeset
    43
  "@filter"   :: [idt, 'a list, bool] => 'a list          ("(1[_\\<in>_ ./ _])")
c7ee913746fd added symbols syntax;
wenzelm
parents: 1908
diff changeset
    44
c7ee913746fd added symbols syntax;
wenzelm
parents: 1908
diff changeset
    45
3342
ec3b55fcb165 New operator "lists" for formalizing sets of lists
paulson
parents: 3320
diff changeset
    46
consts
ec3b55fcb165 New operator "lists" for formalizing sets of lists
paulson
parents: 3320
diff changeset
    47
  lists        :: 'a set => 'a list set
ec3b55fcb165 New operator "lists" for formalizing sets of lists
paulson
parents: 3320
diff changeset
    48
ec3b55fcb165 New operator "lists" for formalizing sets of lists
paulson
parents: 3320
diff changeset
    49
  inductive "lists A"
ec3b55fcb165 New operator "lists" for formalizing sets of lists
paulson
parents: 3320
diff changeset
    50
  intrs
ec3b55fcb165 New operator "lists" for formalizing sets of lists
paulson
parents: 3320
diff changeset
    51
    Nil  "[]: lists A"
ec3b55fcb165 New operator "lists" for formalizing sets of lists
paulson
parents: 3320
diff changeset
    52
    Cons "[| a: A;  l: lists A |] ==> a#l : lists A"
ec3b55fcb165 New operator "lists" for formalizing sets of lists
paulson
parents: 3320
diff changeset
    53
ec3b55fcb165 New operator "lists" for formalizing sets of lists
paulson
parents: 3320
diff changeset
    54
3437
bea2faf1641d Replacing the primrec definition of "length" by a translation to the built-in
paulson
parents: 3401
diff changeset
    55
(*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
    56
  list version as "length".*)
bea2faf1641d Replacing the primrec definition of "length" by a translation to the built-in
paulson
parents: 3401
diff changeset
    57
syntax   length :: 'a list => nat
3507
157be29ad5ba Improved length = size translation.
nipkow
parents: 3465
diff changeset
    58
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
    59
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    60
primrec hd list
3320
3a5e4930fb77 Added `arbitrary'
nipkow
parents: 3196
diff changeset
    61
  "hd([]) = arbitrary"
1898
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    62
  "hd(x#xs) = x"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    63
primrec tl list
3320
3a5e4930fb77 Added `arbitrary'
nipkow
parents: 3196
diff changeset
    64
  "tl([]) = arbitrary"
1898
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    65
  "tl(x#xs) = xs"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    66
primrec ttl list
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    67
  (* a "total" version of tl: *)
1898
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    68
  "ttl([]) = []"
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    69
  "ttl(x#xs) = xs"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    70
primrec "op mem" list
1898
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    71
  "x mem [] = False"
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    72
  "x mem (y#ys) = (if y=x then True else x mem ys)"
3465
e85c24717cad set_of_list -> set
nipkow
parents: 3437
diff changeset
    73
primrec set list
e85c24717cad set_of_list -> set
nipkow
parents: 3437
diff changeset
    74
  "set [] = {}"
e85c24717cad set_of_list -> set
nipkow
parents: 3437
diff changeset
    75
  "set (x#xs) = insert x (set xs)"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    76
primrec list_all list
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    77
  list_all_Nil  "list_all P [] = True"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    78
  list_all_Cons "list_all P (x#xs) = (P(x) & list_all P xs)"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    79
primrec map list
1898
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    80
  "map f [] = []"
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    81
  "map f (x#xs) = f(x)#map f xs"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    82
primrec "op @" list
1898
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    83
  "[] @ ys = ys"
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    84
  "(x#xs)@ys = x#(xs@ys)"
1169
5873833cf37f Added function rev and its properties length_rev, etc.
lcp
parents: 977
diff changeset
    85
primrec rev list
1898
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    86
  "rev([]) = []"
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    87
  "rev(x#xs) = rev(xs) @ [x]"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    88
primrec filter list
1898
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    89
  "filter P [] = []"
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    90
  "filter P (x#xs) = (if P x then x#filter P xs else filter P xs)"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    91
primrec foldl list
1898
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    92
  "foldl f a [] = a"
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    93
  "foldl f a (x#xs) = foldl f (f a x) xs"
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents: 2512
diff changeset
    94
primrec concat list
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents: 2512
diff changeset
    95
  "concat([]) = []"
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents: 2512
diff changeset
    96
  "concat(x#xs) = x @ concat(xs)"
1419
a6a034a47a71 defined take/drop by induction over list rather than nat.
nipkow
parents: 1370
diff changeset
    97
primrec drop list
a6a034a47a71 defined take/drop by induction over list rather than nat.
nipkow
parents: 1370
diff changeset
    98
  drop_Nil  "drop n [] = []"
a6a034a47a71 defined take/drop by induction over list rather than nat.
nipkow
parents: 1370
diff changeset
    99
  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
   100
primrec take list
a6a034a47a71 defined take/drop by induction over list rather than nat.
nipkow
parents: 1370
diff changeset
   101
  take_Nil  "take n [] = []"
a6a034a47a71 defined take/drop by induction over list rather than nat.
nipkow
parents: 1370
diff changeset
   102
  take_Cons "take n (x#xs) = (case n of 0 => [] | Suc(m) => x # take m xs)"
2738
e28a0668dbfe primrec definition for nth
pusch
parents: 2608
diff changeset
   103
primrec nth nat
e28a0668dbfe primrec definition for nth
pusch
parents: 2608
diff changeset
   104
  "nth 0 xs = hd xs"
e28a0668dbfe primrec definition for nth
pusch
parents: 2608
diff changeset
   105
  "nth (Suc n) xs = nth n (tl xs)"
2608
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents: 2512
diff changeset
   106
primrec takeWhile list
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents: 2512
diff changeset
   107
  "takeWhile P [] = []"
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents: 2512
diff changeset
   108
  "takeWhile P (x#xs) = (if P x then x#takeWhile P xs else [])"
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents: 2512
diff changeset
   109
primrec dropWhile list
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents: 2512
diff changeset
   110
  "dropWhile P [] = []"
450c9b682a92 New class "order" and accompanying changes.
nipkow
parents: 2512
diff changeset
   111
  "dropWhile P (x#xs) = (if P x then dropWhile P xs else xs)"
3196
c522bc46aea7 Added pred_list for TFL
paulson
parents: 2738
diff changeset
   112
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   113
end
3507
157be29ad5ba Improved length = size translation.
nipkow
parents: 3465
diff changeset
   114
157be29ad5ba Improved length = size translation.
nipkow
parents: 3465
diff changeset
   115
ML
157be29ad5ba Improved length = size translation.
nipkow
parents: 3465
diff changeset
   116
157be29ad5ba Improved length = size translation.
nipkow
parents: 3465
diff changeset
   117
local
157be29ad5ba Improved length = size translation.
nipkow
parents: 3465
diff changeset
   118
157be29ad5ba Improved length = size translation.
nipkow
parents: 3465
diff changeset
   119
(* translating size::list -> length *)
157be29ad5ba Improved length = size translation.
nipkow
parents: 3465
diff changeset
   120
157be29ad5ba Improved length = size translation.
nipkow
parents: 3465
diff changeset
   121
fun size_tr' (Type ("fun", (Type ("list", _) :: _))) [t] =
157be29ad5ba Improved length = size translation.
nipkow
parents: 3465
diff changeset
   122
      Syntax.const "length" $ t
157be29ad5ba Improved length = size translation.
nipkow
parents: 3465
diff changeset
   123
  | size_tr'   _ _ = raise Match;
157be29ad5ba Improved length = size translation.
nipkow
parents: 3465
diff changeset
   124
157be29ad5ba Improved length = size translation.
nipkow
parents: 3465
diff changeset
   125
in
157be29ad5ba Improved length = size translation.
nipkow
parents: 3465
diff changeset
   126
157be29ad5ba Improved length = size translation.
nipkow
parents: 3465
diff changeset
   127
val typed_print_translation = [("size", size_tr')];
157be29ad5ba Improved length = size translation.
nipkow
parents: 3465
diff changeset
   128
157be29ad5ba Improved length = size translation.
nipkow
parents: 3465
diff changeset
   129
end;