Skip to content
Snippets Groups Projects
Commit 7527bd61 authored by Robbert Krebbers's avatar Robbert Krebbers
Browse files

Rename "inc" of counter into "incr".

parent 3f678b90
No related branches found
No related tags found
No related merge requests found
...@@ -5,12 +5,12 @@ From iris.algebra Require Import frac auth. ...@@ -5,12 +5,12 @@ From iris.algebra Require Import frac auth.
From iris.heap_lang Require Import proofmode notation. From iris.heap_lang Require Import proofmode notation.
Definition newcounter : val := λ: <>, ref #0. Definition newcounter : val := λ: <>, ref #0.
Definition inc : val := Definition incr : val :=
rec: "inc" "l" := rec: "incr" "l" :=
let: "n" := !"l" in let: "n" := !"l" in
if: CAS "l" "n" (#1 + "n") then #() else "inc" "l". if: CAS "l" "n" (#1 + "n") then #() else "incr" "l".
Definition read : val := λ: "l", !"l". Definition read : val := λ: "l", !"l".
Global Opaque newcounter inc get. Global Opaque newcounter incr get.
(** Monotone counter *) (** Monotone counter *)
Class mcounterG Σ := MCounterG { mcounter_inG :> inG Σ (authR mnatUR) }. Class mcounterG Σ := MCounterG { mcounter_inG :> inG Σ (authR mnatUR) }.
...@@ -44,8 +44,8 @@ Section mono_proof. ...@@ -44,8 +44,8 @@ Section mono_proof.
iModIntro. iApply "HΦ". rewrite /mcounter; eauto 10. iModIntro. iApply "HΦ". rewrite /mcounter; eauto 10.
Qed. Qed.
Lemma inc_mono_spec l n : Lemma incr_mono_spec l n :
{{{ mcounter l n }}} inc #l {{{ RET #(); mcounter l (S n) }}}. {{{ mcounter l n }}} incr #l {{{ RET #(); mcounter l (S n) }}}.
Proof. Proof.
iIntros (Φ) "Hl HΦ". iLöb as "IH". wp_rec. iIntros (Φ) "Hl HΦ". iLöb as "IH". wp_rec.
iDestruct "Hl" as (γ) "(% & #? & #Hinv & Hγf)". iDestruct "Hl" as (γ) "(% & #? & #Hinv & Hγf)".
...@@ -122,8 +122,8 @@ Section contrib_spec. ...@@ -122,8 +122,8 @@ Section contrib_spec.
iModIntro. iApply "HΦ". rewrite /ccounter_ctx /ccounter; eauto 10. iModIntro. iApply "HΦ". rewrite /ccounter_ctx /ccounter; eauto 10.
Qed. Qed.
Lemma inc_contrib_spec γ l q n : Lemma incr_contrib_spec γ l q n :
{{{ ccounter_ctx γ l ccounter γ q n }}} inc #l {{{ ccounter_ctx γ l ccounter γ q n }}} incr #l
{{{ RET #(); ccounter γ q (S n) }}}. {{{ RET #(); ccounter γ q (S n) }}}.
Proof. Proof.
iIntros (Φ) "(#(%&?&?) & Hγf) HΦ". iLöb as "IH". wp_rec. iIntros (Φ) "(#(%&?&?) & Hγf) HΦ". iLöb as "IH". wp_rec.
......
...@@ -11,12 +11,12 @@ From iris.heap_lang Require Import proofmode notation. ...@@ -11,12 +11,12 @@ From iris.heap_lang Require Import proofmode notation.
Import uPred. Import uPred.
Definition newcounter : val := λ: <>, ref #0. Definition newcounter : val := λ: <>, ref #0.
Definition inc : val := Definition incr : val :=
rec: "inc" "l" := rec: "incr" "l" :=
let: "n" := !"l" in let: "n" := !"l" in
if: CAS "l" "n" (#1 + "n") then #() else "inc" "l". if: CAS "l" "n" (#1 + "n") then #() else "incr" "l".
Definition read : val := λ: "l", !"l". Definition read : val := λ: "l", !"l".
Global Opaque newcounter inc read. Global Opaque newcounter incr read.
(** The CMRA we need. *) (** The CMRA we need. *)
Inductive M := Auth : nat M | Frag : nat M | Bot. Inductive M := Auth : nat M | Frag : nat M | Bot.
...@@ -103,8 +103,8 @@ Proof. ...@@ -103,8 +103,8 @@ Proof.
iModIntro. rewrite /C; eauto 10. iModIntro. rewrite /C; eauto 10.
Qed. Qed.
Lemma inc_spec l n : Lemma incr_spec l n :
{{ C l n }} inc #l {{ v, v = #() C l (S n) }}. {{ C l n }} incr #l {{ v, v = #() C l (S n) }}.
Proof. Proof.
iIntros "!# Hl /=". iLöb as "IH". wp_rec. iIntros "!# Hl /=". iLöb as "IH". wp_rec.
iDestruct "Hl" as (N γ) "(% & #Hh & #Hinv & Hγf)". iDestruct "Hl" as (N γ) "(% & #Hh & #Hinv & Hγf)".
......
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment