| author | wenzelm |
| Mon, 22 May 2000 16:05:22 +0200 | |
| changeset 8923 | 42fd7abde9aa |
| parent 2642 | 3c3a84cc85a9 |
| child 10835 | f4745d77e620 |
| permissions | -rw-r--r-- |
| 1479 | 1 |
(* Title: HOLCF/ex/Loop.thy |
| 244 | 2 |
ID: $Id$ |
| 1479 | 3 |
Author: Franz Regensburger |
4 |
Copyright 1993 Technische Universitaet Muenchen |
|
| 244 | 5 |
|
6 |
Theory for a loop primitive like while |
|
7 |
*) |
|
8 |
||
| 2642 | 9 |
Loop = Tr + |
| 244 | 10 |
|
11 |
consts |
|
| 1479 | 12 |
step :: "('a -> tr)->('a -> 'a)->'a->'a"
|
13 |
while :: "('a -> tr)->('a -> 'a)->'a->'a"
|
|
| 244 | 14 |
|
|
1168
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
regensbu
parents:
1150
diff
changeset
|
15 |
defs |
| 244 | 16 |
|
| 1479 | 17 |
step_def "step == (LAM b g x. If b`x then g`x else x fi)" |
18 |
while_def "while == (LAM b g. fix`(LAM f x. |
|
|
1168
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
regensbu
parents:
1150
diff
changeset
|
19 |
If b`x then f`(g`x) else x fi))" |
| 244 | 20 |
|
21 |
end |
|
22 |