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
Janno
iris-coq
Commits
1244bcfb
Commit
1244bcfb
authored
Jun 25, 2018
by
Ralf Jung
Browse files
prove a variant of wp_invariance that does not require finding fupd_intro_mask'
parent
42f0842f
Changes
1
Hide whitespace changes
Inline
Side-by-side
theories/program_logic/adequacy.v
View file @
1244bcfb
From
stdpp
Require
Import
namespaces
.
From
iris
.
program_logic
Require
Export
weakestpre
.
From
iris
.
algebra
Require
Import
gmap
auth
agree
gset
coPset
.
From
iris
.
base_logic
.
lib
Require
Import
wsat
.
...
...
@@ -203,3 +204,20 @@ Proof.
iDestruct
"Hwp"
as
(
Istate
)
"(HIstate & Hwp & Hclose)"
.
iApply
(@
wptp_invariance
_
_
(
IrisG
_
_
Hinv
Istate
))
;
eauto
with
iFrame
.
Qed
.
(* An equivalent version that does not require finding [fupd_intro_mask'], but
can be confusing to use. *)
Corollary
wp_invariance'
Σ
Λ
`
{
invPreG
Σ
}
s
e
σ
1
t2
σ
2
φ
:
(
∀
`
{
Hinv
:
invG
Σ
},
(|={
⊤
}=>
∃
stateI
:
state
Λ
→
iProp
Σ
,
let
_
:
irisG
Λ
Σ
:
=
IrisG
_
_
Hinv
stateI
in
stateI
σ
1
∗
WP
e
@
s
;
⊤
{{
_
,
True
}}
∗
(
stateI
σ
2
-
∗
∃
E
,
|={
⊤
,
E
}=>
⌜φ⌝
))%
I
)
→
rtc
step
([
e
],
σ
1
)
(
t2
,
σ
2
)
→
φ
.
Proof
.
intros
Hwp
.
eapply
wp_invariance
;
first
done
.
intros
Hinv
.
iMod
(
Hwp
Hinv
)
as
(
stateI
)
"(? & ? & Hφ)"
.
iModIntro
.
iExists
stateI
.
iFrame
.
iIntros
"Hσ"
.
iDestruct
(
"Hφ"
with
"Hσ"
)
as
(
E
)
">Hφ"
.
iMod
(
fupd_intro_mask'
)
as
"_"
;
last
by
iModIntro
.
solve_ndisj
.
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