src/HOLCF/ex/Stream.thy
author wenzelm
Wed, 09 Jun 2004 18:52:42 +0200
changeset 14898 a25550451b51
parent 14535 7cb26928e70d
child 14981 e73f8140af78
permissions -rw-r--r--
Url.File;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11348
e08a0855af67 added stream length, map, and filter
oheimb
parents: 9169
diff changeset
     1
(*  Title: 	HOLCF/ex/Stream.thy
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 4122
diff changeset
     2
    ID:         $Id$
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
     3
    Author: 	Franz Regensburger, David von Oheimb
12036
wenzelm
parents: 11348
diff changeset
     4
    License:    GPL (GNU GENERAL PUBLIC LICENSE)
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
     5
12036
wenzelm
parents: 11348
diff changeset
     6
General Stream domain.
14535
7cb26928e70d added Streams.thy (with stream concatenation etc.)
oheimb
parents: 14171
diff changeset
     7
TODO: should be integrated with HOLCF/Streams
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
     8
*)
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
     9
11348
e08a0855af67 added stream length, map, and filter
oheimb
parents: 9169
diff changeset
    10
Stream = HOLCF + Nat_Infinity +
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    11
4122
f63c283cefaf * removed "axioms" and "generated by" section
oheimb
parents: 2570
diff changeset
    12
domain 'a stream = "&&" (ft::'a) (lazy rt::'a stream) (infixr 65)
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    13
11348
e08a0855af67 added stream length, map, and filter
oheimb
parents: 9169
diff changeset
    14
consts
e08a0855af67 added stream length, map, and filter
oheimb
parents: 9169
diff changeset
    15
12036
wenzelm
parents: 11348
diff changeset
    16
  smap		:: "('a -> 'b) -> 'a stream -> 'b stream"
wenzelm
parents: 11348
diff changeset
    17
  sfilter	:: "('a -> tr) -> 'a stream -> 'a stream"
wenzelm
parents: 11348
diff changeset
    18
  slen		:: "'a stream => inat"			("#_" [1000] 1000)
11348
e08a0855af67 added stream length, map, and filter
oheimb
parents: 9169
diff changeset
    19
e08a0855af67 added stream length, map, and filter
oheimb
parents: 9169
diff changeset
    20
defs
e08a0855af67 added stream length, map, and filter
oheimb
parents: 9169
diff changeset
    21
14171
0cab06e3bbd0 Extended the notion of letter and digit, such that now one may use greek,
skalberg
parents: 12036
diff changeset
    22
  smap_def	"smap	 == fix\\<cdot>(\\<Lambda> h f s. case s of x && xs => f\\<cdot>x && h\\<cdot>f\\<cdot>xs)"
0cab06e3bbd0 Extended the notion of letter and digit, such that now one may use greek,
skalberg
parents: 12036
diff changeset
    23
  sfilter_def	"sfilter == fix\\<cdot>(\\<Lambda> h p s. case s of x && xs => 
11348
e08a0855af67 added stream length, map, and filter
oheimb
parents: 9169
diff changeset
    24
				     If p\\<cdot>x then x && h\\<cdot>p\\<cdot>xs else h\\<cdot>p\\<cdot>xs fi)"
12036
wenzelm
parents: 11348
diff changeset
    25
  slen_def	"#s == if stream_finite s 
11348
e08a0855af67 added stream length, map, and filter
oheimb
parents: 9169
diff changeset
    26
		      then Fin (LEAST n. stream_take n\\<cdot>s = s) else \\<infinity>"
e08a0855af67 added stream length, map, and filter
oheimb
parents: 9169
diff changeset
    27
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    28
end
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    29
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    30