| author | huffman | 
| Mon, 10 Oct 2005 05:30:02 +0200 | |
| changeset 17816 | 9942c5ed866a | 
| parent 17291 | 94f6113fe9ed | 
| child 19742 | 86f21beabafc | 
| permissions | -rw-r--r-- | 
(* Title: HOLCF/ex/Loop.thy ID: $Id$ Author: Franz Regensburger *) header {* Theory for a loop primitive like while *} theory Loop imports HOLCF begin constdefs step :: "('a -> tr)->('a -> 'a)->'a->'a" "step == (LAM b g x. If b$x then g$x else x fi)" while :: "('a -> tr)->('a -> 'a)->'a->'a" "while == (LAM b g. fix$(LAM f x. If b$x then f$(g$x) else x fi))" ML {* use_legacy_bindings (the_context ()) *} end