src/HOL/List.thy
author lcp
Thu, 29 Jun 1995 16:33:17 +0200
changeset 1169 5873833cf37f
parent 977 5d57287e5e1e
child 1327 6c29cfab679c
permissions -rw-r--r--
Added function rev and its properties length_rev, etc.
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
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     8
*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     9
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    10
List = Arith +
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    11
977
5d57287e5e1e changed syntax of datatype declarations (curried types for constructor
clasohm
parents: 965
diff changeset
    12
datatype 'a list = "[]" ("[]") | "#" 'a ('a list) (infixr 65)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    13
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    14
consts
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
  null      :: "'a list => bool"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    17
  hd        :: "'a list => 'a"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    18
  tl,ttl    :: "'a list => 'a list"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    19
  mem       :: "['a, 'a list] => bool"			(infixl 55)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    20
  list_all  :: "('a => bool) => ('a list => bool)"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    21
  map       :: "('a=>'b) => ('a list => 'b list)"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    22
  "@"	    :: "['a list, 'a list] => 'a list"		(infixr 65)
1169
5873833cf37f Added function rev and its properties length_rev, etc.
lcp
parents: 977
diff changeset
    23
  rev       :: "'a list => 'a list"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    24
  filter    :: "['a => bool, 'a list] => 'a list"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    25
  foldl     :: "[['b,'a] => 'b, 'b, 'a list] => 'b"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    26
  length    :: "'a list => nat"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    27
  flat      :: "'a list list => 'a list"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    28
  nth       :: "[nat, 'a list] => 'a"
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 *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    32
  "@list"   :: "args => 'a list"                        ("[(_)]")
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
  (* Special syntax for list_all and filter *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    35
  "@Alls"	:: "[idt, 'a list, bool] => bool"	("(2Alls _:_./ _)" 10)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    36
  "@filter"	:: "[idt, 'a list, bool] => 'a list"	("(1[_:_ ./ _])")
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
translations
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    39
  "[x, xs]"     == "x#[xs]"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    40
  "[x]"         == "x#[]"
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
  "[x:xs . P]"	== "filter (%x.P) xs"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    43
  "Alls x:xs.P"	== "list_all (%x.P) xs"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    44
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    45
primrec null list
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    46
  null_Nil "null([]) = True"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    47
  null_Cons "null(x#xs) = False"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    48
primrec hd list
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    49
  hd_Nil  "hd([]) = (@x.False)"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    50
  hd_Cons "hd(x#xs) = x"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    51
primrec tl list
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    52
  tl_Nil  "tl([]) = (@x.False)"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    53
  tl_Cons "tl(x#xs) = xs"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    54
primrec ttl list
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    55
  (* a "total" version of tl: *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    56
  ttl_Nil  "ttl([]) = []"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    57
  ttl_Cons "ttl(x#xs) = xs"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    58
primrec "op mem" list
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    59
  mem_Nil  "x mem [] = False"
965
24eef3860714 changed syntax of "if"
clasohm
parents: 923
diff changeset
    60
  mem_Cons "x mem (y#ys) = (if y=x then True else x mem ys)"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    61
primrec list_all list
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    62
  list_all_Nil  "list_all P [] = True"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    63
  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
    64
primrec map list
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    65
  map_Nil  "map f [] = []"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    66
  map_Cons "map f (x#xs) = f(x)#map f xs"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    67
primrec "op @" list
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    68
  append_Nil  "[] @ ys = ys"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    69
  append_Cons "(x#xs)@ys = x#(xs@ys)"
1169
5873833cf37f Added function rev and its properties length_rev, etc.
lcp
parents: 977
diff changeset
    70
primrec rev list
5873833cf37f Added function rev and its properties length_rev, etc.
lcp
parents: 977
diff changeset
    71
  rev_Nil  "rev([]) = []"
5873833cf37f Added function rev and its properties length_rev, etc.
lcp
parents: 977
diff changeset
    72
  rev_Cons "rev(x#xs) = rev(xs) @ [x]"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    73
primrec filter list
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    74
  filter_Nil  "filter P [] = []"
965
24eef3860714 changed syntax of "if"
clasohm
parents: 923
diff changeset
    75
  filter_Cons "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
    76
primrec foldl list
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    77
  foldl_Nil  "foldl f a [] = a"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    78
  foldl_Cons "foldl f a (x#xs) = foldl f (f a x) xs"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    79
primrec length list
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    80
  length_Nil  "length([]) = 0"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    81
  length_Cons "length(x#xs) = Suc(length(xs))"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    82
primrec flat list
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    83
  flat_Nil  "flat([]) = []"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    84
  flat_Cons "flat(x#xs) = x @ flat(xs)"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    85
defs
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    86
  nth_def "nth(n) == nat_rec n hd (%m r xs. r(tl(xs)))"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    87
end