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