Skip to content
GitLab
Menu
Projects
Groups
Snippets
Help
Help
Support
Community forum
Keyboard shortcuts
?
Submit feedback
Contribute to GitLab
Sign in / Register
Toggle navigation
Menu
Open sidebar
Iris
Iris
Commits
eafd77e6
Commit
eafd77e6
authored
Jun 20, 2019
by
Ralf Jung
Browse files
use equality instead of let binding
parent
242cee02
Changes
2
Hide whitespace changes
Inline
Side-by-side
theories/heap_lang/lang.v
View file @
eafd77e6
...
...
@@ -634,11 +634,11 @@ Inductive head_step : expr → state → list observation → expr → state →
[]
(
Val
$
LitV
LitUnit
)
(
state_upd_heap
<[
l
:
=
v
]>
σ
)
[]
|
CompareExchangeS
l
v1
v2
vl
σ
:
|
CompareExchangeS
l
v1
v2
vl
σ
b
:
vals_cas_compare_safe
vl
v1
→
σ
.(
heap
)
!!
l
=
Some
vl
→
(* Crucially, this compares the same way as [EqOp]! *)
let
b
:
=
bool_decide
(
val_for_compare
vl
=
val_for_compare
v1
)
in
b
=
bool_decide
(
val_for_compare
vl
=
val_for_compare
v1
)
→
head_step
(
CompareExchange
(
Val
$
LitV
$
LitLoc
l
)
(
Val
v1
)
(
Val
v2
))
σ
[]
(
Val
$
PairV
(
LitV
$
LitBool
b
)
vl
)
(
if
b
then
state_upd_heap
<[
l
:
=
v2
]>
σ
else
σ
)
...
...
theories/heap_lang/lifting.v
View file @
eafd77e6
...
...
@@ -382,7 +382,7 @@ Proof.
iIntros
(??
Φ
)
">Hl HΦ"
.
iApply
wp_lift_atomic_head_step_no_fork
;
auto
.
iIntros
(
σ
1
κ
κ
s
n
)
"[Hσ Hκs] !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iSplit
;
first
by
eauto
.
iNext
;
iIntros
(
v2'
σ
2
efs
Hstep
)
;
inv_head_step
.
rewrite
/
b
bool_decide_false
//.
rewrite
bool_decide_false
//.
iModIntro
;
iSplit
=>
//.
iFrame
.
by
iApply
"HΦ"
.
Qed
.
Lemma
twp_cas_fail
s
E
l
q
v'
v1
v2
:
...
...
@@ -393,7 +393,7 @@ Proof.
iIntros
(??
Φ
)
"Hl HΦ"
.
iApply
twp_lift_atomic_head_step_no_fork
;
auto
.
iIntros
(
σ
1
κ
s
n
)
"[Hσ Hκs] !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iSplit
;
first
by
eauto
.
iIntros
(
κ
v2'
σ
2
efs
Hstep
)
;
inv_head_step
.
rewrite
/
b
bool_decide_false
//.
rewrite
bool_decide_false
//.
iModIntro
;
iSplit
=>
//.
iSplit
;
first
done
.
iFrame
.
by
iApply
"HΦ"
.
Qed
.
...
...
@@ -405,7 +405,7 @@ Proof.
iIntros
(??
Φ
)
">Hl HΦ"
.
iApply
wp_lift_atomic_head_step_no_fork
;
auto
.
iIntros
(
σ
1
κ
κ
s
n
)
"[Hσ Hκs] !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iSplit
;
first
by
eauto
.
iNext
;
iIntros
(
v2'
σ
2
efs
Hstep
)
;
inv_head_step
.
rewrite
/
b
bool_decide_true
//.
rewrite
bool_decide_true
//.
iMod
(@
gen_heap_update
with
"Hσ Hl"
)
as
"[$ Hl]"
.
iModIntro
.
iSplit
=>//.
iFrame
.
by
iApply
"HΦ"
.
Qed
.
...
...
@@ -417,7 +417,7 @@ Proof.
iIntros
(??
Φ
)
"Hl HΦ"
.
iApply
twp_lift_atomic_head_step_no_fork
;
auto
.
iIntros
(
σ
1
κ
s
n
)
"[Hσ Hκs] !>"
.
iDestruct
(@
gen_heap_valid
with
"Hσ Hl"
)
as
%?.
iSplit
;
first
by
eauto
.
iIntros
(
κ
v2'
σ
2
efs
Hstep
)
;
inv_head_step
.
rewrite
/
b
bool_decide_true
//.
rewrite
bool_decide_true
//.
iMod
(@
gen_heap_update
with
"Hσ Hl"
)
as
"[$ Hl]"
.
iModIntro
.
iSplit
=>//.
iSplit
;
first
done
.
iFrame
.
by
iApply
"HΦ"
.
Qed
.
...
...
Write
Preview
Markdown
is supported
0%
Try again
or
attach a new file
.
Attach a file
Cancel
You are about to add
0
people
to the discussion. Proceed with caution.
Finish editing this message first!
Cancel
Please
register
or
sign in
to comment