author | paulson |
Fri, 16 Jun 2000 13:39:21 +0200 | |
changeset 9083 | b36787a56a1f |
parent 6295 | 351b3c2b0d83 |
child 10064 | 1a77667b21ef |
permissions | -rw-r--r-- |
5252 | 1 |
(* Title: HOL/UNITY/Handshake.thy |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1998 University of Cambridge |
|
5 |
||
6 |
Handshake Protocol |
|
7 |
||
8 |
From Misra, "Asynchronous Compositions of Programs", Section 5.3.2 |
|
9 |
*) |
|
10 |
||
11 |
Handshake = Union + |
|
12 |
||
13 |
record state = |
|
14 |
BB :: bool |
|
15 |
NF :: nat |
|
16 |
NG :: nat |
|
17 |
||
18 |
constdefs |
|
19 |
(*F's program*) |
|
20 |
cmdF :: "(state*state) set" |
|
21 |
"cmdF == {(s,s'). s' = s (|NF:= Suc(NF s), BB:=False|) & BB s}" |
|
22 |
||
5648 | 23 |
F :: "state program" |
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6012
diff
changeset
|
24 |
"F == mk_program ({s. NF s = 0 & BB s}, {cmdF})" |
5252 | 25 |
|
26 |
(*G's program*) |
|
27 |
cmdG :: "(state*state) set" |
|
28 |
"cmdG == {(s,s'). s' = s (|NG:= Suc(NG s), BB:=True|) & ~ BB s}" |
|
29 |
||
5648 | 30 |
G :: "state program" |
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6012
diff
changeset
|
31 |
"G == mk_program ({s. NG s = 0 & BB s}, {cmdG})" |
5252 | 32 |
|
33 |
(*the joint invariant*) |
|
34 |
invFG :: "state set" |
|
35 |
"invFG == {s. NG s <= NF s & NF s <= Suc (NG s) & (BB s = (NF s = NG s))}" |
|
36 |
||
37 |
end |