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
Simon Spies
Iris
Commits
5cbe3c8e
Commit
5cbe3c8e
authored
Dec 09, 2016
by
Jacques-Henri Jourdan
Browse files
Prove [fupd_big_sepL].
parent
87abb52a
Changes
1
Hide whitespace changes
Inline
Side-by-side
base_logic/lib/fancy_updates.v
View file @
5cbe3c8e
...
...
@@ -121,6 +121,12 @@ Qed.
Lemma
fupd_sep
E
P
Q
:
(|={
E
}=>
P
)
∗
(|={
E
}=>
Q
)
={
E
}=
∗
P
∗
Q
.
Proof
.
by
rewrite
fupd_frame_r
fupd_frame_l
fupd_trans
.
Qed
.
Lemma
fupd_big_sepL
{
A
}
E
(
Φ
:
nat
→
A
→
iProp
Σ
)
(
l
:
list
A
)
:
([
∗
list
]
k
↦
x
∈
l
,
|={
E
}=>
Φ
k
x
)
={
E
}=
∗
[
∗
list
]
k
↦
x
∈
l
,
Φ
k
x
.
Proof
.
apply
(
big_opL_forall
(
λ
P
Q
,
P
={
E
}=
∗
Q
))
;
auto
using
fupd_intro
.
intros
P1
P2
HP
Q1
Q2
HQ
.
by
rewrite
HP
HQ
-
fupd_sep
.
Qed
.
Lemma
fupd_big_sepM
`
{
Countable
K
}
{
A
}
E
(
Φ
:
K
→
A
→
iProp
Σ
)
(
m
:
gmap
K
A
)
:
([
∗
map
]
k
↦
x
∈
m
,
|={
E
}=>
Φ
k
x
)
={
E
}=
∗
[
∗
map
]
k
↦
x
∈
m
,
Φ
k
x
.
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