src/HOL/UNITY/Constrains.thy
author oheimb
Fri, 11 Dec 1998 18:56:30 +0100
changeset 6027 9dd06eeda95c
parent 5784 54276fba8420
child 6535 880f31a62784
permissions -rw-r--r--
added new print_mode "xsymbols" for extended symbol support (e.g. genuiely long arrows)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/Constrains
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
     2
    ID:         $Id$
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
     4
    Copyright   1998  University of Cambridge
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
     5
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
     6
Safety relations: restricted to the set of reachable states.
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
     7
*)
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
     8
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
     9
Constrains = UNITY + Traces + 
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    10
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    11
constdefs
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    12
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5620
diff changeset
    13
  Constrains :: "['a set, 'a set] => 'a program set"
fe887910e32e specifications as sets of programs
paulson
parents: 5620
diff changeset
    14
    "Constrains A B == {F. F : constrains (reachable F  Int  A)
fe887910e32e specifications as sets of programs
paulson
parents: 5620
diff changeset
    15
  		                          (reachable F  Int  B)}"
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    16
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5620
diff changeset
    17
  Stable     :: "'a set => 'a program set"
fe887910e32e specifications as sets of programs
paulson
parents: 5620
diff changeset
    18
    "Stable A == Constrains A A"
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    19
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5620
diff changeset
    20
  Unless :: "['a set, 'a set] => 'a program set"
fe887910e32e specifications as sets of programs
paulson
parents: 5620
diff changeset
    21
    "Unless A B == Constrains (A-B) (A Un B)"
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    22
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5620
diff changeset
    23
  Invariant :: "'a set => 'a program set"
fe887910e32e specifications as sets of programs
paulson
parents: 5620
diff changeset
    24
    "Invariant A == {F. Init F <= A} Int Stable A"
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    25
5784
54276fba8420 the Increasing operator
paulson
parents: 5648
diff changeset
    26
  (*Polymorphic in both states and the meaning of <= *)
54276fba8420 the Increasing operator
paulson
parents: 5648
diff changeset
    27
  Increasing :: "['a => 'b::{ord}] => 'a program set"
54276fba8420 the Increasing operator
paulson
parents: 5648
diff changeset
    28
    "Increasing f == INT z. Stable {s. z <= f s}"
54276fba8420 the Increasing operator
paulson
parents: 5648
diff changeset
    29
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    30
end