src/HOL/ex/LList.thy
author clasohm
Fri, 01 Dec 1995 12:03:13 +0100
changeset 1376 92f83b9d17e1
parent 1151 c820b3cc3df0
child 1476 608483c2122a
permissions -rw-r--r--
removed quotes from consts and syntax sections
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     1
(*  Title: 	HOL/LList.thy
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     3
    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     4
    Copyright   1994  University of Cambridge
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     5
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     6
Definition of type 'a llist by a greatest fixed point
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     7
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     8
Shares NIL, CONS, List_case with List.thy
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
     9
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    10
Still needs filter and flatten functions -- hard because they need
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    11
bounds on the amount of lookahead required.
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    12
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    13
Could try (but would it work for the gfp analogue of term?)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    14
  LListD_Fun_def "LListD_Fun(A) == (%Z.diag({Numb(0)}) <++> diag(A) <**> Z)"
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    15
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    16
A nice but complex example would be [ML for the Working Programmer, page 176]
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    17
  from(1) = enumerate (Lmap (Lmap(pack), makeqq(from(1),from(1))))
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    18
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    19
Previous definition of llistD_Fun was explicit:
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    20
  llistD_Fun_def
1151
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 972
diff changeset
    21
   "llistD_Fun(r) == 	
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 972
diff changeset
    22
       {(LNil,LNil)}  Un  	
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 972
diff changeset
    23
       (UN x. (split(%l1 l2.(LCons(x,l1),LCons(x,l2))))``r)"
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    24
*)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    25
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    26
LList = Gfp + SList +
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    27
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    28
types
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    29
  'a llist
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    30
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    31
arities
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    32
   llist :: (term)term
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    33
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    34
consts
1376
92f83b9d17e1 removed quotes from consts and syntax sections
clasohm
parents: 1151
diff changeset
    35
  list_Fun   :: ['a item set, 'a item set] => 'a item set
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    36
  LListD_Fun :: 
1151
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 972
diff changeset
    37
      "[('a item * 'a item)set, ('a item * 'a item)set] => 
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 972
diff changeset
    38
      ('a item * 'a item)set"
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    39
1376
92f83b9d17e1 removed quotes from consts and syntax sections
clasohm
parents: 1151
diff changeset
    40
  llist      :: 'a item set => 'a item set
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    41
  LListD     :: "('a item * 'a item)set => ('a item * 'a item)set"
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    42
  llistD_Fun :: "('a llist * 'a llist)set => ('a llist * 'a llist)set"
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    43
1376
92f83b9d17e1 removed quotes from consts and syntax sections
clasohm
parents: 1151
diff changeset
    44
  Rep_llist  :: 'a llist => 'a item
92f83b9d17e1 removed quotes from consts and syntax sections
clasohm
parents: 1151
diff changeset
    45
  Abs_llist  :: 'a item => 'a llist
92f83b9d17e1 removed quotes from consts and syntax sections
clasohm
parents: 1151
diff changeset
    46
  LNil       :: 'a llist
92f83b9d17e1 removed quotes from consts and syntax sections
clasohm
parents: 1151
diff changeset
    47
  LCons      :: ['a, 'a llist] => 'a llist
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    48
  
