src/ZF/ex/LList.thy
author wenzelm
Fri, 26 Oct 2001 18:16:45 +0200
changeset 11949 38e20c036e37
parent 11354 9b80fe19407f
child 12867 5c900a821a7c
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     1
(*  Title:      ZF/ex/LList.thy
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     2
    ID:         $Id$
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     4
    Copyright   1994  University of Cambridge
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     5
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     6
Codatatype definition of Lazy Lists
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     7
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     8
Equality for llist(A) as a greatest fixed point
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     9
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    10
Functions for Lazy Lists
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    11
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    12
STILL NEEDS:
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    13
co_recursion for defining lconst, flip, etc.
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    14
a typing rule for it, based on some notion of "productivity..."
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    15
*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    16
11354
9b80fe19407f examples files start from Main instead of various ZF theories
paulson
parents: 11316
diff changeset
    17
LList = Main +
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    18
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    19
consts
1401
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 1155
diff changeset
    20
  llist  :: i=>i
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    21
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    22
codatatype
11316
b4e71bd751e4 X-symbols for set theory
paulson
parents: 1478
diff changeset
    23
  "llist(A)" = LNil | LCons ("a \\<in> A", "l \\<in> llist(A)")
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    24
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    25
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    26
(*Coinductive definition of equality*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    27
consts
1401
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 1155
diff changeset
    28
  lleq :: i=>i
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    29
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    30
(*Previously used <*> in the domain and variant pairs as elements.  But
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    31
  standard pairs work just as well.  To use variant pairs, must change prefix
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    32
  a q/Q to the Sigma, Pair and converse rules.*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    33
coinductive
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    34
  domains "lleq(A)" <= "llist(A) * llist(A)"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    35
  intrs
11316
b4e71bd751e4 X-symbols for set theory
paulson
parents: 1478
diff changeset
    36
    LNil  "<LNil, LNil> \\<in> lleq(A)"
b4e71bd751e4 X-symbols for set theory
paulson
parents: 1478
diff changeset
    37
    LCons "[| a \\<in> A; <l,l'>: lleq(A) |] ==> <LCons(a,l), LCons(a,l')>: lleq(A)"
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    38
  type_intrs  "llist.intrs"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    39
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    40
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    41
(*Lazy list functions; flip is not definitional!*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    42
consts
1401
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 1155
diff changeset
    43
  lconst   :: i => i
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 1155
diff changeset
    44
  flip     :: i => i
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    45
753
ec86863e87c8 replaced "rules" by "defs"
lcp
parents: 515
diff changeset
    46
defs
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    47
  lconst_def  "lconst(a) == lfp(univ(a), %l. LCons(a,l))"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    48
753
ec86863e87c8 replaced "rules" by "defs"
lcp
parents: 515
diff changeset
    49
rules
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    50
  flip_LNil   "flip(LNil) = LNil"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    51
11316
b4e71bd751e4 X-symbols for set theory
paulson
parents: 1478
diff changeset
    52
  flip_LCons  "[| x \\<in> bool; l \\<in> llist(bool) |] ==> 
1155
928a16e02f9f removed \...\ inside strings
clasohm
parents: 810
diff changeset
    53
              flip(LCons(x,l)) = LCons(not(x), flip(l))"
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    54
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    55
end
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    56