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
Marianna Rapoport
iris-coq
Commits
9f3a376e
Commit
9f3a376e
authored
Sep 30, 2016
by
Robbert Krebbers
Browse files
Big ops and None.
parent
7227cf28
Changes
1
Hide whitespace changes
Inline
Side-by-side
algebra/cmra_big_op.v
View file @
9f3a376e
...
...
@@ -375,6 +375,33 @@ Section gset.
End
gset
.
End
big_op
.
(** Option *)
Lemma
big_opL_None
{
M
:
cmraT
}
{
A
}
(
f
:
nat
→
A
→
option
M
)
l
:
([
⋅
list
]
k
↦
x
∈
l
,
f
k
x
)
=
None
↔
∀
k
x
,
l
!!
k
=
Some
x
→
f
k
x
=
None
.
Proof
.
revert
f
.
induction
l
as
[|
x
l
IH
]=>
f
//=.
rewrite
big_opL_cons
op_None
IH
.
split
.
-
intros
[??]
[|
k
]
y
?
;
naive_solver
.
-
intros
Hl
.
split
.
by
apply
(
Hl
0
).
intros
k
.
apply
(
Hl
(
S
k
)).
Qed
.
Lemma
big_opM_None
{
M
:
cmraT
}
`
{
Countable
K
}
{
A
}
(
f
:
K
→
A
→
option
M
)
m
:
([
⋅
map
]
k
↦
x
∈
m
,
f
k
x
)
=
None
↔
∀
k
x
,
m
!!
k
=
Some
x
→
f
k
x
=
None
.
Proof
.
induction
m
as
[|
i
x
m
?
IH
]
using
map_ind
=>
//=.
rewrite
-
equiv_None
big_opM_insert
//
equiv_None
op_None
IH
.
split
.
{
intros
[??]
k
y
.
rewrite
lookup_insert_Some
;
naive_solver
.
}
intros
Hm
;
split
.
-
apply
(
Hm
i
).
by
simplify_map_eq
.
-
intros
k
y
?.
apply
(
Hm
k
).
by
simplify_map_eq
.
Qed
.
Lemma
big_opS_None
{
M
:
cmraT
}
`
{
Countable
A
}
(
f
:
A
→
option
M
)
X
:
([
⋅
set
]
x
∈
X
,
f
x
)
=
None
↔
∀
x
,
x
∈
X
→
f
x
=
None
.
Proof
.
induction
X
as
[|
x
X
?
IH
]
using
collection_ind_L
;
[
done
|].
rewrite
-
equiv_None
big_opS_insert
//
equiv_None
op_None
IH
.
set_solver
.
Qed
.
(** Commuting with respect to homomorphisms *)
Lemma
big_opL_commute
{
M1
M2
:
ucmraT
}
{
A
}
(
h
:
M1
→
M2
)
`
{!
UCMRAHomomorphism
h
}
(
f
:
nat
→
A
→
M1
)
l
:
h
([
⋅
list
]
k
↦
x
∈
l
,
f
k
x
)
≡
([
⋅
list
]
k
↦
x
∈
l
,
h
(
f
k
x
)).
...
...
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