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
Rodolphe Lepigre
Iris
Commits
282094bf
Commit
282094bf
authored
Sep 14, 2016
by
Amin Timany
Browse files
Merge branch 'master' of
https://gitlab.mpi-sws.org/FP/iris-coq
parents
5651e5a8
4cf7fb98
Changes
1
Hide whitespace changes
Inline
Side-by-side
algebra/upred.v
View file @
282094bf
...
...
@@ -684,6 +684,15 @@ Proof.
-
by
apply
impl_intro_l
;
rewrite
left_id
.
Qed
.
Lemma
exists_impl_forall
{
A
}
P
(
Ψ
:
A
→
uPred
M
)
:
((
∃
x
:
A
,
Ψ
x
)
→
P
)
⊣
⊢
∀
x
:
A
,
Ψ
x
→
P
.
Proof
.
apply
equiv_spec
;
split
.
-
apply
forall_intro
=>
x
.
by
rewrite
-
exist_intro
.
-
apply
impl_intro_r
,
impl_elim_r'
,
exist_elim
=>
x
.
apply
impl_intro_r
.
by
rewrite
(
forall_elim
x
)
impl_elim_r
.
Qed
.
Lemma
or_and_l
P
Q
R
:
P
∨
Q
∧
R
⊣
⊢
(
P
∨
Q
)
∧
(
P
∨
R
).
Proof
.
apply
(
anti_symm
(
⊢
))
;
first
auto
.
...
...
@@ -817,7 +826,7 @@ Proof.
intros
P
;
unseal
;
split
=>
n
x
Hvalid
;
split
.
-
intros
(
x1
&
x2
&?&
_
&?)
;
cofe_subst
;
eauto
using
uPred_mono
,
cmra_includedN_r
.
-
by
intros
?
;
exists
(
core
x
),
x
;
rewrite
cmra_core_l
.
Qed
.
Qed
.
Global
Instance
sep_comm
:
Comm
(
⊣
⊢
)
(@
uPred_sep
M
).
Proof
.
by
intros
P
Q
;
unseal
;
split
=>
n
x
?
;
split
;
...
...
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