src/ZF/List.thy
author wenzelm
Wed, 07 Apr 1999 15:43:16 +0200
changeset 6411 07e95e4cfefe
parent 6070 032babd0120b
child 9491 1a36151ee2fc
permissions -rw-r--r--
fixed @@;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     1
(*  Title:      ZF/List
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
     2
    ID:         $Id$
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
     4
    Copyright   1994  University of Cambridge
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
     5
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
     6
Lists in Zermelo-Fraenkel Set Theory 
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
     7
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
     8
map is a binding operator -- it applies to meta-level functions, not 
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
     9
object-level functions.  This simplifies the final form of term_rec_conv,
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
    10
although complicating its derivation.
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
    11
*)
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
    12
810
91c68f74f458 removed quotes around "Datatype",
lcp
parents: 753
diff changeset
    13
List = Datatype + 
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
    14
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
    15
consts
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    16
  list       :: i=>i
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
    17
  
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
    18
datatype
581
465075fd257b removal of needless quotes
lcp
parents: 516
diff changeset
    19
  "list(A)" = Nil | Cons ("a:A", "l: list(A)")
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
    20
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
    21
2539
ddd22ceee8cc turned some consts into syntax;
wenzelm
parents: 1926
diff changeset
    22
syntax
ddd22ceee8cc turned some consts into syntax;
wenzelm
parents: 1926
diff changeset
    23
 "[]"        :: i                                       ("[]")
ddd22ceee8cc turned some consts into syntax;
wenzelm
parents: 1926
diff changeset
    24
 "@List"     :: is => i                                 ("[(_)]")
ddd22ceee8cc turned some consts into syntax;
wenzelm
parents: 1926
diff changeset
    25
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
    26
translations
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
    27
  "[x, xs]"     == "Cons(x, [xs])"
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
    28
  "[x]"         == "Cons(x, [])"
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
    29
  "[]"          == "Nil"
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
    30
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
    31
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    32
consts
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    33
  length :: i=>i
1806
12708740f58d Converted to use constdefs instead of defs
paulson
parents: 1478
diff changeset
    34
  hd      :: i=>i
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    35
  tl      :: i=>i
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    36
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    37
primrec
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    38
  "length([]) = 0"
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    39
  "length(Cons(a,l)) = succ(length(l))"
1806
12708740f58d Converted to use constdefs instead of defs
paulson
parents: 1478
diff changeset
    40
  
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    41
primrec
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    42
  "hd(Cons(a,l)) = a"
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    43
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    44
primrec
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    45
  "tl([]) = []"
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    46
  "tl(Cons(a,l)) = l"
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    47
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6053
diff changeset
    48
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6053
diff changeset
    49
consts
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6053
diff changeset
    50
  map        :: [i=>i, i] => i
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6053
diff changeset
    51
  set_of_list :: i=>i
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6053
diff changeset
    52
  "@"        :: [i,i]=>i                        (infixr 60)
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6053
diff changeset
    53
  
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    54
primrec
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    55
  "map(f,[]) = []"
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    56
  "map(f,Cons(a,l)) = Cons(f(a), map(f,l))"
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    57
 
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    58
primrec
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    59
  "set_of_list([]) = 0"
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    60
  "set_of_list(Cons(a,l)) = cons(a, set_of_list(l))"
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    61
   
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    62
primrec
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    63
  "[] @ ys = ys"
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    64
  "(Cons(a,l)) @ ys = Cons(a, l @ ys)"
1806
12708740f58d Converted to use constdefs instead of defs
paulson
parents: 1478
diff changeset
    65
  
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6053
diff changeset
    66
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6053
diff changeset
    67
consts
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6053
diff changeset
    68
  rev :: i=>i
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6053
diff changeset
    69
  flat       :: i=>i
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6053
diff changeset
    70
  list_add   :: i=>i
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6053
diff changeset
    71
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    72
primrec
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    73
  "rev([]) = []"
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    74
  "rev(Cons(a,l)) = rev(l) @ [a]"
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    75
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    76
primrec
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    77
  "flat([]) = []"
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    78
  "flat(Cons(l,ls)) = l @ flat(ls)"
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    79
   
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    80
primrec
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    81
  "list_add([]) = 0"
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    82
  "list_add(Cons(a,l)) = a #+ list_add(l)"
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 3840
diff changeset
    83
       
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6053
diff changeset
    84
consts
1806
12708740f58d Converted to use constdefs instead of defs
paulson
parents: 1478
diff changeset
    85
  drop       :: [i,i]=>i
6070
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6053
diff changeset
    86
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6053
diff changeset
    87
primrec
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6053
diff changeset
    88
  drop_0    "drop(0,l) = l"
032babd0120b ZF: the natural numbers as a datatype
paulson
parents: 6053
diff changeset
    89
  drop_SUCC "drop(succ(i), l) = tl (drop(i,l))"
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
    90
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
    91
end