Skip to content
GitLab
Projects
Groups
Snippets
Help
Loading...
Help
What's new
10
Help
Support
Community forum
Keyboard shortcuts
?
Submit feedback
Contribute to GitLab
Sign in / Register
Toggle navigation
Open sidebar
Simon Spies
Iris
Commits
b79eb4fe
Commit
b79eb4fe
authored
Mar 04, 2018
by
Robbert Krebbers
Browse files
Options
Browse Files
Download
Email Patches
Plain Diff
`Cancelable` instances for gmap/option.
parent
5e608261
Changes
2
Hide whitespace changes
Inline
Side-by-side
Showing
2 changed files
with
10 additions
and
0 deletions
+10
-0
theories/algebra/cmra.v
theories/algebra/cmra.v
+4
-0
theories/algebra/gmap.v
theories/algebra/gmap.v
+6
-0
No files found.
theories/algebra/cmra.v
View file @
b79eb4fe
...
...
@@ -1395,6 +1395,10 @@ Section option.
by
eapply
(
cmra_validN_le
n
)
;
last
lia
.
-
done
.
Qed
.
Global
Instance
option_cancelable
(
ma
:
option
A
)
:
(
∀
a
:
A
,
IdFree
a
)
→
(
∀
a
:
A
,
Cancelable
a
)
→
Cancelable
ma
.
Proof
.
destruct
ma
;
apply
_
.
Qed
.
End
option
.
Arguments
optionR
:
clear
implicits
.
...
...
theories/algebra/gmap.v
View file @
b79eb4fe
...
...
@@ -288,6 +288,12 @@ Proof.
-
by
rewrite
lookup_singleton_ne
//
!(
left_id
None
_
).
Qed
.
Global
Instance
gmap_cancelable
(
m
:
gmap
K
A
)
:
(
∀
x
:
A
,
IdFree
x
)
→
(
∀
x
:
A
,
Cancelable
x
)
→
Cancelable
m
.
Proof
.
intros
??
n
m1
m2
??
i
.
apply
(
cancelableN
(
m
!!
i
))
;
by
rewrite
-!
lookup_op
.
Qed
.
Lemma
insert_op
m1
m2
i
x
y
:
<[
i
:
=
x
⋅
y
]>(
m1
⋅
m2
)
=
<[
i
:
=
x
]>
m1
⋅
<[
i
:
=
y
]>
m2
.
Proof
.
by
rewrite
(
insert_merge
(
⋅
)
m1
m2
i
(
x
⋅
y
)
x
y
).
Qed
.
...
...
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