src/ZF/List.thy
author paulson
Thu, 17 Jan 2002 12:45:52 +0100
changeset 12789 459b5de466b2
parent 9548 15bee2731e43
child 13327 be7105a066d3
permissions -rw-r--r--
new definitions from Sidi Ehmety
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
9548
15bee2731e43 instantiated Cancel_Numerals for "nat" in ZF
paulson
parents: 9491
diff changeset
    13
List = Datatype + ArithSimp +
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
12789
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
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                                       ("[]")
12789
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
    24
 "@List"     :: "is => i"                                 ("[(_)]")
2539
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
12789
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
    33
  length :: "i=>i"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
    34
  hd      :: "i=>i"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
    35
  tl      :: "i=>i"
6053
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
12789
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
    50
  map        :: "[i=>i, i] => i"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
    51
  set_of_list :: "i=>i"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
    52
  "@"        :: "[i,i]=>i"                        (infixr 60)
6070
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
12789
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
    68
  rev :: "i=>i"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
    69
  flat       :: "i=>i"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
    70
  list_add   :: "i=>i"
6070
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
12789
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
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
12789
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
    91
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
    92
(*** Thanks to Sidi Ehmety for the following ***)
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
    93
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
    94
constdefs
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
    95
(* Function `take' returns the first n elements of a list *)
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
    96
  take     :: "[i,i]=>i"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
    97
  "take(n, as) == list_rec(lam n:nat. [],
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
    98
		%a l r. lam n:nat. nat_case([], %m. Cons(a, r`m), n), as)`n"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
    99
  
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   100
(* Function `nth' returns the (n+1)th element in a list (not defined at Nil) *)
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   101
  nth :: "[i, i]=>i"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   102
  "nth(n, as) == list_rec(lam n:nat. 0,
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   103
		%a l r. lam n:nat. nat_case(a, %m. r`m, n), as)`n"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   104
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   105
  list_update :: "[i, i, i]=>i"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   106
  "list_update(xs, i, v) == list_rec(lam n:nat. Nil,
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   107
      %u us vs. lam n:nat. nat_case(Cons(v, us), %m. Cons(u, vs`m), n), xs)`i"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   108
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   109
consts
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   110
  filter :: "[i=>o, i] => i"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   111
  zip :: "[i, i]=>i"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   112
  ziprel :: "[i,i]=>i"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   113
  upt :: "[i, i] =>i"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   114
  
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   115
inductive domains "ziprel(A, B)" <= "list(A)*list(B)*list(A*B)"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   116
intrs   
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   117
  ziprel_Nil1  "ys:list(B) ==><Nil, ys, Nil>:ziprel(A, B)"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   118
  ziprel_Nil2  "xs:list(A) ==><xs, Nil, Nil>:ziprel(A, B)"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   119
  ziprel_Cons "[| <xs, ys, zs>:ziprel(A, B); x:A; y:B |]==>
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   120
	         <Cons(x, xs), Cons(y, ys), Cons(<x, y>, zs)>:ziprel(A,B)"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   121
  type_intrs  "list.intrs"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   122
  
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   123
defs
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   124
  zip_def "zip(xs, ys) ==
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   125
	      THE zs. <xs, ys, zs>:ziprel(set_of_list(xs),set_of_list(ys))"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   126
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   127
primrec
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   128
  "filter(P, Nil) = Nil"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   129
  "filter(P, Cons(x, xs)) =
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   130
     (if P(x) then Cons(x, filter(P, xs)) else filter(P, xs))"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   131
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   132
primrec
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   133
  "upt(i, 0) = Nil"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   134
  "upt(i, succ(j)) = (if i le j then upt(i, j)@[j] else Nil)"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   135
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   136
constdefs
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   137
  sublist :: "[i, i] => i"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   138
    "sublist(xs, A)== 
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   139
     map(fst, (filter(%p. snd(p): A, zip(xs, upt(0,length(xs))))))"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   140
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   141
  min :: "[i,i] =>i"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   142
    "min(x, y) == (if x le y then x else y)"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   143
  
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   144
  max :: "[i, i] =>i"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   145
    "max(x, y) == (if x le y then y else x)"
459b5de466b2 new definitions from Sidi Ehmety
paulson
parents: 9548
diff changeset
   146
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 435
diff changeset
   147
end