src/HOL/UNITY/FP.thy
author paulson
Thu, 26 Aug 1999 11:33:24 +0200
changeset 7359 98a2afab3f86
parent 5648 fe887910e32e
child 13796 19f50fa807ae
permissions -rw-r--r--
extra syntax for JN, making it more like UN
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/FP
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     2
    ID:         $Id$
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     4
    Copyright   1998  University of Cambridge
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     5
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     6
Fixed Point of a Program
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     7
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     8
From Misra, "A Logic for Concurrent Programming", 1994
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     9
*)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    10
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    11
FP = UNITY +
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    12
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    13
constdefs
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    14
5648
fe887910e32e specifications as sets of programs
paulson
parents: 4776
diff changeset
    15
  FP_Orig :: "'a program => 'a set"
fe887910e32e specifications as sets of programs
paulson
parents: 4776
diff changeset
    16
    "FP_Orig F == Union{A. ALL B. F : stable (A Int B)}"
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    17
5648
fe887910e32e specifications as sets of programs
paulson
parents: 4776
diff changeset
    18
  FP :: "'a program => 'a set"
fe887910e32e specifications as sets of programs
paulson
parents: 4776
diff changeset
    19
    "FP F == {s. F : stable {s}}"
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    20
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    21
end