src/HOLCF/ex/Loop.thy
author nipkow
Sun, 09 Apr 2006 14:20:23 +0200
changeset 19376 529b735edbf2
parent 17291 94f6113fe9ed
child 19742 86f21beabafc
permissions -rw-r--r--
Removed old set interval syntax.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1479
21eb5e156d91 expanded tabs
clasohm
parents: 1274
diff changeset
     1
(*  Title:      HOLCF/ex/Loop.thy
244
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
     2
    ID:         $Id$
1479
21eb5e156d91 expanded tabs
clasohm
parents: 1274
diff changeset
     3
    Author:     Franz Regensburger
244
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
     4
*)
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
     5
17291
94f6113fe9ed converted to Isar theory format;
wenzelm
parents: 16232
diff changeset
     6
header {* Theory for a loop primitive like while *}
244
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
     7
17291
94f6113fe9ed converted to Isar theory format;
wenzelm
parents: 16232
diff changeset
     8
theory Loop
94f6113fe9ed converted to Isar theory format;
wenzelm
parents: 16232
diff changeset
     9
imports HOLCF
94f6113fe9ed converted to Isar theory format;
wenzelm
parents: 16232
diff changeset
    10
begin
244
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    11
17291
94f6113fe9ed converted to Isar theory format;
wenzelm
parents: 16232
diff changeset
    12
constdefs
94f6113fe9ed converted to Isar theory format;
wenzelm
parents: 16232
diff changeset
    13
  step  :: "('a -> tr)->('a -> 'a)->'a->'a"
94f6113fe9ed converted to Isar theory format;
wenzelm
parents: 16232
diff changeset
    14
  "step == (LAM b g x. If b$x then g$x else x fi)"
94f6113fe9ed converted to Isar theory format;
wenzelm
parents: 16232
diff changeset
    15
  while :: "('a -> tr)->('a -> 'a)->'a->'a"
94f6113fe9ed converted to Isar theory format;
wenzelm
parents: 16232
diff changeset
    16
  "while == (LAM b g. fix$(LAM f x.
94f6113fe9ed converted to Isar theory format;
wenzelm
parents: 16232
diff changeset
    17
    If b$x then f$(g$x) else x fi))"
244
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    18
17291
94f6113fe9ed converted to Isar theory format;
wenzelm
parents: 16232
diff changeset
    19
ML {* use_legacy_bindings (the_context ()) *}
244
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    20
929fc2c63bd0 HOLCF examples
nipkow
parents:
diff changeset
    21
end
17291
94f6113fe9ed converted to Isar theory format;
wenzelm
parents: 16232
diff changeset
    22