src/HOL/List.thy
author berghofe
Fri, 02 Aug 1996 12:16:11 +0200
changeset 1898 260a9711f507
parent 1824 44254696843a
child 1908 55d8e38262a8
permissions -rw-r--r--
Simplified primrec definitions.
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
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     6
Definition of type 'a list as a datatype. This allows primrec to work.
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     7
1812
debfc40b7756 Addition of setOfList
paulson
parents: 1475
diff changeset
     8
TODO: delete list_all from this theory and from ex/Sorting, etc.
debfc40b7756 Addition of setOfList
paulson
parents: 1475
diff changeset
     9
      use setOfList instead
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    10
*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    11
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    12
List = Arith +
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    13
977
5d57287e5e1e changed syntax of datatype declarations (curried types for constructor
clasohm
parents: 965
diff changeset
    14
datatype 'a list = "[]" ("[]") | "#" 'a ('a list) (infixr 65)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    15
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    16
consts
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    17
1475
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1419
diff changeset
    18
  "@"       :: ['a list, 'a list] => 'a list            (infixr 65)
1370
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1327
diff changeset
    19
  drop      :: [nat, 'a list] => 'a list
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1327
diff changeset
    20
  filter    :: ['a => bool, 'a list] => 'a list
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1327
diff changeset
    21
  flat      :: 'a list list => 'a list
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1327
diff changeset
    22
  foldl     :: [['b,'a] => 'b, 'b, 'a list] => 'b
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1327
diff changeset
    23
  hd        :: 'a list => 'a
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1327
diff changeset
    24
  length    :: 'a list => nat
1812
debfc40b7756 Addition of setOfList
paulson
parents: 1475
diff changeset
    25
  setOfList :: ('a list => 'a set)
1370
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1327
diff changeset
    26
  list_all  :: ('a => bool) => ('a list => bool)
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1327
diff changeset
    27
  map       :: ('a=>'b) => ('a list => 'b list)
1475
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1419
diff changeset
    28
  mem       :: ['a, 'a list] => bool                    (infixl 55)
1370
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1327
diff changeset
    29
  nth       :: [nat, 'a list] => 'a
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1327
diff changeset
    30
  take      :: [nat, 'a list] => 'a list
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1327
diff changeset
    31
  tl,ttl    :: 'a list => 'a list
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1327
diff changeset
    32
  rev       :: 'a list => 'a list
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    33
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    34
syntax
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    35
  (* list Enumeration *)
1370
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1327
diff changeset
    36
  "@list"   :: args => 'a list                        ("[(_)]")
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    37
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    38
  (* Special syntax for list_all and filter *)
1475
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1419
diff changeset
    39
  "@Alls"       :: [idt, 'a list, bool] => bool ("(2Alls _:_./ _)" 10)
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1419
diff changeset
    40
  "@filter"     :: [idt, 'a list, bool] => 'a list      ("(1[_:_ ./ _])")
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    41
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    42
translations
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    43
  "[x, xs]"     == "x#[xs]"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    44
  "[x]"         == "x#[]"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    45
1475
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1419
diff changeset
    46
  "[x:xs . P]"  == "filter (%x.P) xs"
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1419
diff changeset
    47
  "Alls x:xs.P" == "list_all (%x.P) xs"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    48
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    49
primrec hd list
1898
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    50
  "hd([]) = (@x.False)"
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    51
  "hd(x#xs) = x"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    52
primrec tl list
1898
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    53
  "tl([]) = (@x.False)"
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    54
  "tl(x#xs) = xs"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    55
primrec ttl list
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    56
  (* a "total" version of tl: *)
1898
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    57
  "ttl([]) = []"
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    58
  "ttl(x#xs) = xs"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    59
primrec "op mem" list
1898
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    60
  "x mem [] = False"
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    61
  "x mem (y#ys) = (if y=x then True else x mem ys)"
1812
debfc40b7756 Addition of setOfList
paulson
parents: 1475
diff changeset
    62
primrec setOfList list
1898
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    63
  "setOfList [] = {}"
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    64
  "setOfList (x#xs) = insert x (setOfList xs)"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    65
primrec list_all list
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    66
  list_all_Nil  "list_all P [] = True"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    67
  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
    68
primrec map list
1898
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    69
  "map f [] = []"
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    70
  "map f (x#xs) = f(x)#map f xs"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    71
primrec "op @" list
1898
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    72
  "[] @ ys = ys"
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    73
  "(x#xs)@ys = x#(xs@ys)"
1169
5873833cf37f Added function rev and its properties length_rev, etc.
lcp
parents: 977
diff changeset
    74
primrec rev list
1898
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    75
  "rev([]) = []"
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    76
  "rev(x#xs) = rev(xs) @ [x]"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    77
primrec filter list
1898
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    78
  "filter P [] = []"
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    79
  "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
    80
primrec foldl list
1898
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    81
  "foldl f a [] = a"
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    82
  "foldl f a (x#xs) = foldl f (f a x) xs"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    83
primrec length list
1898
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    84
  "length([]) = 0"
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    85
  "length(x#xs) = Suc(length(xs))"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    86
primrec flat list
1898
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    87
  "flat([]) = []"
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    88
  "flat(x#xs) = x @ flat(xs)"
1419
a6a034a47a71 defined take/drop by induction over list rather than nat.
nipkow
parents: 1370
diff changeset
    89
primrec drop list
a6a034a47a71 defined take/drop by induction over list rather than nat.
nipkow
parents: 1370
diff changeset
    90
  drop_Nil  "drop n [] = []"
a6a034a47a71 defined take/drop by induction over list rather than nat.
nipkow
parents: 1370
diff changeset
    91
  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
    92
primrec take list
a6a034a47a71 defined take/drop by induction over list rather than nat.
nipkow
parents: 1370
diff changeset
    93
  take_Nil  "take n [] = []"
a6a034a47a71 defined take/drop by induction over list rather than nat.
nipkow
parents: 1370
diff changeset
    94
  take_Cons "take n (x#xs) = (case n of 0 => [] | Suc(m) => x # take m xs)"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    95
defs
1824
44254696843a Changed argument order of nat_rec.
berghofe
parents: 1812
diff changeset
    96
  nth_def  "nth(n) == nat_rec hd (%m r xs. r(tl(xs))) n"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    97
end
1898
260a9711f507 Simplified primrec definitions.
berghofe
parents: 1824
diff changeset
    98