src/HOL/UNITY/Token.thy
author paulson
Fri, 12 May 2000 15:05:02 +0200
changeset 8862 78643f8449c6
parent 8703 816d8f6513be
permissions -rw-r--r--
NatSimprocs is now a theory, not a file
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/Token
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     2
    ID:         $Id$
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     4
    Copyright   1998  University of Cambridge
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     5
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     6
The Token Ring.
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     7
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     8
From Misra, "A Logic for Concurrent Programming" (1994), sections 5.2 and 13.2.
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     9
*)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    10
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    11
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents: 5253
diff changeset
    12
Token = WFair + 
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    13
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    14
(*process states*)
5232
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 4776
diff changeset
    15
datatype pstate = Hungry | Eating | Thinking
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    16
5232
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 4776
diff changeset
    17
record state =
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 4776
diff changeset
    18
  token :: nat
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 4776
diff changeset
    19
  proc  :: nat => pstate
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    20
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    21
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    22
constdefs
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    23
  HasTok :: nat => state set
5232
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 4776
diff changeset
    24
    "HasTok i == {s. token s = i}"
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    25
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    26
  H :: nat => state set
5232
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 4776
diff changeset
    27
    "H i == {s. proc s i = Hungry}"
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    28
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    29
  E :: nat => state set
5232
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 4776
diff changeset
    30
    "E i == {s. proc s i = Eating}"
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    31
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    32
  T :: nat => state set
5232
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 4776
diff changeset
    33
    "T i == {s. proc s i = Thinking}"
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    34
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    35
5420
b48ab3281944 New approach, using a locale
paulson
parents: 5277
diff changeset
    36
locale Token =
b48ab3281944 New approach, using a locale
paulson
parents: 5277
diff changeset
    37
  fixes
b48ab3281944 New approach, using a locale
paulson
parents: 5277
diff changeset
    38
    N         :: nat	 (*number of nodes in the ring*)
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5608
diff changeset
    39
    F         :: state program
5420
b48ab3281944 New approach, using a locale
paulson
parents: 5277
diff changeset
    40
    nodeOrder :: "nat => (nat*nat)set"
b48ab3281944 New approach, using a locale
paulson
parents: 5277
diff changeset
    41
    next      :: nat => nat
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    42
5420
b48ab3281944 New approach, using a locale
paulson
parents: 5277
diff changeset
    43
  assumes
b48ab3281944 New approach, using a locale
paulson
parents: 5277
diff changeset
    44
    N_positive "0<N"
b48ab3281944 New approach, using a locale
paulson
parents: 5277
diff changeset
    45
6536
281d44905cab made many specification operators infix
paulson
parents: 5931
diff changeset
    46
    TR2  "F : (T i) co (T i Un H i)"
5420
b48ab3281944 New approach, using a locale
paulson
parents: 5277
diff changeset
    47
6536
281d44905cab made many specification operators infix
paulson
parents: 5931
diff changeset
    48
    TR3  "F : (H i) co (H i Un E i)"
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    49
6536
281d44905cab made many specification operators infix
paulson
parents: 5931
diff changeset
    50
    TR4  "F : (H i - HasTok i) co (H i)"
5420
b48ab3281944 New approach, using a locale
paulson
parents: 5277
diff changeset
    51
6536
281d44905cab made many specification operators infix
paulson
parents: 5931
diff changeset
    52
    TR5  "F : (HasTok i) co (HasTok i Un -(E i))"
5420
b48ab3281944 New approach, using a locale
paulson
parents: 5277
diff changeset
    53
6536
281d44905cab made many specification operators infix
paulson
parents: 5931
diff changeset
    54
    TR6  "F : (H i Int HasTok i) leadsTo (E i)"
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    55
6536
281d44905cab made many specification operators infix
paulson
parents: 5931
diff changeset
    56
    TR7  "F : (HasTok i) leadsTo (HasTok (next i))"
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    57
5420
b48ab3281944 New approach, using a locale
paulson
parents: 5277
diff changeset
    58
  defines
b48ab3281944 New approach, using a locale
paulson
parents: 5277
diff changeset
    59
    nodeOrder_def
b48ab3281944 New approach, using a locale
paulson
parents: 5277
diff changeset
    60
      "nodeOrder j == (inv_image less_than (%i. ((j+N)-i) mod N))  Int
8703
816d8f6513be Times -> <*>
nipkow
parents: 6536
diff changeset
    61
		      (lessThan N <*> lessThan N)"
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    62
5420
b48ab3281944 New approach, using a locale
paulson
parents: 5277
diff changeset
    63
    next_def
b48ab3281944 New approach, using a locale
paulson
parents: 5277
diff changeset
    64
      "next i == (Suc i) mod N"
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    65
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    66
end