| author | huffman |
| Wed, 08 Jun 2005 01:40:39 +0200 | |
| changeset 16319 | 1ff2965cc2e7 |
| parent 16232 | 8a12e11d222b |
| child 17291 | 94f6113fe9ed |
| permissions | -rw-r--r-- |
| 1479 | 1 |
(* Title: HOLCF/ex/Loop.thy |
| 244 | 2 |
ID: $Id$ |
| 1479 | 3 |
Author: Franz Regensburger |
| 244 | 4 |
|
5 |
Theory for a loop primitive like while |
|
6 |
*) |
|
7 |
||
| 16232 | 8 |
Loop = Tr + Fix + |
| 244 | 9 |
|
10 |
consts |
|
| 1479 | 11 |
step :: "('a -> tr)->('a -> 'a)->'a->'a"
|
12 |
while :: "('a -> tr)->('a -> 'a)->'a->'a"
|
|
| 244 | 13 |
|
|
1168
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
regensbu
parents:
1150
diff
changeset
|
14 |
defs |
| 244 | 15 |
|
| 10835 | 16 |
step_def "step == (LAM b g x. If b$x then g$x else x fi)" |
17 |
while_def "while == (LAM b g. fix$(LAM f x. |
|
18 |
If b$x then f$(g$x) else x fi))" |
|
| 244 | 19 |
|
20 |
end |
|
21 |