Skip to content
GitLab
Projects
Groups
Snippets
Help
Loading...
Help
What's new
10
Help
Support
Community forum
Keyboard shortcuts
?
Submit feedback
Contribute to GitLab
Sign in / Register
Toggle navigation
Open sidebar
Jonas Kastberg
iris
Commits
5d3c508d
Commit
5d3c508d
authored
Oct 26, 2018
by
Robbert Krebbers
Browse files
Options
Browse Files
Download
Email Patches
Plain Diff
New lemma `step_fupdN_wand`, and use instead of `step_fupdN_mono`.
parent
d9d67406
Changes
2
Hide whitespace changes
Inline
Side-by-side
Showing
2 changed files
with
13 additions
and
5 deletions
+13
-5
theories/base_logic/lib/fancy_updates.v
theories/base_logic/lib/fancy_updates.v
+3
-4
theories/bi/updates.v
theories/bi/updates.v
+10
-1
No files found.
theories/base_logic/lib/fancy_updates.v
View file @
5d3c508d
...
...
@@ -78,8 +78,8 @@ Proof.
iPoseProof
(
Hiter
Hinv
)
as
"H"
.
clear
Hiter
.
destruct
n
as
[|
n
].
-
iApply
fupd_plainly_mask_empty
.
iMod
"H"
as
%?
;
auto
.
-
i
PoseProof
(
step_fupdN_
mono
_
_
_
_
(|={
⊤
}=>
⌜φ⌝
)%
I
with
"H"
)
as
"H'"
.
{
iIntros
"H"
.
by
iApply
fupd_plain_mask_empty
.
}
-
i
Destruct
(
step_fupdN_
wand
_
_
_
_
(|={
⊤
}=>
⌜φ⌝
)%
I
with
"H
[]
"
)
as
"H'"
.
{
by
iApply
fupd_plain_mask_empty
.
}
rewrite
-
step_fupdN_S_fupd
.
iMod
(
step_fupdN_plain
with
"H'"
)
as
"Hφ"
.
iModIntro
.
iNext
.
rewrite
-
later_laterN
laterN_later
.
...
...
@@ -92,6 +92,5 @@ Lemma step_fupdN_soundness' `{invPreG Σ} φ n :
Proof
.
iIntros
(
Hiter
).
eapply
(
step_fupdN_soundness
_
n
).
iIntros
(
Hinv
).
iPoseProof
(
Hiter
Hinv
)
as
"Hiter"
.
iApply
(
step_fupdN_mono
with
"Hiter"
).
iIntros
(?).
iMod
(
fupd_intro_mask'
_
∅
)
as
"_"
;
auto
.
iApply
(
step_fupdN_wand
with
"Hiter"
).
by
iApply
(
fupd_mask_weaken
_
_
_
).
Qed
.
theories/bi/updates.v
View file @
5d3c508d
...
...
@@ -359,11 +359,20 @@ Section fupd_derived.
Qed
.
Lemma
step_fupdN_mono
E1
E2
n
P
Q
:
(
P
⊢
Q
)
→
(|={
E1
,
E2
}
▷
=>^
n
P
)
⊢
(|={
E1
,
E2
}
▷
=>^
n
Q
).
(
P
⊢
Q
)
→
(|={
E1
,
E2
}
▷
=>^
n
P
)
⊢
(|={
E1
,
E2
}
▷
=>^
n
Q
).
Proof
.
intros
HPQ
.
induction
n
as
[|
n
IH
]=>
//=.
rewrite
IH
//.
Qed
.
Lemma
step_fupdN_wand
E1
E2
n
P
Q
:
(|={
E1
,
E2
}
▷
=>^
n
P
)
-
∗
(
P
-
∗
Q
)
-
∗
(|={
E1
,
E2
}
▷
=>^
n
Q
).
Proof
.
apply
wand_intro_l
.
induction
n
as
[|
n
IH
]=>
/=.
{
by
rewrite
wand_elim_l
.
}
rewrite
-
IH
-
fupd_frame_l
later_sep
-
fupd_frame_l
.
by
apply
sep_mono
;
first
apply
later_intro
.
Qed
.
Lemma
step_fupdN_S_fupd
n
E
P
:
(|={
E
,
∅
}
▷
=>^(
S
n
)
P
)
⊣
⊢
(|={
E
,
∅
}
▷
=>^(
S
n
)
|={
E
}=>
P
).
Proof
.
...
...
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