src/ZF/ex/LListFn.thy
author wenzelm
Tue, 21 Sep 1999 17:31:20 +0200
changeset 7566 c5a3f980a7af
parent 120 09287f26bfb8
permissions -rw-r--r--
accomodate refined facts handling;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     1
(*  Title: 	ZF/ex/llist-fn.thy
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     3
    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1993  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     6
Functions for Lazy Lists in Zermelo-Fraenkel Set Theory 
120
09287f26bfb8 changed all co- and co_ to co
lcp
parents: 0
diff changeset
     7
09287f26bfb8 changed all co- and co_ to co
lcp
parents: 0
diff changeset
     8
STILL NEEDS:
09287f26bfb8 changed all co- and co_ to co
lcp
parents: 0
diff changeset
     9
co_recursion for defining lconst, flip, etc.
09287f26bfb8 changed all co- and co_ to co
lcp
parents: 0
diff changeset
    10
a typing rule for it, based on some notion of "productivity..."
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    11
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    12
120
09287f26bfb8 changed all co- and co_ to co
lcp
parents: 0
diff changeset
    13
LListFn = LList + LList_Eq +
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    14
consts
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    15
  lconst   :: "i => i"
120
09287f26bfb8 changed all co- and co_ to co
lcp
parents: 0
diff changeset
    16
  flip     :: "i => i"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    17
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    18
rules
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    19
  lconst_def  "lconst(a) == lfp(univ(a), %l. LCons(a,l))"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    20
120
09287f26bfb8 changed all co- and co_ to co
lcp
parents: 0
diff changeset
    21
  flip_LNil   "flip(LNil) = LNil"
09287f26bfb8 changed all co- and co_ to co
lcp
parents: 0
diff changeset
    22
09287f26bfb8 changed all co- and co_ to co
lcp
parents: 0
diff changeset
    23
  flip_LCons  "[| x:bool; l: llist(bool) |] ==> \
09287f26bfb8 changed all co- and co_ to co
lcp
parents: 0
diff changeset
    24
\              flip(LCons(x,l)) = LCons(not(x), flip(l))"
09287f26bfb8 changed all co- and co_ to co
lcp
parents: 0
diff changeset
    25
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    26
end