src/HOLCF/ex/Loop.thy
author clasohm
Fri, 24 Nov 1995 11:46:23 +0100
changeset 1367 78bdb2d04771
parent 1274 ea0668a1c0ba
child 1479 21eb5e156d91
permissions -rw-r--r--
fixed make_html bug
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1274
ea0668a1c0ba added 8bit pragmas
regensbu
parents: 1168
diff changeset
     1
(*  Title:	HOLCF/ex/Loop.thy
244
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
     2
    ID:         $Id$
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
     3
    Author: 	Franz Regensburger
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
     4
    Copyright	1993 Technische Universitaet Muenchen
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
     5
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
     6
Theory for a loop primitive like while
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
     7
*)
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
     8
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
     9
Loop = Tr2 +
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    10
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    11
consts
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    12
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    13
	step  :: "('a -> tr)->('a -> 'a)->'a->'a"
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    14
	while :: "('a -> tr)->('a -> 'a)->'a->'a"
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    15
1168
74be52691d62 The curried version of HOLCF is now just called HOLCF. The old
regensbu
parents: 1150
diff changeset
    16
defs
244
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    17
1168
74be52691d62 The curried version of HOLCF is now just called HOLCF. The old
regensbu
parents: 1150
diff changeset
    18
  step_def	"step == (LAM b g x. If b`x then g`x else x fi)"
74be52691d62 The curried version of HOLCF is now just called HOLCF. The old
regensbu
parents: 1150
diff changeset
    19
  while_def	"while == (LAM b g. fix`(LAM f x.
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
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    21
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    22
end
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    23