-
Notifications
You must be signed in to change notification settings - Fork 0
/
spec_counter.v
59 lines (46 loc) · 1.64 KB
/
spec_counter.v
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
From iris.base_logic.lib Require Import invariants.
From smr.program_logic Require Import atomic.
From smr.lang Require Import proofmode notation.
From iris.prelude Require Import options.
Definition CounterT Σ : Type :=
∀ (γ : gname) (z : Z), iProp Σ.
Definition IsCounterT Σ (N : namespace) : Type :=
∀ (γ : gname) (c : loc), iProp Σ.
Local Open Scope Z.
Section spec.
Context `{!heapGS Σ}.
Context (counterN : namespace).
Variables
(counter_new : val)
(counter_inc : val).
Variables
(Counter : CounterT Σ)
(IsCounter : IsCounterT Σ counterN).
Definition counter_new_spec' : Prop :=
{{{ True }}}
counter_new #()
{{{ γ c, RET #c; IsCounter γ c ∗ Counter γ 0 }}}.
Definition counter_inc_spec' : Prop :=
⊢ ∀ γ c,
IsCounter γ c -∗
<<{ ∀∀ (x : Z), Counter γ x }>>
counter_inc #c @ ⊤,↑counterN,∅
<<{ Counter γ (x + 1) | RET #x }>>.
End spec.
Record counter_code : Type := CounterCode {
counter_new : val;
counter_inc : val;
}.
Record counter_spec {Σ} `{!heapGS Σ} {counterN : namespace}
: Type := CounterSpec {
counter_spec_code :> counter_code;
Counter: CounterT Σ;
IsCounter : IsCounterT Σ counterN;
Counter_Timeless : ∀ γ c, Timeless (Counter γ c);
IsCounter_Persistent : ∀ γ l, Persistent (IsCounter γ l);
counter_new_spec : counter_new_spec' counterN counter_spec_code.(counter_new) Counter IsCounter;
counter_inc_spec : counter_inc_spec' counterN counter_spec_code.(counter_inc) Counter IsCounter;
}.
Global Arguments counter_spec : clear implicits.
Global Arguments counter_spec _ {_} _ : assert.
Global Existing Instances Counter_Timeless IsCounter_Persistent.