| author | oheimb |
| Sat, 15 Feb 1997 17:55:11 +0100 | |
| changeset 2638 | 6c6a44b5f757 |
| parent 1479 | 21eb5e156d91 |
| child 2642 | 3c3a84cc85a9 |
| 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 |
||
9 |
Loop = Tr2 + |
|
10 |
||
11 |
consts |
|
12 |
||
| 1479 | 13 |
step :: "('a -> tr)->('a -> 'a)->'a->'a"
|
14 |
while :: "('a -> tr)->('a -> 'a)->'a->'a"
|
|
| 244 | 15 |
|
|
1168
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
regensbu
parents:
1150
diff
changeset
|
16 |
defs |
| 244 | 17 |
|
| 1479 | 18 |
step_def "step == (LAM b g x. If b`x then g`x else x fi)" |
19 |
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
|
20 |
If b`x then f`(g`x) else x fi))" |
| 244 | 21 |
|
22 |
end |
|
23 |