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
Iris
Iris
Commits
90f773c0
Commit
90f773c0
authored
Nov 29, 2016
by
Robbert Krebbers
Browse files
Prove that from_option is timeless.
parent
ce32b224
Pipeline
#3172
passed with stage
in 10 minutes and 39 seconds
Changes
1
Pipelines
1
Hide whitespace changes
Inline
Side-by-side
base_logic/derived.v
View file @
90f773c0
...
...
@@ -794,6 +794,9 @@ Proof.
apply
except_0_mono
.
rewrite
internal_eq_sym
.
apply
(
internal_eq_rewrite
b
a
(
uPred_ownM
))
;
first
apply
_;
auto
.
Qed
.
Global
Instance
from_option_timeless
{
A
}
P
(
Ψ
:
A
→
uPred
M
)
(
mx
:
option
A
)
:
(
∀
x
,
TimelessP
(
Ψ
x
))
→
TimelessP
P
→
TimelessP
(
from_option
Ψ
P
mx
).
Proof
.
destruct
mx
;
apply
_
.
Qed
.
(* Persistence *)
Global
Instance
pure_persistent
φ
:
PersistentP
(
⌜φ⌝
:
uPred
M
)%
I
.
...
...
Write
Preview
Supports
Markdown
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