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
George Pirlea
Iris
Commits
352292a8
Commit
352292a8
authored
Mar 04, 2018
by
Jacques-Henri Jourdan
Browse files
We usually fail faster to prove IntoPure than Absorbing/Affine.
parent
3b03bb6d
Changes
1
Hide whitespace changes
Inline
Side-by-side
theories/proofmode/class_instances.v
View file @
352292a8
...
...
@@ -613,9 +613,9 @@ Proof.
apply
pure_elim_l
=>
H
φ
.
by
rewrite
-(
exist_intro
H
φ
).
Qed
.
Global
Instance
into_exist_sep_pure
P
Q
φ
:
TCOr
(
Affine
P
)
(
Absorbing
Q
)
→
IntoPureT
P
φ
→
IntoExist
(
P
∗
Q
)
(
λ
_
:
φ
,
Q
).
IntoPureT
P
φ
→
TCOr
(
Affine
P
)
(
Absorbing
Q
)
→
IntoExist
(
P
∗
Q
)
(
λ
_
:
φ
,
Q
).
Proof
.
intros
?
(
φ
'
&->&?).
rewrite
/
IntoExist
.
intros
(
φ
'
&->&?)
?
.
rewrite
/
IntoExist
.
eapply
(
pure_elim
φ
'
)
;
[
by
rewrite
(
into_pure
P
)
;
apply
sep_elim_l
,
_
|]=>?.
rewrite
-
exist_intro
//.
apply
sep_elim_r
,
_
.
Qed
.
...
...
@@ -658,10 +658,10 @@ Proof.
intros
(
φ
'
&->&?).
by
rewrite
/
FromForall
-
pure_impl_forall
(
into_pure
P
).
Qed
.
Global
Instance
from_forall_wand_pure
P
Q
φ
:
TCOr
(
Affine
P
)
(
Absorbing
Q
)
→
IntoPureT
P
φ
→
IntoPureT
P
φ
→
TCOr
(
Affine
P
)
(
Absorbing
Q
)
→
FromForall
(
P
-
∗
Q
)%
I
(
λ
_
:
φ
,
Q
)%
I
.
Proof
.
intros
[|]
(
φ
'
&->&?)
;
rewrite
/
FromForall
;
apply
wand_intro_r
.
intros
(
φ
'
&->&?)
[|]
;
rewrite
/
FromForall
;
apply
wand_intro_r
.
-
rewrite
-(
affine_affinely
P
)
(
into_pure
P
)
-
persistent_and_affinely_sep_r
.
apply
pure_elim_r
=>?.
by
rewrite
forall_elim
.
-
by
rewrite
(
into_pure
P
)
-
pure_wand_forall
wand_elim_l
.
...
...
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