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
be9f621b
Commit
be9f621b
authored
Feb 13, 2016
by
Ralf Jung
Browse files
heap tuning
parent
1c8ba455
Changes
2
Hide whitespace changes
Inline
Side-by-side
algebra/cmra.v
View file @
be9f621b
...
...
@@ -338,6 +338,9 @@ Proof. by rewrite equiv_dist=>?? n; apply (local_updateN L). Qed.
Global
Instance
local_update_op
x
:
LocalUpdate
(
λ
_
,
True
)
(
op
x
).
Proof
.
split
.
apply
_
.
by
intros
n
y1
y2
_
_;
rewrite
assoc
.
Qed
.
Global
Instance
local_update_id
:
LocalUpdate
(
λ
_
,
True
)
(@
id
A
).
Proof
.
split
;
auto
with
typeclass_instances
.
Qed
.
(** ** Updates *)
Global
Instance
cmra_update_preorder
:
PreOrder
(@
cmra_update
A
).
Proof
.
split
.
by
intros
x
y
.
intros
x
y
y'
??
z
?
;
naive_solver
.
Qed
.
...
...
heap_lang/heap.v
View file @
be9f621b
...
...
@@ -44,7 +44,7 @@ Section heap.
Definition
heap_ctx
(
γ
:
gname
)
:
iPropG
heap_lang
Σ
:
=
auth_ctx
HeapI
γ
N
heap_inv
.
Global
Instance
heap_inv_
ne
:
Proper
((
≡
)
==>
(
≡
))
heap_inv
.
Global
Instance
heap_inv_
proper
:
Proper
((
≡
)
==>
(
≡
))
heap_inv
.
Proof
.
move
=>?
?
EQ
.
rewrite
/
heap_inv
/
from_heap
.
(* TODO I guess we need some lemma about omap? *)
...
...
@@ -91,8 +91,8 @@ Section heap.
apply
later_mono
,
wand_intro_l
.
rewrite
left_id
const_equiv
//
left_id
.
by
rewrite
-
later_intro
.
Unshelve
.
(* TODO Make it so that this becomes a goal, not shelved. *)
{
eexists
.
eauto
.
}
{
split
;
first
by
apply
_
.
done
.
}
Qed
.
End
heap
.
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