1376
92f83b9d17e1 removed quotes from consts and syntax sections
clasohm
parents: 1151
diff changeset
    49
  llist_case :: ['b, ['a, 'a llist]=>'b, 'a llist] => 'b
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    50
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    51
  LList_corec_fun :: "[nat, 'a=>unit+('b item * 'a), 'a] => 'b item"
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    52
  LList_corec     :: "['a, 'a => unit + ('b item * 'a)] => 'b item"
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    53
  llist_corec     :: "['a, 'a => unit + ('b * 'a)] => 'b llist"
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    54
1376
92f83b9d17e1 removed quotes from consts and syntax sections
clasohm
parents: 1151
diff changeset
    55
  Lmap	     :: ('a item => 'b item) => ('a item => 'b item)
92f83b9d17e1 removed quotes from consts and syntax sections
clasohm
parents: 1151
diff changeset
    56
  lmap	     :: ('a=>'b) => ('a llist => 'b llist)
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    57
1376
92f83b9d17e1 removed quotes from consts and syntax sections
clasohm
parents: 1151
diff changeset
    58
  iterates   :: ['a => 'a, 'a] => 'a llist
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    59
1376
92f83b9d17e1 removed quotes from consts and syntax sections
clasohm
parents: 1151
diff changeset
    60
  Lconst     :: 'a item => 'a item
92f83b9d17e1 removed quotes from consts and syntax sections
clasohm
parents: 1151
diff changeset
    61
  Lappend    :: ['a item, 'a item] => 'a item
92f83b9d17e1 removed quotes from consts and syntax sections
clasohm
parents: 1151
diff changeset
    62
  lappend    :: ['a llist, 'a llist] => 'a llist
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    63
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    64
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    65
coinductive "llist(A)"
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    66
  intrs
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    67
    NIL_I  "NIL: llist(A)"
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    68
    CONS_I "[| a: A;  M: llist(A) |] ==> CONS a M : llist(A)"
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    69
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    70
coinductive "LListD(r)"
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    71
  intrs
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 969
diff changeset
    72
    NIL_I  "(NIL, NIL) : LListD(r)"
1151
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 972
diff changeset
    73
    CONS_I "[| (a,b): r;  (M,N) : LListD(r)   
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 972
diff changeset
    74
	    |] ==> (CONS a M, CONS b N) : LListD(r)"
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    75
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    76
defs
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    77
  (*Now used exclusively for abbreviating the coinduction rule*)
1151
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 972
diff changeset
    78
  list_Fun_def   "list_Fun A X ==   
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 972
diff changeset
    79
		  {z. z = NIL | (? M a. z = CONS a M & a : A & M : X)}"
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    80
1151
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 972
diff changeset
    81
  LListD_Fun_def "LListD_Fun r X ==   
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 972
diff changeset
    82
		  {z. z = (NIL, NIL) |   
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 972
diff changeset
    83
		      (? M N a b. z = (CONS a M, CONS b N) &   
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 972
diff changeset
    84
		                  (a, b) : r & (M, N) : X)}"
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    85
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    86
  (*defining the abstract constructors*)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    87
  LNil_def  	"LNil == Abs_llist(NIL)"
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    88
  LCons_def 	"LCons x xs == Abs_llist(CONS (Leaf x) (Rep_llist xs))"
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    89
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    90
  llist_case_def
1151
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 972
diff changeset
    91
   "llist_case c d l == 
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 972
diff changeset
    92
       List_case c (%x y. d (Inv Leaf x) (Abs_llist y)) (Rep_llist l)"
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    93
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    94
  LList_corec_fun_def
1151
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 972
diff changeset
    95
    "LList_corec_fun k f == 
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 972
diff changeset
    96
     nat_rec k (%x. {}) 			
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 972
diff changeset
    97
	      (%j r x. sum_case (%u.NIL) (split(%z w. CONS z (r w))) (f x))"
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    98
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
    99
  LList_corec_def
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   100
    "LList_corec a f == UN k. LList_corec_fun k f a"
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   101
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   102
  llist_corec_def
1151
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 972
diff changeset
   103
   "llist_corec a f == 
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 972
diff changeset
   104
       Abs_llist(LList_corec a (%z.sum_case (%x.Inl(x)) 
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 972
diff changeset
   105
                                    (split(%v w. Inr((Leaf(v), w)))) (f z)))"
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   106
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   107
  llistD_Fun_def
1151
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 972
diff changeset
   108
   "llistD_Fun(r) == 	
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 972
diff changeset
   109
	prod_fun Abs_llist Abs_llist ``  	
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 972
diff changeset
   110
                LListD_Fun (diag(range Leaf))	
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 972
diff changeset
   111
		            (prod_fun Rep_llist Rep_llist `` r)"
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   112
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   113
  Lconst_def	"Lconst(M) == lfp(%N. CONS M N)"     
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   114
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   115
  Lmap_def
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 969
diff changeset
   116
   "Lmap f M == LList_corec M (List_case (Inl ()) (%x M'. Inr((f(x), M'))))"
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   117
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   118
  lmap_def
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 969
diff changeset
   119
   "lmap f l == llist_corec l (llist_case (Inl ()) (%y z. Inr((f(y), z))))"
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   120
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 969
diff changeset
   121
  iterates_def	"iterates f a == llist_corec a (%x. Inr((x, f(x))))"     
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   122
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   123
(*Append generates its result by applying f, where
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 969
diff changeset
   124
    f((NIL,NIL)) = Inl(())
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 969
diff changeset
   125
    f((NIL, CONS N1 N2)) = Inr((N1, (NIL,N2))
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 969
diff changeset
   126
    f((CONS M1 M2, N))    = Inr((M1, (M2,N))
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   127
*)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   128
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   129
  Lappend_def
1151
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 972
diff changeset
   130
   "Lappend M N == LList_corec (M,N)	   				     
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 972
diff changeset
   131
     (split(List_case (List_case (Inl ()) (%N1 N2. Inr((N1, (NIL,N2))))) 
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 972
diff changeset
   132
                      (%M1 M2 N. Inr((M1, (M2,N))))))"
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   133
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   134
  lappend_def
1151
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 972
diff changeset
   135
   "lappend l n == llist_corec (l,n)	   				     
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 972
diff changeset
   136
   (split(llist_case (llist_case (Inl ()) (%n1 n2. Inr((n1, (LNil,n2))))) 
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 972
diff changeset
   137
                     (%l1 l2 n. Inr((l1, (l2,n))))))"
969
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   138
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   139
rules
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   140
    (*faking a type definition...*)
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   141
  Rep_llist 	    "Rep_llist(xs): llist(range(Leaf))"
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   142
  Rep_llist_inverse "Abs_llist(Rep_llist(xs)) = xs"
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   143
  Abs_llist_inverse "M: llist(range(Leaf)) ==> Rep_llist(Abs_llist(M)) = M"
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   144
b051e2fc2e34 converted ex with curried function application
clasohm
parents:
diff changeset
   145
end