src/HOL/Induct/LList.thy
author paulson
Thu, 26 Nov 1998 17:40:10 +0100
changeset 5977 9f0c8869cf71
parent 3842 b55686a7b22c
child 6382 8b0c9205da75
permissions -rw-r--r--
tidied up list definitions, using type 'a option instead of unit + 'a, also using real typedefs instead of faking them with extra rules
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
     1
(*  Title:      HOL/LList.thy
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
     2
    ID:         $Id$
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
     4
    Copyright   1994  University of Cambridge
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
     5
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
     6
Definition of type 'a llist by a greatest fixed point
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
     7
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
     8
Shares NIL, CONS, List_case with List.thy
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
     9
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    10
Still needs filter and flatten functions -- hard because they need
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    11
bounds on the amount of lookahead required.
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    12
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    13
Could try (but would it work for the gfp analogue of term?)
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3120
diff changeset
    14
  LListD_Fun_def "LListD_Fun(A) == (%Z. diag({Numb(0)}) <++> diag(A) <**> Z)"
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    15
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    16
A nice but complex example would be [ML for the Working Programmer, page 176]
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    17
  from(1) = enumerate (Lmap (Lmap(pack), makeqq(from(1),from(1))))
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    18
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    19
Previous definition of llistD_Fun was explicit:
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    20
  llistD_Fun_def
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    21
   "llistD_Fun(r) ==    
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    22
       {(LNil,LNil)}  Un        
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    23
       (UN x. (split(%l1 l2.(LCons(x,l1),LCons(x,l2))))``r)"
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    24
*)
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    25
5977
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    26
LList = Main + SList +
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    27
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    28
consts
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    29
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    30
  llist      :: 'a item set => 'a item set
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    31
  LListD     :: "('a item * 'a item)set => ('a item * 'a item)set"
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    32
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    33
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    34
coinductive "llist(A)"
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    35
  intrs
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    36
    NIL_I  "NIL: llist(A)"
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    37
    CONS_I "[| a: A;  M: llist(A) |] ==> CONS a M : llist(A)"
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    38
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    39
coinductive "LListD(r)"
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    40
  intrs
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    41
    NIL_I  "(NIL, NIL) : LListD(r)"
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    42
    CONS_I "[| (a,b): r;  (M,N) : LListD(r)   
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    43
            |] ==> (CONS a M, CONS b N) : LListD(r)"
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    44
5977
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    45
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    46
typedef (LList)
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    47
  'a llist = "llist(range Leaf)" (llist.NIL_I)
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    48
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    49
constdefs
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    50
  (*Now used exclusively for abbreviating the coinduction rule*)
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    51
  list_Fun   :: ['a item set, 'a item set] => 'a item set
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    52
     "list_Fun A X == {z. z = NIL | (? M a. z = CONS a M & a : A & M : X)}"
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    53
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    54
  LListD_Fun :: 
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    55
      "[('a item * 'a item)set, ('a item * 'a item)set] => 
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    56
       ('a item * 'a item)set"
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    57
    "LListD_Fun r X ==   
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    58
       {z. z = (NIL, NIL) |   
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    59
           (? M N a b. z = (CONS a M, CONS b N) & (a, b) : r & (M, N) : X)}"
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    60
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    61
  (*the abstract constructors*)
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    62
  LNil       :: 'a llist
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    63
    "LNil == Abs_LList NIL"
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    64
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    65
  LCons      :: ['a, 'a llist] => 'a llist
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    66
    "LCons x xs == Abs_LList(CONS (Leaf x) (Rep_LList xs))"
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    67
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    68
  llist_case :: ['b, ['a, 'a llist]=>'b, 'a llist] => 'b
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    69
    "llist_case c d l == 
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    70
       List_case c (%x y. d (inv Leaf x) (Abs_LList y)) (Rep_LList l)"
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    71
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    72
  LList_corec_fun :: "[nat, 'a=> ('b item * 'a) option, 'a] => 'b item"
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    73
    "LList_corec_fun k f == 
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    74
     nat_rec (%x. {})                         
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    75
             (%j r x. case f x of None    => NIL
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    76
                                | Some(z,w) => CONS z (r w)) 
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    77
             k"
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    78
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    79
  LList_corec     :: "['a, 'a => ('b item * 'a) option] => 'b item"
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    80
    "LList_corec a f == UN k. LList_corec_fun k f a"
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    81
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    82
  llist_corec     :: "['a, 'a => ('b * 'a) option] => 'b llist"
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    83
    "llist_corec a f == 
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    84
       Abs_LList(LList_corec a 
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    85
                 (%z. case f z of None      => None
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    86
                                | Some(v,w) => Some(Leaf(v), w)))"
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    87
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    88
  llistD_Fun :: "('a llist * 'a llist)set => ('a llist * 'a llist)set"
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    89
    "llistD_Fun(r) ==    
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    90
        prod_fun Abs_LList Abs_LList ``         
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    91
                LListD_Fun (diag(range Leaf))   
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    92
                            (prod_fun Rep_LList Rep_LList `` r)"
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    93
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    94
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    95
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
    96
(*The case syntax for type 'a llist*)
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    97
translations
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3120
diff changeset
    98
  "case p of LNil => a | LCons x l => b" == "llist_case a (%x l. b) p"
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
    99
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   100
5977
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
   101
(** Sample function definitions.  Item-based ones start with L ***)
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   102
5977
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
   103
constdefs
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
   104
  Lmap       :: ('a item => 'b item) => ('a item => 'b item)
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
   105
    "Lmap f M == LList_corec M (List_case None (%x M'. Some((f(x), M'))))"
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   106
5977
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
   107
  lmap       :: ('a=>'b) => ('a llist => 'b llist)
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
   108
    "lmap f l == llist_corec l (%z. case z of LNil => None 
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
   109
                                           | LCons y z => Some(f(y), z))"
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   110
5977
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
   111
  iterates   :: ['a => 'a, 'a] => 'a llist
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
   112
    "iterates f a == llist_corec a (%x. Some((x, f(x))))"     
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   113
5977
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
   114
  Lconst     :: 'a item => 'a item
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
   115
    "Lconst(M) == lfp(%N. CONS M N)"     
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   116
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   117
(*Append generates its result by applying f, where
5977
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
   118
    f((NIL,NIL))          = None
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
   119
    f((NIL, CONS N1 N2))  = Some((N1, (NIL,N2))
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
   120
    f((CONS M1 M2, N))    = Some((M1, (M2,N))
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   121
*)
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   122
5977
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
   123
  Lappend    :: ['a item, 'a item] => 'a item
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   124
   "Lappend M N == LList_corec (M,N)                                         
5977
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
   125
     (split(List_case (List_case None (%N1 N2. Some((N1, (NIL,N2))))) 
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
   126
                      (%M1 M2 N. Some((M1, (M2,N))))))"
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   127
5977
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
   128
  lappend    :: ['a llist, 'a llist] => 'a llist
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
   129
    "lappend l n == llist_corec (l,n)                                         
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
   130
       (split(llist_case (llist_case None (%n1 n2. Some((n1, (LNil,n2))))) 
9f0c8869cf71 tidied up list definitions, using type 'a option instead of
paulson
parents: 3842
diff changeset
   131
                         (%l1 l2 n. Some((l1, (l2,n))))))"
3120
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   132
c58423c20740 New directory to contain examples of (co)inductive definitions
paulson
parents:
diff changeset
   133
end