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
Lennard Gäher
Iris
Commits
bd9a61e5
Commit
bd9a61e5
authored
May 14, 2020
by
Ralf Jung
Browse files
make inv_heapG part of heapG
parent
e2407f9f
Changes
7
Hide whitespace changes
Inline
Side-by-side
tests/heap_lang.v
View file @
bd9a61e5
...
...
@@ -204,7 +204,7 @@ Section tests.
End
tests
.
Section
notation_tests
.
Context
`
{!
heapG
Σ
,
!
inv_heapG
Σ
}.
Context
`
{!
heapG
Σ
}.
(* Make sure these parse and type-check. *)
Lemma
inv_mapsto_own_test
(
l
:
loc
)
:
⊢
l
↦
_
(
λ
_
,
True
)
#
5
.
Abort
.
...
...
@@ -274,4 +274,4 @@ Section error_tests.
End
error_tests
.
Lemma
heap_e_adequate
σ
:
adequate
NotStuck
heap_e
σ
(
λ
v
_
,
v
=
#
2
).
Proof
.
eapply
(
heap_adequacy
heap
Σ
)=>
?.
by
a
pply
heap_e_spec
.
Qed
.
Proof
.
eapply
(
heap_adequacy
heap
Σ
)=>
?.
iIntros
"_"
.
by
iA
pply
heap_e_spec
.
Qed
.
tests/one_shot.v
View file @
bd9a61e5
...
...
@@ -124,7 +124,7 @@ Definition clientΣ : gFunctors := #[ heapΣ; one_shotΣ; spawnΣ ].
(** This lemma implicitly shows that these functors are enough to meet
all library assumptions. *)
Lemma
client_adequate
σ
:
adequate
NotStuck
client
σ
(
λ
_
_
,
True
).
Proof
.
apply
(
heap_adequacy
client
Σ
)=>
?.
a
pply
client_safe
.
Qed
.
Proof
.
apply
(
heap_adequacy
client
Σ
)=>
?.
iIntros
"_"
.
iA
pply
client_safe
.
Qed
.
(* Since we check the output of the test files, this means
our test suite will fail if we ever accidentally add an axiom
...
...
tests/one_shot_once.v
View file @
bd9a61e5
...
...
@@ -142,4 +142,4 @@ Definition clientΣ : gFunctors := #[ heapΣ; one_shotΣ; spawnΣ ].
(** This lemma implicitly shows that these functors are enough to meet
all library assumptions. *)
Lemma
client_adequate
σ
:
adequate
NotStuck
client
σ
(
λ
_
_
,
True
).
Proof
.
apply
(
heap_adequacy
client
Σ
)=>
?.
a
pply
client_safe
.
Qed
.
Proof
.
apply
(
heap_adequacy
client
Σ
)=>
?.
iIntros
"_"
.
iA
pply
client_safe
.
Qed
.
theories/base_logic/lib/gen_inv_heap.v
View file @
bd9a61e5
...
...
@@ -113,7 +113,7 @@ Section to_inv_heap.
Qed
.
End
to_inv_heap
.
Lemma
inv_heap_init
{
L
V
:
Type
}
`
{
Countable
L
,
!
invG
Σ
,
!
gen_heapG
L
V
Σ
,
!
inv_heapPreG
L
V
Σ
}
E
:
Lemma
inv_heap_init
(
L
V
:
Type
)
`
{
Countable
L
,
!
invG
Σ
,
!
gen_heapG
L
V
Σ
,
!
inv_heapPreG
L
V
Σ
}
E
:
⊢
|==>
∃
_
:
inv_heapG
L
V
Σ
,
|={
E
}=>
inv_heap_inv
L
V
.
Proof
.
iMod
(
own_alloc
(
●
(
to_inv_heap
∅
)))
as
(
γ
)
"H●"
.
...
...
theories/heap_lang/adequacy.v
View file @
bd9a61e5
...
...
@@ -7,23 +7,25 @@ Set Default Proof Using "Type".
Class
heapPreG
Σ
:
=
HeapPreG
{
heap_preG_iris
:
>
invPreG
Σ
;
heap_preG_heap
:
>
gen_heapPreG
loc
val
Σ
;
heap_preG_inv_heap
:
>
inv_heapPreG
loc
val
Σ
;
heap_preG_proph
:
>
proph_mapPreG
proph_id
(
val
*
val
)
Σ
;
}.
Definition
heap
Σ
:
gFunctors
:
=
#[
inv
Σ
;
gen_heap
Σ
loc
val
;
proph_map
Σ
proph_id
(
val
*
val
)].
#[
inv
Σ
;
gen_heap
Σ
loc
val
;
inv_heap
Σ
loc
val
;
proph_map
Σ
proph_id
(
val
*
val
)].
Instance
subG_heapPreG
{
Σ
}
:
subG
heap
Σ
Σ
→
heapPreG
Σ
.
Proof
.
solve_inG
.
Qed
.
Definition
heap_adequacy
Σ
`
{!
heapPreG
Σ
}
s
e
σ
φ
:
(
∀
`
{!
heapG
Σ
},
⊢
WP
e
@
s
;
⊤
{{
v
,
⌜φ
v
⌝
}})
→
(
∀
`
{!
heapG
Σ
},
⊢
inv_heap_inv
-
∗
WP
e
@
s
;
⊤
{{
v
,
⌜φ
v
⌝
}})
→
adequate
s
e
σ
(
λ
v
_
,
φ
v
).
Proof
.
intros
Hwp
;
eapply
(
wp_adequacy
_
_
)
;
iIntros
(??)
""
.
iMod
(
gen_heap_init
σ
.(
heap
))
as
(?)
"Hh"
.
iMod
(
inv_heap_init
loc
val
)
as
(?)
">Hi"
.
iMod
(
proph_map_init
κ
s
σ
.(
used_proph_id
))
as
(?)
"Hp"
.
iModIntro
.
iExists
(
λ
σ
κ
s
,
(
gen_heap_ctx
σ
.(
heap
)
∗
proph_map_ctx
κ
s
σ
.(
used_proph_id
))%
I
),
(
λ
_
,
True
%
I
).
iFrame
.
iApply
(
Hwp
(
HeapG
_
_
_
_
)).
iFrame
.
iApply
(
Hwp
(
HeapG
_
_
_
_
_
)).
done
.
Qed
.
theories/heap_lang/lifting.v
View file @
bd9a61e5
...
...
@@ -11,6 +11,7 @@ Set Default Proof Using "Type".
Class
heapG
Σ
:
=
HeapG
{
heapG_invG
:
invG
Σ
;
heapG_gen_heapG
:
>
gen_heapG
loc
val
Σ
;
heapG_inv_heapG
:
>
inv_heapG
loc
val
Σ
;
heapG_proph_mapG
:
>
proph_mapG
proph_id
(
val
*
val
)
Σ
;
}.
...
...
@@ -34,7 +35,6 @@ Notation "l ↦□ I" := (inv_mapsto (L:=loc) (V:=val) l I%stdpp%type)
(
at
level
20
,
format
"l ↦□ I"
)
:
bi_scope
.
Notation
"l ↦_ I v"
:
=
(
inv_mapsto_own
(
L
:
=
loc
)
(
V
:
=
val
)
l
v
I
%
stdpp
%
type
)
(
at
level
20
,
I
at
level
9
,
format
"l ↦_ I v"
)
:
bi_scope
.
Notation
inv_heapG
:
=
(
inv_heapG
loc
val
).
Notation
inv_heap_inv
:
=
(
inv_heap_inv
loc
val
).
(** The tactic [inv_head_step] performs inversion on hypotheses of the shape
...
...
theories/heap_lang/total_adequacy.v
View file @
bd9a61e5
...
...
@@ -5,15 +5,16 @@ From iris.heap_lang Require Import proofmode notation.
Set
Default
Proof
Using
"Type"
.
Definition
heap_total
Σ
`
{!
heapPreG
Σ
}
s
e
σ
φ
:
(
∀
`
{!
heapG
Σ
},
⊢
WP
e
@
s
;
⊤
[{
v
,
⌜φ
v
⌝
}])
→
(
∀
`
{!
heapG
Σ
},
⊢
inv_heap_inv
-
∗
WP
e
@
s
;
⊤
[{
v
,
⌜φ
v
⌝
}])
→
sn
erased_step
([
e
],
σ
).
Proof
.
intros
Hwp
;
eapply
(
twp_total
_
_
)
;
iIntros
(?)
""
.
iMod
(
gen_heap_init
σ
.(
heap
))
as
(?)
"Hh"
.
iMod
(
inv_heap_init
loc
val
)
as
(?)
">Hi"
.
iMod
(
proph_map_init
[]
σ
.(
used_proph_id
))
as
(?)
"Hp"
.
iModIntro
.
iExists
(
λ
σ
κ
s
_
,
(
gen_heap_ctx
σ
.(
heap
)
∗
proph_map_ctx
κ
s
σ
.(
used_proph_id
))%
I
),
(
λ
_
,
True
%
I
)
;
iFrame
.
iApply
(
Hwp
(
HeapG
_
_
_
_
)).
iApply
(
Hwp
(
HeapG
_
_
_
_
_
)).
done
.
Qed
.
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