How to finish the TLAPS proof for a refinement mapping involving records?

别等时光非礼了梦想. 提交于 2019-12-12 23:12:43

问题


I have some difficulty in proving a refinement mapping involving records. Below are the simplified illustrating TLA specs@github (Note that this post is also in tlaplus-googlegroup, without replies yet.):


SimpleVoting.tla:

It maintains for each participant a maxBal which is a natural number. In IncreaseMaxBal(p, b), maxBal[p] is increased to a larger value b.

---------------------------- MODULE SimpleVoting ----------------------------
EXTENDS Naturals
-----------------------------------------------------------------------------
CONSTANT Participant

VARIABLE maxBal

TypeOK == maxBal \in [Participant -> Nat]
-----------------------------------------------------------------------------
Init == maxBal = [p \in Participant |-> 0]

IncreaseMaxBal(p, b) ==
  /\ maxBal[p] < b
  /\ maxBal' = [maxBal EXCEPT ![p] = b]
-----------------------------------------------------------------------------
Next == \E p \in Participant, b \in Nat : IncreaseMaxBal(p, b)

Spec == Init /\ [][Next]_maxBal
=============================================================================

Record.tla:

It maintains a 2D "array" state, where state[p][q] is the State of q from the view of p and a State is a record: State == [maxBal : Nat, maxVBal : Nat].

In Prepare(p, b), state[p][p].maxBal is increased to a larger value b.

------------------------------- MODULE Record -------------------------------
EXTENDS Naturals, TLAPS
---------------------------------------------------------------------------
CONSTANTS Participant  \* the set of partipants

VARIABLES state \* state[p][q]: the state of q \in Participant from the view of p \in Participant

State == [maxBal: Nat, maxVBal: Nat]

TypeOK == state \in [Participant -> [Participant -> State]]
---------------------------------------------------------------------------
InitState == [maxBal |-> 0, maxVBal |-> 0]

Init == state = [p \in Participant |-> [q \in Participant |-> InitState]] 

Prepare(p, b) == 
    /\ state[p][p].maxBal < b
    /\ state' = [state EXCEPT ![p][p].maxBal = b]
---------------------------------------------------------------------------
Next == \E p \in Participant, b \in Nat : Prepare(p, b)

Spec == Init /\ [][Next]_state
---------------------------------------------------------------------------

Intuitively, Record maintains maxBal[p] of SimpleVoting as state[p][p].maxBal. Therefore, I want to show that Record refines SimpleVoting under the following refinement mapping:

maxBal == [p \in Participant |-> state[p][p].maxBal]

SV == INSTANCE SimpleVoting

However, step <3>2 in the following proof fails.

THEOREM Spec => SV!Spec
  <1>1. Init => SV!Init
    BY DEF Init, SV!Init, maxBal, InitState
  <1>2. [Next]_state => [SV!Next]_maxBal
    <2>1. UNCHANGED state => UNCHANGED maxBal
      BY DEF maxBal
    <2>2. Next => SV!Next
      <3> SUFFICES ASSUME NEW p \in Participant, NEW b \in Nat,
                          Prepare(p, b)
                   PROVE  SV!IncreaseMaxBal(p, b)
        BY DEF Next, SV!Next
      <3>1. maxBal[p] < b
        BY DEF Prepare, maxBal
      <3>2. maxBal' = [maxBal EXCEPT ![p] = b] \* failed here!
        BY DEF Prepare, maxBal
      <3>3. QED
        BY <3>1, <3>2 DEF SV!IncreaseMaxBal
    <2>3. QED
      BY <2>1, <2>2
  <1>3. QED

The obligation at <3>2 is as follows. Isn't the state' = [state EXCEPT ![p] = ...] in the assumption the same as the conclusion [p_1 \in Participant |-> state[p_1][p_1].maxBal]' ...? What is missing? What is wrong with my proof?

ASSUME NEW CONSTANT Participant,
       NEW VARIABLE state,
       NEW CONSTANT p \in Participant,
       NEW CONSTANT b \in Nat,
       /\ state[p][p].maxBal < b
       /\ state'
          = [state EXCEPT
               ![p] = [state[p] EXCEPT
                         ![p] = [state[p][p] EXCEPT !.maxBal = b]]] 
PROVE  [p_1 \in Participant |-> state[p_1][p_1].maxBal]'
       = [[p_1 \in Participant |-> state[p_1][p_1].maxBal] EXCEPT ![p] = b]

来源:https://stackoverflow.com/questions/57534064/how-to-finish-the-tlaps-proof-for-a-refinement-mapping-involving-records

易学教程内所有资源均来自网络或用户发布的内容,如有违反法律规定的内容欢迎反馈
该文章没有解决你所遇到的问题?点击提问,说说你的问题,让更多的人一起探讨吧!