src/HOL/UNITY/Network.thy
author paulson
Thu, 26 Aug 1999 11:33:24 +0200
changeset 7359 98a2afab3f86
parent 5111 8f4b72f0c15d
permissions -rw-r--r--
extra syntax for JN, making it more like UN
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 4776
diff changeset
     1
(*  Title:      HOL/UNITY/Network
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 4776
diff changeset
     2
    ID:         $Id$
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 4776
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 4776
diff changeset
     4
    Copyright   1998  University of Cambridge
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 4776
diff changeset
     5
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 4776
diff changeset
     6
The Communication Network
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 4776
diff changeset
     7
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 4776
diff changeset
     8
From Misra, "A Logic for Concurrent Programming" (1994), section 5.7
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 4776
diff changeset
     9
*)
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 4776
diff changeset
    10
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    11
Network = UNITY +
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    12
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 4776
diff changeset
    13
(*The state assigns a number to each process variable*)
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 4776
diff changeset
    14
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    15
datatype pvar = Sent | Rcvd | Idle
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    16
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    17
datatype pname = Aproc | Bproc
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    18
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    19
types state = "pname * pvar => nat"
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    20
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    21
end