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
c301e52e
Verified
Commit
c301e52e
authored
Aug 13, 2019
by
Paolo G. Giarrusso
Browse files
Flip cmra_morphism_core
parent
e3dd5307
Changes
2
Hide whitespace changes
Inline
Side-by-side
CHANGELOG.md
View file @
c301e52e
...
...
@@ -162,7 +162,8 @@ 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`
,
`cmra_morphism_pcore`
.
`Cinl_op`
,
`Cinr_op`
,
`cmra_morphism_op`
,
`cmra_morphism_pcore`
,
`cmra_morphism_core`
.
*
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 @
c301e52e
...
...
@@ -768,7 +768,7 @@ Qed.
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
).
Lemma
cmra_morphism_core
x
:
f
(
core
x
)
≡
core
(
f
x
).
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
.
...
...
@@ -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
.
by
rewrite
(
cmra_morphism_core
(
f
i
)).
-
intros
.
apply
Some_proper
=>
i
.
apply
(
cmra_morphism_core
(
f
i
)).
-
intros
g1
g2
i
.
by
rewrite
/
discrete_fun_map
discrete_fun_lookup_op
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