src/HOL/UNITY/UNITY.thy
author wenzelm
Fri, 09 Nov 2001 00:09:47 +0100
changeset 12114 a8e860c86252
parent 10834 a7897aebbffc
child 13797 baefae13ad37
permissions -rw-r--r--
eliminated old "symbols" syntax, use "xsymbols" instead;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/UNITY
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
The basic UNITY theory (revised version, based upon the "co" operator)
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
8948
b797cfa3548d restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
paulson
parents: 6823
diff changeset
    11
UNITY = Main + 
6535
880f31a62784 eliminated theory UNITY/Traces
paulson
parents: 5804
diff changeset
    12
880f31a62784 eliminated theory UNITY/Traces
paulson
parents: 5804
diff changeset
    13
typedef (Program)
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 8948
diff changeset
    14
  'a program = "{(init:: 'a set, acts :: ('a * 'a)set set,
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 8948
diff changeset
    15
		  allowed :: ('a * 'a)set set). Id:acts & Id: allowed}"
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    16
6536
281d44905cab made many specification operators infix
paulson
parents: 6535
diff changeset
    17
consts
6823
97babc436a41 new-style infix directives
paulson
parents: 6713
diff changeset
    18
  constrains :: "['a set, 'a set] => 'a program set"  (infixl "co"     60)
97babc436a41 new-style infix directives
paulson
parents: 6713
diff changeset
    19
  op_unless  :: "['a set, 'a set] => 'a program set"  (infixl "unless" 60)
6536
281d44905cab made many specification operators infix
paulson
parents: 6535
diff changeset
    20
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    21
constdefs
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 8948
diff changeset
    22
    mk_program :: "('a set * ('a * 'a)set set * ('a * 'a)set set)
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 8948
diff changeset
    23
		   => 'a program"
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 8948
diff changeset
    24
    "mk_program == %(init, acts, allowed).
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 8948
diff changeset
    25
                      Abs_Program (init, insert Id acts, insert Id allowed)"
6535
880f31a62784 eliminated theory UNITY/Traces
paulson
parents: 5804
diff changeset
    26
880f31a62784 eliminated theory UNITY/Traces
paulson
parents: 5804
diff changeset
    27
  Init :: "'a program => 'a set"
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 8948
diff changeset
    28
    "Init F == (%(init, acts, allowed). init) (Rep_Program F)"
6535
880f31a62784 eliminated theory UNITY/Traces
paulson
parents: 5804
diff changeset
    29
880f31a62784 eliminated theory UNITY/Traces
paulson
parents: 5804
diff changeset
    30
  Acts :: "'a program => ('a * 'a)set set"
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 8948
diff changeset
    31
    "Acts F == (%(init, acts, allowed). acts) (Rep_Program F)"
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 8948
diff changeset
    32
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 8948
diff changeset
    33
  AllowedActs :: "'a program => ('a * 'a)set set"
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 8948
diff changeset
    34
    "AllowedActs F == (%(init, acts, allowed). allowed) (Rep_Program F)"
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 8948
diff changeset
    35
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 8948
diff changeset
    36
  Allowed :: "'a program => 'a program set"
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 8948
diff changeset
    37
    "Allowed F == {G. Acts G <= AllowedActs F}"
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    38
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5253
diff changeset
    39
  stable     :: "'a set => 'a program set"
6536
281d44905cab made many specification operators infix
paulson
parents: 6535
diff changeset
    40
    "stable A == A co A"
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    41
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5253
diff changeset
    42
  strongest_rhs :: "['a program, 'a set] => 'a set"
6536
281d44905cab made many specification operators infix
paulson
parents: 6535
diff changeset
    43
    "strongest_rhs F A == Inter {B. F : A co B}"
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    44
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5253
diff changeset
    45
  invariant :: "'a set => 'a program set"
fe887910e32e specifications as sets of programs
paulson
parents: 5253
diff changeset
    46
    "invariant A == {F. Init F <= A} Int stable A"
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    47
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5253
diff changeset
    48
  (*Polymorphic in both states and the meaning of <= *)
6713
614a76ce9bc6 increasing makes sense only for partial orderings
paulson
parents: 6536
diff changeset
    49
  increasing :: "['a => 'b::{order}] => 'a program set"
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5253
diff changeset
    50
    "increasing f == INT z. stable {s. z <= f s}"
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    51
6536
281d44905cab made many specification operators infix
paulson
parents: 6535
diff changeset
    52
281d44905cab made many specification operators infix
paulson
parents: 6535
diff changeset
    53
defs
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10797
diff changeset
    54
  constrains_def "A co B == {F. ALL act: Acts F. act``A <= B}"
6536
281d44905cab made many specification operators infix
paulson
parents: 6535
diff changeset
    55
281d44905cab made many specification operators infix
paulson
parents: 6535
diff changeset
    56
  unless_def     "A unless B == (A-B) co (A Un B)"
281d44905cab made many specification operators infix
paulson
parents: 6535
diff changeset
    57
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    58
end