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
George Pirlea
Iris
Commits
b7a767a7
Commit
b7a767a7
authored
Mar 29, 2019
by
Ralf Jung
Browse files
make fupd_plain_soundness more like soundness_bupd_plain
parent
d840f8cb
Changes
2
Hide whitespace changes
Inline
Side-by-side
theories/base_logic/lib/fancy_updates.v
View file @
b7a767a7
...
...
@@ -60,9 +60,9 @@ Proof.
Qed
.
Lemma
fupd_plain_soundness
`
{!
invPreG
Σ
}
E1
E2
(
P
:
iProp
Σ
)
`
{!
Plain
P
}
:
(
∀
`
{
Hinv
:
!
invG
Σ
},
(|={
E1
,
E2
}=>
P
)
%
I
)
→
(
▷
P
)%
I
.
(
∀
`
{
Hinv
:
!
invG
Σ
},
bi_emp_valid
(|={
E1
,
E2
}=>
P
))
→
bi_emp_valid
P
.
Proof
.
iIntros
(
Hfupd
).
iMod
wsat_alloc
as
(
Hinv
)
"[Hw HE]"
.
iIntros
(
Hfupd
).
apply
soundness_later
.
iMod
wsat_alloc
as
(
Hinv
)
"[Hw HE]"
.
iAssert
(|={
⊤
,
E2
}=>
P
)%
I
as
"H"
.
{
iMod
fupd_intro_mask'
;
last
iApply
Hfupd
.
done
.
}
rewrite
uPred_fupd_eq
/
uPred_fupd_def
.
...
...
@@ -74,7 +74,7 @@ Lemma step_fupdN_soundness `{!invPreG Σ} φ n :
φ
.
Proof
.
intros
Hiter
.
apply
(
soundness
(
M
:
=
iResUR
Σ
)
_
(
S
(
S
n
)
))
;
simpl
.
apply
(
soundness
(
M
:
=
iResUR
Σ
)
_
(
S
n
))
;
simpl
.
apply
(
fupd_plain_soundness
⊤
⊤
_
)=>
Hinv
.
iPoseProof
(
Hiter
Hinv
)
as
"H"
.
clear
Hiter
.
destruct
n
as
[|
n
].
...
...
theories/program_logic/total_adequacy.v
View file @
b7a767a7
...
...
@@ -123,7 +123,7 @@ Theorem twp_total Σ Λ `{!invPreG Σ} s e σ Φ :
stateI
σ
[]
0
∗
WP
e
@
s
;
⊤
[{
Φ
}])%
I
)
→
sn
erased_step
([
e
],
σ
).
(* i.e. ([e], σ) is strongly normalizing *)
Proof
.
intros
Hwp
.
apply
(
soundness
(
M
:
=
iResUR
Σ
)
_
2
)
;
simpl
.
intros
Hwp
.
apply
(
soundness
(
M
:
=
iResUR
Σ
)
_
1
)
;
simpl
.
apply
(
fupd_plain_soundness
⊤
⊤
_
)=>
Hinv
.
iMod
(
Hwp
)
as
(
stateI
fork_post
)
"[Hσ H]"
.
iApply
(@
twptp_total
_
_
(
IrisG
_
_
Hinv
stateI
fork_post
)
with
"Hσ"
).
...
...
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