src/HOLCF/ex/Stream.thy
author wenzelm
Fri, 06 Oct 2000 17:35:58 +0200
changeset 10168 50be659d4222
parent 9169 85a47aa21f74
child 11348 e08a0855af67
permissions -rw-r--r--
final tuning;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
9169
85a47aa21f74 tidying and unbatchifying
paulson
parents: 4122
diff changeset
     1
(*  Title: 	HOLCF/ex//Stream.thy
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
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
     4
    Copyright   1993, 1995 Technische Universitaet Muenchen
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
     5
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
     6
general Stream domain
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
     7
*)
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
     8
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
     9
Stream = HOLCF + 
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    10
4122
f63c283cefaf * removed "axioms" and "generated by" section
oheimb
parents: 2570
diff changeset
    11
domain 'a stream = "&&" (ft::'a) (lazy rt::'a stream) (infixr 65)
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    12
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    13
end
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    14
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    15