Skip to content
GitLab
Menu
Projects
Groups
Snippets
Help
Help
Support
Community forum
Keyboard shortcuts
?
Submit feedback
Contribute to GitLab
Sign in / Register
Toggle navigation
Menu
Open sidebar
Tej Chajed
iris
Commits
d8a82b35
Commit
d8a82b35
authored
Jan 19, 2016
by
Ralf Jung
Browse files
show that we can construct an IParam from a single, constant cmra (not using the functor)
parent
c9ae33d4
Changes
1
Hide whitespace changes
Inline
Side-by-side
iris/parameter.v
View file @
d8a82b35
Require
Export
modures
.
cmra
iris
.
language
.
Set
Bullet
Behavior
"Strict Subproofs"
.
Record
iParam
:
=
IParam
{
iexpr
:
Type
;
ival
:
Type
;
istate
:
Type
;
icmra
:
cofeT
→
cmraT
;
ilanguage
:
Language
iexpr
ival
istate
;
icmra
:
cofeT
→
cmraT
;
icmra_empty
A
:
Empty
(
icmra
A
)
;
icmra_empty_spec
A
:
RAIdentity
(
icmra
A
)
;
icmra_map
{
A
B
}
(
f
:
A
-
n
>
B
)
:
icmra
A
-
n
>
icmra
B
;
...
...
@@ -15,6 +17,7 @@ Record iParam := IParam {
icmra_map
(
g
◎
f
)
x
≡
icmra_map
g
(
icmra_map
f
x
)
;
icmra_map_mono
{
A
B
}
(
f
:
A
-
n
>
B
)
:
CMRAMonotone
(
icmra_map
f
)
}.
Arguments
IParam
_
_
_
_
_
{
_
_
}
_
{
_
_
_
_
}.
Existing
Instances
ilanguage
.
Existing
Instances
icmra_empty
icmra_empty_spec
icmra_map_ne
icmra_map_mono
.
...
...
@@ -24,4 +27,16 @@ Proof.
by
intros
?
;
apply
equiv_dist
=>
n
;
apply
icmra_map_ne
=>
?
;
apply
equiv_dist
.
Qed
.
Canonical
Structure
istateC
Σ
:
=
leibnizC
(
istate
Σ
).
\ No newline at end of file
Definition
IParamConst
(
iexpr
ival
istate
:
Type
)
(
ilanguage
:
Language
iexpr
ival
istate
)
(
icmra
:
cmraT
)
{
icmra_empty
:
Empty
icmra
}
{
icmra_empty_spec
:
RAIdentity
icmra
}
:
iParam
.
eapply
(
IParam
iexpr
ival
istate
ilanguage
(
fun
_
=>
icmra
)
(
fun
_
_
_
=>
cid
)).
Unshelve
.
-
by
intros
.
-
by
intros
.
Defined
.
Canonical
Structure
istateC
Σ
:
=
leibnizC
(
istate
Σ
).
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