src/HOLCF/dlist.thy
author clasohm
Tue, 07 Nov 1995 13:15:04 +0100
changeset 1323 ae24fa249266
parent 298 3a0485439396
permissions -rw-r--r--
added leading "." to HTML filenames
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
298
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
     1
(*  Title: 	HOLCF/dlist.thy
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
     2
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
     3
    Author: 	Franz Regensburger
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
     4
    ID:         $ $
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
     5
    Copyright   1994 Technische Universitaet Muenchen
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
     6
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
     7
Theory for lists
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
     8
*)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
     9
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    10
Dlist = Stream2 +
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    11
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    12
types dlist 1
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    13
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    14
(* ----------------------------------------------------------------------- *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    15
(* arity axiom is validated by semantic reasoning                          *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    16
(* partial ordering is implicit in the isomorphism axioms and their cont.  *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    17
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    18
arities dlist::(pcpo)pcpo
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    19
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    20
consts
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    21
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    22
(* ----------------------------------------------------------------------- *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    23
(* essential constants                                                     *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    24
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    25
dlist_rep	:: "('a dlist) -> (one ++ 'a ** 'a dlist)"
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    26
dlist_abs	:: "(one ++ 'a ** 'a dlist) -> ('a dlist)"
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    27
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    28
(* ----------------------------------------------------------------------- *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    29
(* abstract constants and auxiliary constants                              *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    30
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    31
dlist_copy	:: "('a dlist -> 'a dlist) ->'a dlist -> 'a dlist"
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    32
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    33
dnil            :: "'a dlist"
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    34
dcons		:: "'a -> 'a dlist -> 'a dlist"
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    35
dlist_when	:: " 'b -> ('a -> 'a dlist -> 'b) -> 'a dlist -> 'b"
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    36
is_dnil    	:: "'a dlist -> tr"
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    37
is_dcons	:: "'a dlist -> tr"
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    38
dhd		:: "'a dlist -> 'a"
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    39
dtl		:: "'a dlist -> 'a dlist"
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    40
dlist_take	:: "nat => 'a dlist -> 'a dlist"
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    41
dlist_finite	:: "'a dlist => bool"
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    42
dlist_bisim	:: "('a dlist => 'a dlist => bool) => bool"
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    43
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    44
rules
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    45
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    46
(* ----------------------------------------------------------------------- *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    47
(* axiomatization of recursive type 'a dlist                               *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    48
(* ----------------------------------------------------------------------- *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    49
(* ('a dlist,dlist_abs) is the initial F-algebra where                     *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    50
(* F is the locally continuous functor determined by domain equation       *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    51
(* X = one ++ 'a ** X                                                      *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    52
(* ----------------------------------------------------------------------- *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    53
(* dlist_abs is an isomorphism with inverse dlist_rep                      *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    54
(* identity is the least endomorphism on 'a dlist                          *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    55
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    56
dlist_abs_iso	"dlist_rep[dlist_abs[x]] = x"
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    57
dlist_rep_iso	"dlist_abs[dlist_rep[x]] = x"
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    58
dlist_copy_def	"dlist_copy == (LAM f. dlist_abs oo \
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    59
\ 		(when[sinl][sinr oo (ssplit[LAM x y. x ## f[y]])])\
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    60
\                                oo dlist_rep)"
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    61
dlist_reach	"(fix[dlist_copy])[x]=x"
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    62
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    63
(* ----------------------------------------------------------------------- *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    64
(* properties of additional constants                                      *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    65
(* ----------------------------------------------------------------------- *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    66
(* constructors                                                            *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    67
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    68
dnil_def	"dnil  == dlist_abs[sinl[one]]"
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    69
dcons_def	"dcons == (LAM x l. dlist_abs[sinr[x##l]])"
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    70
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    71
(* ----------------------------------------------------------------------- *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    72
(* discriminator functional                                                *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    73
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    74
dlist_when_def 
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    75
"dlist_when == (LAM f1 f2 l.\
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    76
\   when[LAM x.f1][ssplit[LAM x l.f2[x][l]]][dlist_rep[l]])"
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    77
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    78
(* ----------------------------------------------------------------------- *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    79
(* discriminators and selectors                                            *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    80
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    81
is_dnil_def	"is_dnil  == dlist_when[TT][LAM x l.FF]"
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    82
is_dcons_def	"is_dcons == dlist_when[FF][LAM x l.TT]"
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    83
dhd_def		"dhd == dlist_when[UU][LAM x l.x]"
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    84
dtl_def		"dtl == dlist_when[UU][LAM x l.l]"
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    85
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    86
(* ----------------------------------------------------------------------- *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    87
(* the taker for dlists                                                   *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    88
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    89
dlist_take_def "dlist_take == (%n.iterate(n,dlist_copy,UU))"
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    90
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    91
(* ----------------------------------------------------------------------- *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    92
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    93
dlist_finite_def	"dlist_finite == (%s.? n.dlist_take(n)[s]=s)"
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    94
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    95
(* ----------------------------------------------------------------------- *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    96
(* definition of bisimulation is determined by domain equation             *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    97
(* simplification and rewriting for abstract constants yields def below    *)
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    98
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
    99
dlist_bisim_def "dlist_bisim ==\
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
   100
\ ( %R.!l1 l2.\
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
   101
\ 	R(l1,l2) -->\
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
   102
\  ((l1=UU & l2=UU) |\
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
   103
\   (l1=dnil & l2=dnil) |\
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
   104
\   (? x l11 l21. x~=UU & l11~=UU & l21~=UU & \
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
   105
\               l1=dcons[x][l11] & l2 = dcons[x][l21] & R(l11,l21))))"
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
   106
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
   107
end
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
   108
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
   109
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
   110
3a0485439396 structural induction for strict lists
nipkow
parents:
diff changeset
   111