Skip to content
GitLab
Projects
Groups
Snippets
/
Help
Help
Support
Community forum
Keyboard shortcuts
?
Submit feedback
Contribute to GitLab
Sign in / Register
Toggle navigation
Menu
Open sidebar
Dmitry Khalanskiy
Iris
Commits
aeff566c
Commit
aeff566c
authored
Aug 23, 2016
by
Robbert Krebbers
Browse files
Persistence of invariant in wp_invariance is not needed.
parent
46cf3677
Changes
1
Hide whitespace changes
Inline
Side-by-side
program_logic/adequacy.v
View file @
aeff566c
...
...
@@ -128,13 +128,12 @@ Proof.
Qed
.
Lemma
wptp_invariance
n
e1
e2
t1
t2
σ
1
σ
2
I
φ
Φ
:
PersistentP
I
→
nsteps
step
n
(
e1
::
t1
,
σ
1
)
(
t2
,
σ
2
)
→
(
I
={
⊤
,
∅
}=>
∃
σ
'
,
ownP
σ
'
∧
■
φ
σ
'
)
→
I
★
world
σ
1
★
WP
e1
{{
Φ
}}
★
wptp
t1
⊢
Nat
.
iter
(
S
(
S
n
))
(
λ
P
,
|=
r
=>
▷
P
)
(
■
φ
σ
2
).
Proof
.
intros
?
?
HI
.
rewrite
wptp_steps
//.
intros
?
HI
.
rewrite
wptp_steps
//.
rewrite
(
Nat_iter_S_r
(
S
n
))
rvs_iter_frame_l
.
apply
rvs_iter_mono
.
iIntros
"[HI H]"
.
iDestruct
"H"
as
(
e2'
t2'
)
"(% & (Hw&HE&Hσ) & _)"
;
subst
.
...
...
@@ -163,13 +162,12 @@ Proof.
Qed
.
Theorem
wp_invariance
Σ
`
{
irisPreG
Λ
Σ
}
e
σ
1
t2
σ
2
I
φ
Φ
:
PersistentP
I
→
(
∀
`
{
irisG
Λ
Σ
},
ownP
σ
1
={
⊤
}=>
I
★
WP
e
{{
Φ
}})
→
(
∀
`
{
irisG
Λ
Σ
},
I
={
⊤
,
∅
}=>
∃
σ
'
,
ownP
σ
'
∧
■
φ
σ
'
)
→
rtc
step
([
e
],
σ
1
)
(
t2
,
σ
2
)
→
φ
σ
2
.
Proof
.
intros
?
Hwp
HI
[
n
?]%
rtc_nsteps
.
intros
Hwp
HI
[
n
?]%
rtc_nsteps
.
eapply
(
adequacy
(
M
:
=
iResUR
Σ
)
_
(
S
(
S
(
S
n
))))
;
iIntros
""
.
rewrite
Nat_iter_S
.
iVs
(
iris_alloc
σ
1
)
as
(?)
"(Hw & HE & ? & Hσ)"
.
rewrite
pvs_eq
in
Hwp
.
...
...
Write
Preview
Supports
Markdown
0%
Try again
or
attach a new 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