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
George Pirlea
Iris
Commits
e46aa078
Commit
e46aa078
authored
Aug 06, 2016
by
Robbert Krebbers
Browse files
Add missing heap_lang/adequacy.
parent
5e5e90ea
Changes
1
Hide whitespace changes
Inline
Side-by-side
heap_lang/adequacy.v
0 → 100644
View file @
e46aa078
From
iris
.
program_logic
Require
Export
weakestpre
adequacy
.
From
iris
.
heap_lang
Require
Export
heap
.
From
iris
.
program_logic
Require
Import
auth
ownership
.
From
iris
.
heap_lang
Require
Import
proofmode
notation
.
From
iris
.
proofmode
Require
Import
tactics
weakestpre
.
Definition
heapGF
:
gFunctorList
:
=
[
authGF
heapUR
;
irisGF
heap_lang
].
Definition
heap_adequacy
Σ
`
{
irisPreG
heap_lang
Σ
,
authG
Σ
heapUR
}
e
σ
φ
:
(
∀
`
{
heapG
Σ
},
heap_ctx
⊢
WP
e
{{
v
,
■
φ
v
}})
→
adequate
e
σ
φ
.
Proof
.
intros
Hwp
;
eapply
(
wp_adequacy
Σ
)
;
iIntros
(?)
"Hσ"
.
iVs
(
auth_alloc
(
ownP
∘
of_heap
)
heapN
_
(
to_heap
σ
)
with
"[Hσ]"
)
as
(
γ
)
"[??]"
.
-
auto
using
to_heap_valid
.
-
rewrite
/=
(
from_to_heap
σ
)
;
auto
.
-
iApply
(
Hwp
(
HeapG
_
_
_
γ
)).
by
rewrite
/
heap_ctx
.
Qed
.
\ No newline at end of file
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