| author | berghofe |
| Tue, 07 Aug 2001 19:26:42 +0200 | |
| changeset 11470 | d3a3be6660f9 |
| parent 10835 | f4745d77e620 |
| child 12036 | 49f6c49454c2 |
| 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