src/HOL/Induct/LList.thy
changeset 5977 9f0c8869cf71
parent 3842 b55686a7b22c
child 6382 8b0c9205da75
equal deleted inserted replaced
5976:44290b71a85f 5977:9f0c8869cf71
    21    "llistD_Fun(r) ==    
    21    "llistD_Fun(r) ==    
    22        {(LNil,LNil)}  Un        
    22        {(LNil,LNil)}  Un        
    23        (UN x. (split(%l1 l2.(LCons(x,l1),LCons(x,l2))))``r)"
    23        (UN x. (split(%l1 l2.(LCons(x,l1),LCons(x,l2))))``r)"
    24 *)
    24 *)
    25 
    25 
    26 LList = Gfp + SList +
    26 LList = Main + SList +
    27 
       
    28 types
       
    29   'a llist
       
    30 
       
    31 arities
       
    32    llist :: (term)term
       
    33 
    27 
    34 consts
    28 consts
    35   list_Fun   :: ['a item set, 'a item set] => 'a item set
       
    36   LListD_Fun :: 
       
    37       "[('a item * 'a item)set, ('a item * 'a item)set] => 
       
    38       ('a item * 'a item)set"
       
    39 
    29 
    40   llist      :: 'a item set => 'a item set
    30   llist      :: 'a item set => 'a item set
    41   LListD     :: "('a item * 'a item)set => ('a item * 'a item)set"
    31   LListD     :: "('a item * 'a item)set => ('a item * 'a item)set"
    42   llistD_Fun :: "('a llist * 'a llist)set => ('a llist * 'a llist)set"
       
    43 
       
    44   Rep_llist  :: 'a llist => 'a item
       
    45   Abs_llist  :: 'a item => 'a llist
       
    46   LNil       :: 'a llist
       
    47   LCons      :: ['a, 'a llist] => 'a llist
       
    48   
       
    49   llist_case :: ['b, ['a, 'a llist]=>'b, 'a llist] => 'b
       
    50 
       
    51   LList_corec_fun :: "[nat, 'a=>unit+('b item * 'a), 'a] => 'b item"
       
    52   LList_corec     :: "['a, 'a => unit + ('b item * 'a)] => 'b item"
       
    53   llist_corec     :: "['a, 'a => unit + ('b * 'a)] => 'b llist"
       
    54 
       
    55   Lmap       :: ('a item => 'b item) => ('a item => 'b item)
       
    56   lmap       :: ('a=>'b) => ('a llist => 'b llist)
       
    57 
       
    58   iterates   :: ['a => 'a, 'a] => 'a llist
       
    59 
       
    60   Lconst     :: 'a item => 'a item
       
    61   Lappend    :: ['a item, 'a item] => 'a item
       
    62   lappend    :: ['a llist, 'a llist] => 'a llist
       
    63 
    32 
    64 
    33 
    65 coinductive "llist(A)"
    34 coinductive "llist(A)"
    66   intrs
    35   intrs
    67     NIL_I  "NIL: llist(A)"
    36     NIL_I  "NIL: llist(A)"
    71   intrs
    40   intrs
    72     NIL_I  "(NIL, NIL) : LListD(r)"
    41     NIL_I  "(NIL, NIL) : LListD(r)"
    73     CONS_I "[| (a,b): r;  (M,N) : LListD(r)   
    42     CONS_I "[| (a,b): r;  (M,N) : LListD(r)   
    74             |] ==> (CONS a M, CONS b N) : LListD(r)"
    43             |] ==> (CONS a M, CONS b N) : LListD(r)"
    75 
    44 
       
    45 
       
    46 typedef (LList)
       
    47   'a llist = "llist(range Leaf)" (llist.NIL_I)
       
    48 
       
    49 constdefs
       
    50   (*Now used exclusively for abbreviating the coinduction rule*)
       
    51   list_Fun   :: ['a item set, 'a item set] => 'a item set
       
    52      "list_Fun A X == {z. z = NIL | (? M a. z = CONS a M & a : A & M : X)}"
       
    53 
       
    54   LListD_Fun :: 
       
    55       "[('a item * 'a item)set, ('a item * 'a item)set] => 
       
    56        ('a item * 'a item)set"
       
    57     "LListD_Fun r X ==   
       
    58        {z. z = (NIL, NIL) |   
       
    59            (? M N a b. z = (CONS a M, CONS b N) & (a, b) : r & (M, N) : X)}"
       
    60 
       
    61   (*the abstract constructors*)
       
    62   LNil       :: 'a llist
       
    63     "LNil == Abs_LList NIL"
       
    64 
       
    65   LCons      :: ['a, 'a llist] => 'a llist
       
    66     "LCons x xs == Abs_LList(CONS (Leaf x) (Rep_LList xs))"
       
    67 
       
    68   llist_case :: ['b, ['a, 'a llist]=>'b, 'a llist] => 'b
       
    69     "llist_case c d l == 
       
    70        List_case c (%x y. d (inv Leaf x) (Abs_LList y)) (Rep_LList l)"
       
    71 
       
    72   LList_corec_fun :: "[nat, 'a=> ('b item * 'a) option, 'a] => 'b item"
       
    73     "LList_corec_fun k f == 
       
    74      nat_rec (%x. {})                         
       
    75              (%j r x. case f x of None    => NIL
       
    76                                 | Some(z,w) => CONS z (r w)) 
       
    77              k"
       
    78 
       
    79   LList_corec     :: "['a, 'a => ('b item * 'a) option] => 'b item"
       
    80     "LList_corec a f == UN k. LList_corec_fun k f a"
       
    81 
       
    82   llist_corec     :: "['a, 'a => ('b * 'a) option] => 'b llist"
       
    83     "llist_corec a f == 
       
    84        Abs_LList(LList_corec a 
       
    85                  (%z. case f z of None      => None
       
    86                                 | Some(v,w) => Some(Leaf(v), w)))"
       
    87 
       
    88   llistD_Fun :: "('a llist * 'a llist)set => ('a llist * 'a llist)set"
       
    89     "llistD_Fun(r) ==    
       
    90         prod_fun Abs_LList Abs_LList ``         
       
    91                 LListD_Fun (diag(range Leaf))   
       
    92                             (prod_fun Rep_LList Rep_LList `` r)"
       
    93 
       
    94 
       
    95 
       
    96 (*The case syntax for type 'a llist*)
    76 translations
    97 translations
    77   "case p of LNil => a | LCons x l => b" == "llist_case a (%x l. b) p"
    98   "case p of LNil => a | LCons x l => b" == "llist_case a (%x l. b) p"
    78 
    99 
    79 
   100 
    80 defs
   101 (** Sample function definitions.  Item-based ones start with L ***)
    81   (*Now used exclusively for abbreviating the coinduction rule*)
       
    82   list_Fun_def   "list_Fun A X ==   
       
    83                   {z. z = NIL | (? M a. z = CONS a M & a : A & M : X)}"
       
    84 
   102 
    85   LListD_Fun_def "LListD_Fun r X ==   
   103 constdefs
    86                   {z. z = (NIL, NIL) |   
   104   Lmap       :: ('a item => 'b item) => ('a item => 'b item)
    87                       (? M N a b. z = (CONS a M, CONS b N) &   
   105     "Lmap f M == LList_corec M (List_case None (%x M'. Some((f(x), M'))))"
    88                                   (a, b) : r & (M, N) : X)}"
       
    89 
   106 
    90   (*defining the abstract constructors*)
   107   lmap       :: ('a=>'b) => ('a llist => 'b llist)
    91   LNil_def      "LNil == Abs_llist(NIL)"
   108     "lmap f l == llist_corec l (%z. case z of LNil => None 
    92   LCons_def     "LCons x xs == Abs_llist(CONS (Leaf x) (Rep_llist xs))"
   109                                            | LCons y z => Some(f(y), z))"
    93 
   110 
    94   llist_case_def
   111   iterates   :: ['a => 'a, 'a] => 'a llist
    95    "llist_case c d l == 
   112     "iterates f a == llist_corec a (%x. Some((x, f(x))))"     
    96        List_case c (%x y. d (inv Leaf x) (Abs_llist y)) (Rep_llist l)"
       
    97 
   113 
    98   LList_corec_fun_def
   114   Lconst     :: 'a item => 'a item
    99     "LList_corec_fun k f == 
   115     "Lconst(M) == lfp(%N. CONS M N)"     
   100      nat_rec (%x. {})                         
       
   101              (%j r x. case f x of Inl u    => NIL
       
   102                                 | Inr(z,w) => CONS z (r w)) 
       
   103              k"
       
   104 
       
   105   LList_corec_def
       
   106     "LList_corec a f == UN k. LList_corec_fun k f a"
       
   107 
       
   108   llist_corec_def
       
   109    "llist_corec a f == 
       
   110        Abs_llist(LList_corec a 
       
   111                  (%z. case f z of Inl x    => Inl(x)
       
   112                                | Inr(v,w) => Inr(Leaf(v), w)))"
       
   113 
       
   114   llistD_Fun_def
       
   115    "llistD_Fun(r) ==    
       
   116         prod_fun Abs_llist Abs_llist ``         
       
   117                 LListD_Fun (diag(range Leaf))   
       
   118                             (prod_fun Rep_llist Rep_llist `` r)"
       
   119 
       
   120   Lconst_def    "Lconst(M) == lfp(%N. CONS M N)"     
       
   121 
       
   122   Lmap_def
       
   123    "Lmap f M == LList_corec M (List_case (Inl ()) (%x M'. Inr((f(x), M'))))"
       
   124 
       
   125   lmap_def
       
   126    "lmap f l == llist_corec l (%z. case z of LNil => (Inl ()) 
       
   127                                            | LCons y z => Inr(f(y), z))"
       
   128 
       
   129   iterates_def  "iterates f a == llist_corec a (%x. Inr((x, f(x))))"     
       
   130 
   116 
   131 (*Append generates its result by applying f, where
   117 (*Append generates its result by applying f, where
   132     f((NIL,NIL))          = Inl(())
   118     f((NIL,NIL))          = None
   133     f((NIL, CONS N1 N2))  = Inr((N1, (NIL,N2))
   119     f((NIL, CONS N1 N2))  = Some((N1, (NIL,N2))
   134     f((CONS M1 M2, N))    = Inr((M1, (M2,N))
   120     f((CONS M1 M2, N))    = Some((M1, (M2,N))
   135 *)
   121 *)
   136 
   122 
   137   Lappend_def
   123   Lappend    :: ['a item, 'a item] => 'a item
   138    "Lappend M N == LList_corec (M,N)                                         
   124    "Lappend M N == LList_corec (M,N)                                         
   139      (split(List_case (List_case (Inl ()) (%N1 N2. Inr((N1, (NIL,N2))))) 
   125      (split(List_case (List_case None (%N1 N2. Some((N1, (NIL,N2))))) 
   140                       (%M1 M2 N. Inr((M1, (M2,N))))))"
   126                       (%M1 M2 N. Some((M1, (M2,N))))))"
   141 
   127 
   142   lappend_def
   128   lappend    :: ['a llist, 'a llist] => 'a llist
   143    "lappend l n == llist_corec (l,n)                                         
   129     "lappend l n == llist_corec (l,n)                                         
   144    (split(llist_case (llist_case (Inl ()) (%n1 n2. Inr((n1, (LNil,n2))))) 
   130        (split(llist_case (llist_case None (%n1 n2. Some((n1, (LNil,n2))))) 
   145                      (%l1 l2 n. Inr((l1, (l2,n))))))"
   131                          (%l1 l2 n. Some((l1, (l2,n))))))"
   146 
       
   147 rules
       
   148     (*faking a type definition...*)
       
   149   Rep_llist         "Rep_llist(xs): llist(range(Leaf))"
       
   150   Rep_llist_inverse "Abs_llist(Rep_llist(xs)) = xs"
       
   151   Abs_llist_inverse "M: llist(range(Leaf)) ==> Rep_llist(Abs_llist(M)) = M"
       
   152 
   132 
   153 end
   133 end