| author | wenzelm | 
| Tue, 13 Jun 2000 18:33:34 +0200 | |
| changeset 9058 | 7856a01119fb | 
| parent 6823 | 97babc436a41 | 
| child 13797 | baefae13ad37 | 
| 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 | |
| 6535 | 9 | Constrains = UNITY + | 
| 10 | ||
| 11 | consts traces :: "['a set, ('a * 'a)set set] => ('a * 'a list) set"
 | |
| 12 | ||
| 13 | (*Initial states and program => (final state, reversed trace to it)... | |
| 14 | Arguments MUST be curried in an inductive definition*) | |
| 15 | ||
| 16 | inductive "traces init acts" | |
| 17 | intrs | |
| 18 | (*Initial trace is empty*) | |
| 19 | Init "s: init ==> (s,[]) : traces init acts" | |
| 20 | ||
| 21 | Acts "[| act: acts; (s,evs) : traces init acts; (s,s'): act |] | |
| 22 | ==> (s', s#evs) : traces init acts" | |
| 23 | ||
| 24 | ||
| 25 | consts reachable :: "'a program => 'a set" | |
| 26 | ||
| 27 | inductive "reachable F" | |
| 28 | intrs | |
| 29 | Init "s: Init F ==> s : reachable F" | |
| 30 | ||
| 31 | Acts "[| act: Acts F; s : reachable F; (s,s'): act |] | |
| 32 | ==> s' : reachable F" | |
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 33 | |
| 6536 | 34 | consts | 
| 6823 | 35 | Constrains :: "['a set, 'a set] => 'a program set" (infixl "Co" 60) | 
| 36 | op_Unless :: "['a set, 'a set] => 'a program set" (infixl "Unless" 60) | |
| 6536 | 37 | |
| 38 | defs | |
| 39 | Constrains_def | |
| 6575 | 40 |     "A Co B == {F. F : (reachable F Int A)  co  B}"
 | 
| 6536 | 41 | |
| 42 | Unless_def | |
| 43 | "A Unless B == (A-B) Co (A Un B)" | |
| 44 | ||
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 45 | constdefs | 
| 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 46 | |
| 5648 | 47 | Stable :: "'a set => 'a program set" | 
| 6536 | 48 | "Stable A == A Co A" | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 49 | |
| 6570 | 50 | (*Always is the weak form of "invariant"*) | 
| 51 | Always :: "'a set => 'a program set" | |
| 52 |     "Always A == {F. Init F <= A} Int Stable A"
 | |
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 53 | |
| 5784 | 54 | (*Polymorphic in both states and the meaning of <= *) | 
| 6705 | 55 |   Increasing :: "['a => 'b::{order}] => 'a program set"
 | 
| 5784 | 56 |     "Increasing f == INT z. Stable {s. z <= f s}"
 | 
| 57 | ||
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: diff
changeset | 58 | end |