src/HOLCF/explicit_domains/Stream.thy
changeset 1274 ea0668a1c0ba
child 1479 21eb5e156d91
equal deleted inserted replaced
1273:6960ec882bca 1274:ea0668a1c0ba
       
     1 (* 
       
     2     ID:         $Id$
       
     3     Author: 	Franz Regensburger
       
     4     Copyright   1993 Technische Universitaet Muenchen
       
     5 
       
     6 Theory for streams without defined empty stream 
       
     7   'a stream = 'a ** ('a stream)u
       
     8 
       
     9 The type is axiomatized as the least solution of the domain equation above.
       
    10 The functor term that specifies the domain equation is: 
       
    11 
       
    12   FT = <**,K_{'a},U>
       
    13 
       
    14 For details see chapter 5 of:
       
    15 
       
    16 [Franz Regensburger] HOLCF: Eine konservative Erweiterung von HOL um LCF,
       
    17                      Dissertation, Technische Universit"at M"unchen, 1994
       
    18 *)
       
    19 
       
    20 Stream = Dnat2 +
       
    21 
       
    22 types stream 1
       
    23 
       
    24 (* ----------------------------------------------------------------------- *)
       
    25 (* arity axiom is validated by semantic reasoning                          *)
       
    26 (* partial ordering is implicit in the isomorphism axioms and their cont.  *)
       
    27 
       
    28 arities stream::(pcpo)pcpo
       
    29 
       
    30 consts
       
    31 
       
    32 (* ----------------------------------------------------------------------- *)
       
    33 (* essential constants                                                     *)
       
    34 
       
    35 stream_rep	:: "('a stream) -> ('a ** ('a stream)u)"
       
    36 stream_abs	:: "('a ** ('a stream)u) -> ('a stream)"
       
    37 
       
    38 (* ----------------------------------------------------------------------- *)
       
    39 (* abstract constants and auxiliary constants                              *)
       
    40 
       
    41 stream_copy	:: "('a stream -> 'a stream) ->'a stream -> 'a stream"
       
    42 
       
    43 scons		:: "'a -> 'a stream -> 'a stream"
       
    44 stream_when	:: "('a -> 'a stream -> 'b) -> 'a stream -> 'b"
       
    45 is_scons	:: "'a stream -> tr"
       
    46 shd		:: "'a stream -> 'a"
       
    47 stl		:: "'a stream -> 'a stream"
       
    48 stream_take	:: "nat => 'a stream -> 'a stream"
       
    49 stream_finite	:: "'a stream => bool"
       
    50 stream_bisim	:: "('a stream => 'a stream => bool) => bool"
       
    51 
       
    52 rules
       
    53 
       
    54 (* ----------------------------------------------------------------------- *)
       
    55 (* axiomatization of recursive type 'a stream                              *)
       
    56 (* ----------------------------------------------------------------------- *)
       
    57 (* ('a stream,stream_abs) is the initial F-algebra where                   *)
       
    58 (* F is the locally continuous functor determined by functor term FT.      *)
       
    59 (* domain equation: 'a stream = 'a ** ('a stream)u                         *)
       
    60 (* functor term:    FT = <**,K_{'a},U>                                     *)
       
    61 (* ----------------------------------------------------------------------- *)
       
    62 (* stream_abs is an isomorphism with inverse stream_rep                    *)
       
    63 (* identity is the least endomorphism on 'a stream                         *)
       
    64 
       
    65 stream_abs_iso	"stream_rep`(stream_abs`x) = x"
       
    66 stream_rep_iso	"stream_abs`(stream_rep`x) = x"
       
    67 stream_copy_def	"stream_copy == (LAM f. stream_abs oo 
       
    68  		(ssplit`(LAM x y. (|x , (lift`(up oo f))`y|) )) oo stream_rep)"
       
    69 stream_reach	"(fix`stream_copy)`x = x"
       
    70 
       
    71 defs
       
    72 (* ----------------------------------------------------------------------- *)
       
    73 (* properties of additional constants                                      *)
       
    74 (* ----------------------------------------------------------------------- *)
       
    75 (* constructors                                                            *)
       
    76 
       
    77 scons_def	"scons == (LAM x l. stream_abs`(| x, up`l |))"
       
    78 
       
    79 (* ----------------------------------------------------------------------- *)
       
    80 (* discriminator functional                                                *)
       
    81 
       
    82 stream_when_def 
       
    83 "stream_when == (LAM f l.ssplit `(LAM x l.f`x`(lift`ID`l)) `(stream_rep`l))"
       
    84 
       
    85 (* ----------------------------------------------------------------------- *)
       
    86 (* discriminators and selectors                                            *)
       
    87 
       
    88 is_scons_def	"is_scons == stream_when`(LAM x l.TT)"
       
    89 shd_def		"shd == stream_when`(LAM x l.x)"
       
    90 stl_def		"stl == stream_when`(LAM x l.l)"
       
    91 
       
    92 (* ----------------------------------------------------------------------- *)
       
    93 (* the taker for streams                                                   *)
       
    94 
       
    95 stream_take_def "stream_take == (%n.iterate n stream_copy UU)"
       
    96 
       
    97 (* ----------------------------------------------------------------------- *)
       
    98 
       
    99 stream_finite_def	"stream_finite == (%s.? n.stream_take n `s=s)"
       
   100 
       
   101 (* ----------------------------------------------------------------------- *)
       
   102 (* definition of bisimulation is determined by domain equation             *)
       
   103 (* simplification and rewriting for abstract constants yields def below    *)
       
   104 
       
   105 stream_bisim_def "stream_bisim ==
       
   106 (%R.!s1 s2.
       
   107  	R s1 s2 -->
       
   108   ((s1=UU & s2=UU) |
       
   109   (? x s11 s21. x~=UU & s1=scons`x`s11 & s2 = scons`x`s21 & R s11 s21)))"
       
   110 
       
   111 end
       
   112 
       
   113 
       
   114 
       
   115