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
Rodolphe Lepigre
Iris
Commits
7209226c
Commit
7209226c
authored
Jan 19, 2016
by
Robbert Krebbers
Browse files
Invariants are AlwaysStable.
parent
f5f9e7ff
Changes
1
Hide whitespace changes
Inline
Side-by-side
iris/ownership.v
View file @
7209226c
...
...
@@ -25,15 +25,14 @@ Proof.
apply
(
_:
Proper
(
_
==>
_
)
iProp_unfold
),
Later_contractive
in
HPQ
.
by
unfold
inv
;
rewrite
HPQ
.
Qed
.
Lemma
inv_duplicate
i
P
:
inv
i
P
≡
(
inv
i
P
★
inv
i
P
)%
I
.
Proof
.
by
rewrite
/
inv
-
uPred
.
own_op
Res_op
map_op_singleton
agree_idempotent
!(
left_id
_
_
).
Qed
.
Lemma
always_inv
i
P
:
(
□
inv
i
P
)%
I
≡
inv
i
P
.
Proof
.
by
apply
uPred
.
always_own
;
rewrite
Res_unit
!
ra_unit_empty
map_unit_singleton
.
Qed
.
Global
Instance
inv_always_stable
i
P
:
AlwaysStable
(
inv
i
P
).
Proof
.
by
rewrite
/
AlwaysStable
always_inv
.
Qed
.
Lemma
inv_sep_dup
i
P
:
inv
i
P
≡
(
inv
i
P
★
inv
i
P
)%
I
.
Proof
.
apply
(
uPred
.
always_sep_dup'
_
).
Qed
.
(* physical state *)
Lemma
ownP_twice
σ
1
σ
2
:
(
ownP
σ
1
★
ownP
σ
2
:
iProp
Σ
)
⊑
False
.
...
...
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