Skip to content
GitLab
Explore
Sign in
Primary navigation
Search or go to…
Project
Iris
Manage
Activity
Members
Labels
Plan
Issues
Issue boards
Milestones
Wiki
Code
Merge requests
Repository
Branches
Commits
Tags
Repository graph
Compare revisions
Build
Pipelines
Jobs
Pipeline schedules
Artifacts
Deploy
Releases
Package Registry
Model registry
Operate
Terraform modules
Monitor
Service Desk
Analyze
Value stream analytics
Contributor analytics
CI/CD analytics
Repository analytics
Model experiments
Help
Help
Support
GitLab documentation
Compare GitLab plans
Community forum
Contribute to GitLab
Provide feedback
Terms and privacy
Keyboard shortcuts
?
Snippets
Groups
Projects
Show more breadcrumbs
Gaëtan Gilbert
Iris
Commits
df51d7fb
Commit
df51d7fb
authored
5 years ago
by
Robbert Krebbers
Browse files
Options
Downloads
Patches
Plain Diff
Add missing `Proper` instances for non-expansiveness of map operations.
Copied from std++, but adapted from `≡` to `≡{n}≡`.
parent
62d3e6cb
No related branches found
Branches containing commit
No related tags found
No related merge requests found
Changes
1
Hide whitespace changes
Inline
Side-by-side
Showing
1 changed file
theories/algebra/gmap.v
+38
-18
38 additions, 18 deletions
theories/algebra/gmap.v
with
38 additions
and
18 deletions
theories/algebra/gmap.v
+
38
−
18
View file @
df51d7fb
...
@@ -44,32 +44,28 @@ Proof. intros ? m m' ? i. by apply (discrete _). Qed.
...
@@ -44,32 +44,28 @@ Proof. intros ? m m' ? i. by apply (discrete _). Qed.
Global
Instance
gmapO_leibniz
:
LeibnizEquiv
A
→
LeibnizEquiv
gmapO
.
Global
Instance
gmapO_leibniz
:
LeibnizEquiv
A
→
LeibnizEquiv
gmapO
.
Proof
.
intros
;
change
(
LeibnizEquiv
(
gmap
K
A
));
apply
_
.
Qed
.
Proof
.
intros
;
change
(
LeibnizEquiv
(
gmap
K
A
));
apply
_
.
Qed
.
Global
Instance
lookup_ne
k
:
Global
Instance
lookup_ne
k
:
NonExpansive
(
lookup
k
:
gmap
K
A
→
option
A
)
.
NonExpansive
(
lookup
k
:
gmap
K
A
→
option
A
)
.
Proof
.
by
intros
m1
m2
.
Qed
.
Proof
.
by
intros
m1
m2
.
Qed
.
Global
Instance
lookup_proper
k
:
Global
Instance
partial_alter_ne
n
:
Proper
((
≡
)
==>
(
≡
))
(
lookup
k
:
gmap
K
A
→
option
A
)
:=
_
.
Proper
((
dist
n
==>
dist
n
)
==>
(
=
)
==>
dist
n
==>
dist
n
)
Global
Instance
alter_ne
(
f
:
A
→
A
)
(
k
:
K
)
n
:
(
partial_alter
(
M
:=
gmap
K
A
))
.
Proper
(
dist
n
==>
dist
n
)
f
→
Proper
(
dist
n
==>
dist
n
)
(
alter
f
k
)
.
Proof
.
Proof
.
intros
?
m
m'
Hm
k'
.
by
intros
f1
f2
Hf
i
?
<-
m1
m2
Hm
j
;
destruct
(
decide
(
i
=
j
))
as
[
->
|];
by
destruct
(
decide
(
k
=
k'
));
simplify_map_eq
;
rewrite
(
Hm
k'
)
.
rewrite
?lookup_partial_alter
?lookup_partial_alter_ne
//
;
Qed
.
try
apply
Hf
;
apply
lookup_ne
.
Global
Instance
insert_ne
i
:
NonExpansive2
(
insert
(
M
:=
gmap
K
A
)
i
)
.
Proof
.
intros
n
x
y
?
m
m'
?
j
;
destruct
(
decide
(
i
=
j
));
simplify_map_eq
;
[
by
constructor
|
by
apply
lookup_ne
]
.
Qed
.
Qed
.
Global
Instance
singleton_ne
i
:
Global
Instance
insert_ne
i
:
NonExpansive2
(
insert
(
M
:=
gmap
K
A
)
i
)
.
NonExpansive
(
singletonM
i
:
A
→
gmap
K
A
)
.
Proof
.
intros
n
x
y
?
m
m'
?
j
;
apply
partial_alter_ne
;
by
try
constructor
.
Qed
.
Global
Instance
singleton_ne
i
:
NonExpansive
(
singletonM
i
:
A
→
gmap
K
A
)
.
Proof
.
by
intros
????;
apply
insert_ne
.
Qed
.
Proof
.
by
intros
????;
apply
insert_ne
.
Qed
.
Global
Instance
delete_ne
i
:
Global
Instance
delete_ne
i
:
NonExpansive
(
delete
(
M
:=
gmap
K
A
)
i
)
.
NonExpansive
(
delete
(
M
:=
gmap
K
A
)
i
)
.
Proof
.
Proof
.
intros
n
m
m'
?
j
;
destruct
(
decide
(
i
=
j
));
simplify_map_eq
;
intros
n
m
m'
?
j
;
destruct
(
decide
(
i
=
j
));
simplify_map_eq
;
[
by
constructor
|
by
apply
lookup_ne
]
.
[
by
constructor
|
by
apply
lookup_ne
]
.
Qed
.
Qed
.
Global
Instance
alter_ne
(
f
:
A
→
A
)
(
k
:
K
)
n
:
Proper
(
dist
n
==>
dist
n
)
f
→
Proper
(
dist
n
==>
dist
n
)
(
alter
f
k
)
.
Proof
.
intros
?
m
m'
Hm
k'
.
by
apply
partial_alter_ne
;
[
solve_proper
|..]
.
Qed
.
Global
Instance
gmap_empty_discrete
:
Discrete
(
∅
:
gmap
K
A
)
.
Global
Instance
gmap_empty_discrete
:
Discrete
(
∅
:
gmap
K
A
)
.
Proof
.
Proof
.
...
@@ -99,6 +95,30 @@ End cofe.
...
@@ -99,6 +95,30 @@ End cofe.
Arguments
gmapO
_
{_
_}
_
.
Arguments
gmapO
_
{_
_}
_
.
(** Non-expansiveness of higher-order map functions and big-ops *)
Lemma
merge_ne
`{
Countable
K
}
{
A
B
C
:
ofeT
}
(
f
g
:
option
A
→
option
B
→
option
C
)
`{
!
DiagNone
f
,
!
DiagNone
g
}
n
:
((
dist
n
)
==>
(
dist
n
)
==>
(
dist
n
))
%
signature
f
g
→
((
dist
n
)
==>
(
dist
n
)
==>
(
dist
n
))
%
signature
(
merge
(
M
:=
gmap
K
)
f
)
(
merge
g
)
.
Proof
.
by
intros
Hf
??
Hm1
??
Hm2
i
;
rewrite
!
lookup_merge
//
;
apply
Hf
.
Qed
.
Instance
union_with_proper
`{
Countable
K
}
{
A
:
ofeT
}
n
:
Proper
(((
dist
n
)
==>
(
dist
n
)
==>
(
dist
n
))
==>
(
dist
n
)
==>
(
dist
n
)
==>
(
dist
n
))
(
union_with
(
M
:=
gmap
K
A
))
.
Proof
.
intros
??
Hf
??
Hm1
??
Hm2
i
;
apply
(
merge_ne
_
_);
auto
.
by
do
2
destruct
1
;
first
[
apply
Hf
|
constructor
]
.
Qed
.
Instance
map_fmap_proper
`{
Countable
K
}
{
A
B
:
ofeT
}
(
f
:
A
→
B
)
n
:
Proper
(
dist
n
==>
dist
n
)
f
→
Proper
(
dist
n
==>
dist
n
)
(
fmap
(
M
:=
gmap
K
)
f
)
.
Proof
.
intros
?
m
m'
?
k
;
rewrite
!
lookup_fmap
.
by
repeat
f_equiv
.
Qed
.
Instance
map_zip_with_proper
`{
Countable
K
}
{
A
B
C
:
ofeT
}
(
f
:
A
→
B
→
C
)
n
:
Proper
(
dist
n
==>
dist
n
==>
dist
n
)
f
→
Proper
(
dist
n
==>
dist
n
==>
dist
n
)
(
map_zip_with
(
M
:=
gmap
K
)
f
)
.
Proof
.
intros
Hf
m1
m1'
Hm1
m2
m2'
Hm2
.
apply
merge_ne
;
try
done
.
destruct
1
;
destruct
1
;
repeat
f_equiv
;
constructor
||
done
.
Qed
.
(* CMRA *)
(* CMRA *)
Section
cmra
.
Section
cmra
.
Context
`{
Countable
K
}
{
A
:
cmraT
}
.
Context
`{
Countable
K
}
{
A
:
cmraT
}
.
...
...
This diff is collapsed.
Click to expand it.
Preview
0%
Loading
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!
Save comment
Cancel
Please
register
or
sign in
to comment