List.thy
author nipkow
Wed, 08 Feb 1995 11:34:11 +0100
changeset 210 1a3d3b5b5d15
parent 203 d465d3be2744
child 212 2740293cc458
permissions -rw-r--r--
More rewrite rules.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
196
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
     1
(*  Title:      HOL/List.thy
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
196
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
     3
    Author:     Tobias Nipkow
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
     4
    Copyright   1994 TU Muenchen
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     5
196
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
     6
Definition of type 'a list as a datatype. This allows primrec to work.
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     7
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     8
*)
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     9
196
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    10
List = Arith +
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    11
196
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    12
datatype 'a list = "[]" ("[]") | "#"('a,'a list) (infixr 65)
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    13
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    14
consts
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    15
113
0b9b8eb74101 HOL/List: rotated arguments of List_case, list_case
lcp
parents: 64
diff changeset
    16
  null      :: "'a list => bool"
0b9b8eb74101 HOL/List: rotated arguments of List_case, list_case
lcp
parents: 64
diff changeset
    17
  hd        :: "'a list => 'a"
0b9b8eb74101 HOL/List: rotated arguments of List_case, list_case
lcp
parents: 64
diff changeset
    18
  tl,ttl    :: "'a list => 'a list"
196
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    19
  mem       :: "['a, 'a list] => bool"			(infixl 55)
113
0b9b8eb74101 HOL/List: rotated arguments of List_case, list_case
lcp
parents: 64
diff changeset
    20
  list_all  :: "('a => bool) => ('a list => bool)"
0b9b8eb74101 HOL/List: rotated arguments of List_case, list_case
lcp
parents: 64
diff changeset
    21
  map       :: "('a=>'b) => ('a list => 'b list)"
196
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    22
  "@"	    :: "['a list, 'a list] => 'a list"		(infixr 65)
113
0b9b8eb74101 HOL/List: rotated arguments of List_case, list_case
lcp
parents: 64
diff changeset
    23
  filter    :: "['a => bool, 'a list] => 'a list"
196
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    24
  foldl     :: "[['b,'a] => 'b, 'b, 'a list] => 'b"
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    25
  length    :: "'a list => nat"
203
d465d3be2744 Added "nth" and some lemmas.
nipkow
parents: 196
diff changeset
    26
  nth       :: "[nat, 'a list] => 'a"
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    27
196
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    28
syntax
128
89669c58e506 INSTALLATION OF INDUCTIVE DEFINITIONS
lcp
parents: 113
diff changeset
    29
  (* list Enumeration *)
196
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    30
  "@list"   :: "args => 'a list"                        ("[(_)]")
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    31
34
7d437bed7765 added conj_assoc to HOL_ss
nipkow
parents: 13
diff changeset
    32
  (* Special syntax for list_all and filter *)
39
91614f0eb250 id -> idt in type of @filter and @Alls
nipkow
parents: 34
diff changeset
    33
  "@Alls"	:: "[idt, 'a list, bool] => bool"	("(2Alls _:_./ _)" 10)
91614f0eb250 id -> idt in type of @filter and @Alls
nipkow
parents: 34
diff changeset
    34
  "@filter"	:: "[idt, 'a list, bool] => 'a list"	("(1[_:_ ./ _])")
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    35
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    36
translations
48
21291189b51e changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents: 42
diff changeset
    37
  "[x, xs]"     == "x#[xs]"
21291189b51e changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents: 42
diff changeset
    38
  "[x]"         == "x#[]"
42
87f6e8b93221 added translations for "case xs of"
nipkow
parents: 40
diff changeset
    39
34
7d437bed7765 added conj_assoc to HOL_ss
nipkow
parents: 13
diff changeset
    40
  "[x:xs . P]"	== "filter(%x.P,xs)"
7d437bed7765 added conj_assoc to HOL_ss
nipkow
parents: 13
diff changeset
    41
  "Alls x:xs.P"	== "list_all(%x.P,xs)"
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    42
196
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    43
primrec null list
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    44
  null_Nil "null([]) = True"
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    45
  null_Cons "null(x#xs) = False"
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    46
primrec hd list
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    47
  hd_Nil  "hd([]) = (@x.False)"
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    48
  hd_Cons "hd(x#xs) = x"
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    49
primrec tl list
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    50
  tl_Nil  "tl([]) = (@x.False)"
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    51
  tl_Cons "tl(x#xs) = xs"
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    52
primrec ttl list
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    53
  (* a "total" version of tl: *)
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    54
  ttl_Nil  "ttl([]) = []"
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    55
  ttl_Cons "ttl(x#xs) = xs"
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    56
primrec "op mem" list
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    57
  mem_Nil  "x mem [] = False"
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    58
  mem_Cons "x mem (y#ys) = if(y=x, True, x mem ys)"
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    59
primrec list_all list
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    60
  list_all_Nil  "list_all(P,[]) = True"
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    61
  list_all_Cons "list_all(P,x#xs) = (P(x) & list_all(P,xs))"
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    62
primrec map list
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    63
  map_Nil  "map(f,[]) = []"
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    64
  map_Cons "map(f,x#xs) = f(x)#map(f,xs)"
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    65
primrec "op @" list
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    66
  append_Nil  "[] @ ys = ys"
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    67
  append_Cons "(x#xs)@ys = x#(xs@ys)"
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    68
primrec filter list
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    69
  filter_Nil  "filter(P,[]) = []"
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    70
  filter_Cons "filter(P,x#xs) = if(P(x), x#filter(P,xs), filter(P,xs))"
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    71
primrec foldl list
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    72
  foldl_Nil  "foldl(f,a,[]) = a"
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    73
  foldl_Cons "foldl(f,a,x#xs) = foldl(f, f(a,x), xs)"
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    74
primrec length list
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    75
  length_Nil  "length([]) = 0"
61620d959717 Moved the old List to ex and replaced it by one defined via
nipkow
parents: 128
diff changeset
    76
  length_Cons "length(x#xs) = Suc(length(xs))"
203
d465d3be2744 Added "nth" and some lemmas.
nipkow
parents: 196
diff changeset
    77
defs
d465d3be2744 Added "nth" and some lemmas.
nipkow
parents: 196
diff changeset
    78
  nth_def "nth(n) == nat_rec(n, hd, %m r xs. r(tl(xs)))"
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    79
end