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
e3dd5307
Verified
Commit
e3dd5307
authored
Aug 13, 2019
by
Paolo G. Giarrusso
Browse files
Flip cmra_morphism_pcore
parent
a51d5e1f
Changes
3
Hide whitespace changes
Inline
Side-by-side
CHANGELOG.md
View file @
e3dd5307
...
...
@@ -162,7 +162,7 @@ Changes in Coq:
*
Lemma
`prop_ext`
works in both directions; its default direction is the
opposite of what it used to be.
*
Make direction of
`f_op`
rewrite lemmas more consistent: Flip
`pair_op`
,
`Cinl_op`
,
`Cinr_op`
,
`cmra_morphism_op`
.
`Cinl_op`
,
`Cinr_op`
,
`cmra_morphism_op`
,
`cmra_morphism_pcore`
.
*
Rename
`C`
suffixes into
`O`
since we no longer use COFEs but OFEs. Also
rename
`ofe_fun`
into
`discrete_fun`
and the corresponding notation
`-c>`
into
`-d>`
. The renaming can be automatically done using the following script (on
...
...
theories/algebra/cmra.v
View file @
e3dd5307
...
...
@@ -251,7 +251,7 @@ Hint Mode CmraDiscrete ! : typeclass_instances.
Class
CmraMorphism
{
A
B
:
cmraT
}
(
f
:
A
→
B
)
:
=
{
cmra_morphism_ne
:
>
NonExpansive
f
;
cmra_morphism_validN
n
x
:
✓
{
n
}
x
→
✓
{
n
}
f
x
;
cmra_morphism_pcore
x
:
pcore
(
f
x
)
≡
f
<$>
pcore
x
;
cmra_morphism_pcore
x
:
f
<$>
pcore
x
≡
pcore
(
f
x
)
;
cmra_morphism_op
x
y
:
f
(
x
⋅
y
)
≡
f
x
⋅
f
y
}.
Arguments
cmra_morphism_validN
{
_
_
}
_
{
_
}
_
_
_
.
...
...
@@ -761,7 +761,7 @@ Proof.
split
.
-
apply
_
.
-
move
=>
n
x
Hx
/=.
by
apply
cmra_morphism_validN
,
cmra_morphism_validN
.
-
move
=>
x
/=.
by
rewrite
2
!
cmra_morphism_pcore
option_fmap_compos
e
.
-
move
=>
x
/=.
by
rewrite
option_fmap_compose
!
cmra_morphism_pcor
e
.
-
move
=>
x
y
/=.
by
rewrite
!
cmra_morphism_op
.
Qed
.
...
...
@@ -769,7 +769,7 @@ Section cmra_morphism.
Local
Set
Default
Proof
Using
"Type*"
.
Context
{
A
B
:
cmraT
}
(
f
:
A
→
B
)
`
{!
CmraMorphism
f
}.
Lemma
cmra_morphism_core
x
:
core
(
f
x
)
≡
f
(
core
x
).
Proof
.
unfold
core
,
core'
.
rewrite
cmra_morphism_pcore
.
by
destruct
(
pcore
x
).
Qed
.
Proof
.
unfold
core
,
core'
.
rewrite
-
cmra_morphism_pcore
.
by
destruct
(
pcore
x
).
Qed
.
Lemma
cmra_morphism_monotone
x
y
:
x
≼
y
→
f
x
≼
f
y
.
Proof
.
intros
[
z
->].
exists
(
f
z
).
by
rewrite
cmra_morphism_op
.
Qed
.
Lemma
cmra_morphism_monotoneN
n
x
y
:
x
≼
{
n
}
y
→
f
x
≼
{
n
}
f
y
.
...
...
@@ -1207,8 +1207,8 @@ Instance prod_map_cmra_morphism {A A' B B' : cmraT} (f : A → A') (g : B → B'
Proof
.
split
;
first
apply
_
.
-
by
intros
n
x
[??]
;
split
;
simpl
;
apply
cmra_morphism_validN
.
-
intros
x
.
etrans
.
apply
(
reflexivity
(
mbind
_
_
)).
etrans
;
la
st
apply
(
reflexivity
(
_
<$>
mbind
_
_
)).
simpl
.
-
intros
x
.
etrans
;
last
apply
(
reflexivity
(
mbind
_
_
)).
etrans
;
fir
st
apply
(
reflexivity
(
_
<$>
mbind
_
_
)).
simpl
.
assert
(
Hf
:
=
cmra_morphism_pcore
f
(
x
.
1
)).
destruct
(
pcore
(
f
(
x
.
1
))),
(
pcore
(
x
.
1
))
;
inversion_clear
Hf
=>//=.
assert
(
Hg
:
=
cmra_morphism_pcore
g
(
x
.
2
)).
...
...
@@ -1597,7 +1597,7 @@ Instance discrete_fun_map_cmra_morphism {A} {B1 B2 : A → ucmraT} (f : ∀ x, B
Proof
.
split
;
first
apply
_
.
-
intros
n
g
Hg
x
;
rewrite
/
discrete_fun_map
;
apply
(
cmra_morphism_validN
(
f
_
)),
Hg
.
-
intros
.
apply
Some_proper
=>
i
.
apply
(
cmra_morphism_core
(
f
i
)).
-
intros
.
apply
Some_proper
=>
i
.
by
rewrite
(
cmra_morphism_core
(
f
i
)).
-
intros
g1
g2
i
.
by
rewrite
/
discrete_fun_map
discrete_fun_lookup_op
cmra_morphism_op
.
Qed
.
...
...
theories/algebra/csum.v
View file @
e3dd5307
...
...
@@ -366,7 +366,7 @@ Instance csum_map_cmra_morphism {A A' B B' : cmraT} (f : A → A') (g : B → B'
Proof
.
split
;
try
apply
_
.
-
intros
n
[
a
|
b
|]
;
simpl
;
auto
using
cmra_morphism_validN
.
-
move
=>
[
a
|
b
|]=>//=
;
rewrite
cmra_morphism_pcore
;
by
destruct
pcore
.
-
move
=>
[
a
|
b
|]=>//=
;
rewrite
-
cmra_morphism_pcore
;
by
destruct
pcore
.
-
intros
[
xa
|
ya
|]
[
xb
|
yb
|]=>//=
;
by
rewrite
cmra_morphism_op
.
Qed
.
...
...
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