src/ZF/UNITY/Mutex.thy
author paulson
Tue, 27 May 2003 11:39:03 +0200
changeset 14046 6616e6c53d48
parent 12195 ed2893765a08
child 15634 bca33c49b083
permissions -rw-r--r--
updating ZF-UNITY with Sidi's new material
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     1
(*  Title:      ZF/UNITY/Mutex.thy
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     2
    ID:         $Id$
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     3
    Author:     Sidi O Ehmety, Computer Laboratory
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     4
    Copyright   2001  University of Cambridge
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     5
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     6
Based on "A Family of 2-Process Mutual Exclusion Algorithms" by J Misra
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     7
12195
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
     8
Variables' types are introduced globally so that type verification
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
     9
reduces to the usual ZF typechecking: an ill-tyed expressions reduce to the empty set.
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    10
*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    11
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    12
Mutex = SubstAx + 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    13
consts
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    14
  p, m, n, u, v :: i
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    15
  
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    16
translations
14046
6616e6c53d48 updating ZF-UNITY with Sidi's new material
paulson
parents: 12195
diff changeset
    17
  "p" == "Var([0])"
6616e6c53d48 updating ZF-UNITY with Sidi's new material
paulson
parents: 12195
diff changeset
    18
  "m" == "Var([1])"
6616e6c53d48 updating ZF-UNITY with Sidi's new material
paulson
parents: 12195
diff changeset
    19
  "n" == "Var([0,0])"
6616e6c53d48 updating ZF-UNITY with Sidi's new material
paulson
parents: 12195
diff changeset
    20
  "u" == "Var([0,1])"
6616e6c53d48 updating ZF-UNITY with Sidi's new material
paulson
parents: 12195
diff changeset
    21
  "v" == "Var([1,0])"
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    22
  
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    23
rules (** Type declarations  **)
14046
6616e6c53d48 updating ZF-UNITY with Sidi's new material
paulson
parents: 12195
diff changeset
    24
  p_type  "type_of(p)=bool & default_val(p)=0"
6616e6c53d48 updating ZF-UNITY with Sidi's new material
paulson
parents: 12195
diff changeset
    25
  m_type  "type_of(m)=int  & default_val(m)=#0"
6616e6c53d48 updating ZF-UNITY with Sidi's new material
paulson
parents: 12195
diff changeset
    26
  n_type  "type_of(n)=int  & default_val(n)=#0"
6616e6c53d48 updating ZF-UNITY with Sidi's new material
paulson
parents: 12195
diff changeset
    27
  u_type  "type_of(u)=bool & default_val(u)=0"
6616e6c53d48 updating ZF-UNITY with Sidi's new material
paulson
parents: 12195
diff changeset
    28
  v_type  "type_of(v)=bool & default_val(v)=0"
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    29
  
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    30
constdefs
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    31
  (** The program for process U **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    32
   U0 :: i
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    33
    "U0 == {<s,t>:state*state. t = s(u:=1, m:=#1) & s`m = #0}"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    34
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    35
  U1 :: i
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    36
  "U1 == {<s,t>:state*state. t = s(p:= s`v, m:=#2) &  s`m = #1}"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    37
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    38
  U2 :: i
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    39
    "U2 == {<s,t>:state*state. t = s(m:=#3) & s`p=0 & s`m = #2}"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    40
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    41
  U3 :: i
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    42
    "U3 == {<s,t>:state*state. t=s(u:=0, m:=#4) & s`m = #3}"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    43
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    44
  U4 :: i
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    45
    "U4 == {<s,t>:state*state. t = s(p:=1, m:=#0) & s`m = #4}"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    46
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    47
  
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    48
   (** The program for process V **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    49
  
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    50
  V0 :: i
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    51
    "V0 == {<s,t>:state*state. t = s (v:=1, n:=#1) & s`n = #0}"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    52
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    53
  V1 :: i
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    54
    "V1 == {<s,t>:state*state. t = s(p:=not(s`u), n:=#2) & s`n = #1}"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    55
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    56
  V2 :: i
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    57
    "V2 == {<s,t>:state*state. t  = s(n:=#3) & s`p=1 & s`n = #2}"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    58
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    59
  V3 :: i
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    60
  "V3 == {<s,t>:state*state. t = s (v:=0, n:=#4) & s`n = #3}"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    61
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    62
  V4 :: i
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    63
    "V4 == {<s,t>:state*state. t  = s (p:=0, n:=#0) & s`n = #4}"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    64
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    65
 Mutex :: i
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    66
 "Mutex == mk_program({s:state. s`u=0 & s`v=0 & s`m = #0 & s`n = #0},
14046
6616e6c53d48 updating ZF-UNITY with Sidi's new material
paulson
parents: 12195
diff changeset
    67
              {U0, U1, U2, U3, U4, V0, V1, V2, V3, V4}, Pow(state*state))"
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    68
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    69
  (** The correct invariants **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    70
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    71
  IU :: i
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    72
    "IU == {s:state. (s`u = 1<->(#1 $<= s`m & s`m $<= #3))
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    73
	             & (s`m = #3 --> s`p=0)}"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    74
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    75
  IV :: i
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    76
    "IV == {s:state. (s`v = 1 <-> (#1 $<= s`n & s`n $<= #3))
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    77
	              & (s`n = #3 --> s`p=1)}"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    78
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    79
  (** The faulty invariant (for U alone) **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    80
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    81
  bad_IU :: i
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    82
    "bad_IU == {s:state. (s`u = 1 <-> (#1 $<= s`m & s`m  $<= #3))&
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    83
                   (#3 $<= s`m & s`m $<= #4 --> s`p=0)}"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    84
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    85
end