| author | wenzelm | 
| Thu, 27 Aug 1998 16:54:55 +0200 | |
| changeset 5393 | 7299e531d481 | 
| parent 5313 | 1861a564d7e2 | 
| child 5620 | 3ac11c4af76a | 
| permissions | -rw-r--r-- | 
| 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 | |
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 13 | Constrains :: "['a program, 'a set, 'a set] => bool" | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 14 | "Constrains prg A B == | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 15 | constrains (Acts prg) | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 16 | (reachable prg Int A) | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 17 | (reachable prg Int B)" | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 18 | |
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 19 | Stable :: "'a program => 'a set => bool" | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 20 | "Stable prg A == Constrains prg A A" | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 21 | |
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 22 | Unless :: "['a program, 'a set, 'a set] => bool" | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 23 | "Unless prg A B == Constrains prg (A-B) (A Un B)" | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 24 | |
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 25 | Invariant :: "['a program, 'a set] => bool" | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 26 | "Invariant prg A == (Init prg) <= A & Stable prg A" | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 27 | |
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 28 | end |