| author | wenzelm | 
| Thu, 08 Jan 1998 18:19:48 +0100 | |
| changeset 4538 | 0f40d6e7897d | 
| parent 2642 | 3c3a84cc85a9 | 
| child 10835 | f4745d77e620 | 
| permissions | -rw-r--r-- | 
(* Title: HOLCF/ex/Loop.thy ID: $Id$ Author: Franz Regensburger Copyright 1993 Technische Universitaet Muenchen Theory for a loop primitive like while *) Loop = Tr + consts step :: "('a -> tr)->('a -> 'a)->'a->'a" while :: "('a -> tr)->('a -> 'a)->'a->'a" defs step_def "step == (LAM b g x. If b`x then g`x else x fi)" while_def "while == (LAM b g. fix`(LAM f x. If b`x then f`(g`x) else x fi))" end