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
Dan Frumin
iris-coq
Commits
732b3d49
Commit
732b3d49
authored
Mar 08, 2016
by
Ralf Jung
Browse files
rename CMRA identity -> CMRA unit
parent
c2c84732
Changes
16
Hide whitespace changes
Inline
Side-by-side
algebra/auth.v
View file @
732b3d49
...
...
@@ -156,12 +156,12 @@ Proof. uPred.unseal. by destruct x as [[]]. Qed.
(
**
The
notations
◯
and
●
only
work
for
CMRAs
with
an
empty
element
.
So
,
in
what
follows
,
we
assume
we
have
an
empty
element
.
*
)
Context
`
{
Empty
A
,
!
CMRA
Ident
it
y
A
}
.
Context
`
{
Empty
A
,
!
CMRA
Un
it
A
}
.
Global
Instance
auth_cmra_
ident
it
y
:
CMRA
Ident
it
y
authR
.
Global
Instance
auth_cmra_
un
it
:
CMRA
Un
it
authR
.
Proof
.
split
;
simpl
.
-
by
apply
(
@
cmra_
empty
_valid
A
_
).
-
by
apply
(
@
cmra_
unit
_valid
A
_
).
-
by
intros
x
;
constructor
;
rewrite
/=
left_id
.
-
apply
_.
Qed
.
...
...
algebra/cmra.v
View file @
732b3d49
...
...
@@ -124,15 +124,15 @@ Section cmra_mixin.
Proof
.
apply
(
mixin_cmra_extend
_
(
cmra_mixin
A
)).
Qed
.
End
cmra_mixin
.
(
**
*
CMRAs
with
a
global
ident
it
y
element
*
)
(
**
*
CMRAs
with
a
un
it
element
*
)
(
**
We
use
the
notation
∅
because
for
most
instances
(
maps
,
sets
,
etc
)
the
`empty
'
element
is
the
global
ident
it
y
.
*
)
Class
CMRA
Ident
it
y
(
A
:
cmraT
)
`
{
Empty
A
}
:=
{
cmra_
empty
_valid
:
✓
∅
;
cmra_
empty
_left_id
:>
LeftId
(
≡
)
∅
(
⋅
);
cmra_
empty
_timeless
:>
Timeless
∅
`empty
'
element
is
the
un
it
.
*
)
Class
CMRA
Un
it
(
A
:
cmraT
)
`
{
Empty
A
}
:=
{
cmra_
unit
_valid
:
✓
∅
;
cmra_
unit
_left_id
:>
LeftId
(
≡
)
∅
(
⋅
);
cmra_
unit
_timeless
:>
Timeless
∅
}
.
Instance
cmra_
ident
it
y
_inhabited
`
{
CMRA
Ident
it
y
A
}
:
Inhabited
A
:=
populate
∅
.
Instance
cmra_
un
it_inhabited
`
{
CMRA
Un
it
A
}
:
Inhabited
A
:=
populate
∅
.
(
**
*
Discrete
CMRAs
*
)
Class
CMRADiscrete
(
A
:
cmraT
)
:=
{
...
...
@@ -347,20 +347,20 @@ Lemma cmra_discrete_update `{CMRADiscrete A} (x y : A) :
(
∀
z
,
✓
(
x
⋅
z
)
→
✓
(
y
⋅
z
))
→
x
~~>
y
.
Proof
.
intros
?
n
.
by
setoid_rewrite
<-
cmra_discrete_valid_iff
.
Qed
.
(
**
**
RAs
with
a
n
empty
element
*
)
Section
ident
it
y
.
Context
`
{
Empty
A
,
!
CMRA
Ident
it
y
A
}
.
Lemma
cmra_
empty
_validN
n
:
✓
{
n
}
∅
.
Proof
.
apply
cmra_valid_validN
,
cmra_
empty
_valid
.
Qed
.
Lemma
cmra_
empty
_leastN
n
x
:
∅
≼
{
n
}
x
.
(
**
**
RAs
with
a
unit
element
*
)
Section
un
it
.
Context
`
{
Empty
A
,
!
CMRA
Un
it
A
}
.
Lemma
cmra_
unit
_validN
n
:
✓
{
n
}
∅
.
Proof
.
apply
cmra_valid_validN
,
cmra_
unit
_valid
.
Qed
.
Lemma
cmra_
unit
_leastN
n
x
:
∅
≼
{
n
}
x
.
Proof
.
by
exists
x
;
rewrite
left_id
.
Qed
.
Lemma
cmra_
empty
_least
x
:
∅
≼
x
.
Lemma
cmra_
unit
_least
x
:
∅
≼
x
.
Proof
.
by
exists
x
;
rewrite
left_id
.
Qed
.
Global
Instance
cmra_
empty
_right_id
:
RightId
(
≡
)
∅
(
⋅
).
Global
Instance
cmra_
unit
_right_id
:
RightId
(
≡
)
∅
(
⋅
).
Proof
.
by
intros
x
;
rewrite
(
comm
op
)
left_id
.
Qed
.
Lemma
cmra_core_
empty
:
core
∅
≡
∅
.
Lemma
cmra_core_
unit
:
core
∅
≡
∅
.
Proof
.
by
rewrite
-{
2
}
(
cmra_core_l
∅
)
right_id
.
Qed
.
End
ident
it
y
.
End
un
it
.
(
**
**
Local
updates
*
)
Global
Instance
local_update_proper
Lv
(
L
:
A
→
A
)
:
...
...
@@ -422,13 +422,13 @@ Qed.
Lemma
cmra_update_id
x
:
x
~~>
x
.
Proof
.
intro
.
auto
.
Qed
.
Section
ident
it
y
_updates
.
Context
`
{
Empty
A
,
!
CMRA
Ident
it
y
A
}
.
Lemma
cmra_update_
empty
x
:
x
~~>
∅
.
Section
un
it_updates
.
Context
`
{
Empty
A
,
!
CMRA
Un
it
A
}
.
Lemma
cmra_update_
unit
x
:
x
~~>
∅
.
Proof
.
intros
n
z
;
rewrite
left_id
;
apply
cmra_validN_op_r
.
Qed
.
Lemma
cmra_update_
empty
_alt
y
:
∅
~~>
y
↔
∀
x
,
x
~~>
y
.
Proof
.
split
;
[
intros
;
trans
∅
|
];
auto
using
cmra_update_
empty
.
Qed
.
End
ident
it
y
_updates
.
Lemma
cmra_update_
unit
_alt
y
:
∅
~~>
y
↔
∀
x
,
x
~~>
y
.
Proof
.
split
;
[
intros
;
trans
∅
|
];
auto
using
cmra_update_
unit
.
Qed
.
End
un
it_updates
.
End
cmra
.
(
**
*
Properties
about
monotone
functions
*
)
...
...
@@ -531,7 +531,7 @@ Section unit.
Proof
.
by
split
.
Qed
.
Canonical
Structure
unitR
:
cmraT
:=
Eval
cbv
[
unitC
discreteR
cofe_car
]
in
discreteR
unit_ra
.
Global
Instance
unit_cmra_
ident
it
y
:
CMRA
Ident
it
y
unitR
.
Global
Instance
unit_cmra_
un
it
:
CMRA
Un
it
unitR
.
Global
Instance
unit_cmra_discrete
:
CMRADiscrete
unitR
.
Proof
.
by
apply
discrete_cmra_discrete
.
Qed
.
End
unit
.
...
...
@@ -582,11 +582,11 @@ Section prod.
by
exists
((
z1
.1
,
z2
.1
),(
z1
.2
,
z2
.2
)).
Qed
.
Canonical
Structure
prodR
:
cmraT
:=
CMRAT
prod_cofe_mixin
prod_cmra_mixin
.
Global
Instance
prod_cmra_
ident
it
y
`
{
Empty
A
,
Empty
B
}
:
CMRA
Ident
it
y
A
→
CMRA
Ident
it
y
B
→
CMRA
Ident
it
y
prodR
.
Global
Instance
prod_cmra_
un
it
`
{
Empty
A
,
Empty
B
}
:
CMRA
Un
it
A
→
CMRA
Un
it
B
→
CMRA
Un
it
prodR
.
Proof
.
split
.
-
split
;
apply
cmra_
empty
_valid
.
-
split
;
apply
cmra_
unit
_valid
.
-
by
split
;
rewrite
/=
left_id
.
-
by
intros
?
[
??
];
split
;
apply
(
timeless
_
).
Qed
.
...
...
algebra/cmra_big_op.v
View file @
732b3d49
...
...
@@ -11,7 +11,7 @@ Instance: Params (@big_opM) 5.
(
**
*
Properties
about
big
ops
*
)
Section
big_op
.
Context
`
{
CMRA
Ident
it
y
A
}
.
Context
`
{
CMRA
Un
it
A
}
.
(
**
*
Big
ops
*
)
Lemma
big_op_nil
:
big_op
(
@
nil
A
)
=
∅
.
...
...
algebra/cmra_tactics.v
View file @
732b3d49
...
...
@@ -3,7 +3,7 @@ From algebra Require Import cmra_big_op.
(
**
*
Simple
solver
for
validity
and
inclusion
by
reflection
*
)
Module
ra_reflection
.
Section
ra_reflection
.
Context
`
{
CMRA
Ident
it
y
A
}
.
Context
`
{
CMRA
Un
it
A
}
.
Inductive
expr
:=
|
EVar
:
nat
→
expr
...
...
algebra/excl.v
View file @
732b3d49
...
...
@@ -128,7 +128,7 @@ Proof.
end
;
destruct
y1
,
y2
;
inversion_clear
Hx
;
repeat
constructor
.
Qed
.
Canonical
Structure
exclR
:
cmraT
:=
CMRAT
excl_cofe_mixin
excl_cmra_mixin
.
Global
Instance
excl_cmra_
ident
it
y
:
CMRA
Ident
it
y
exclR
.
Global
Instance
excl_cmra_
un
it
:
CMRA
Un
it
exclR
.
Proof
.
split
.
done
.
by
intros
[].
apply
_.
Qed
.
Global
Instance
excl_cmra_discrete
:
Discrete
A
→
CMRADiscrete
exclR
.
Proof
.
split
.
apply
_.
by
intros
[].
Qed
.
...
...
algebra/fin_maps.v
View file @
732b3d49
...
...
@@ -159,7 +159,7 @@ Proof.
by
symmetry
;
apply
option_op_positive_dist_r
with
(
m1
!!
i
).
Qed
.
Canonical
Structure
mapR
:
cmraT
:=
CMRAT
map_cofe_mixin
map_cmra_mixin
.
Global
Instance
map_cmra_
ident
it
y
:
CMRA
Ident
it
y
mapR
.
Global
Instance
map_cmra_
un
it
:
CMRA
Un
it
mapR
.
Proof
.
split
.
-
by
intros
i
;
rewrite
lookup_empty
.
...
...
@@ -194,7 +194,7 @@ Lemma map_insert_valid m i x : ✓ x → ✓ m → ✓ <[i:=x]>m.
Proof
.
by
intros
??
j
;
destruct
(
decide
(
i
=
j
));
simplify_map_eq
.
Qed
.
Lemma
map_singleton_validN
n
i
x
:
✓
{
n
}
(
{
[
i
:=
x
]
}
:
gmap
K
A
)
↔
✓
{
n
}
x
.
Proof
.
split
;
[
|
by
intros
;
apply
map_insert_validN
,
cmra_
empty
_validN
].
split
;
[
|
by
intros
;
apply
map_insert_validN
,
cmra_
unit
_validN
].
by
move
=>/
(
_
i
);
simplify_map_eq
.
Qed
.
Lemma
map_singleton_valid
i
x
:
✓
(
{
[
i
:=
x
]
}
:
gmap
K
A
)
↔
✓
x
.
...
...
@@ -232,7 +232,7 @@ Proof.
-
intros
(
y
&
Hi
&?
);
rewrite
map_includedN_spec
=>
j
.
destruct
(
decide
(
i
=
j
));
simplify_map_eq
.
+
rewrite
Hi
.
by
apply
(
includedN_preserving
_
),
cmra_included_includedN
.
+
apply
:
cmra_
empty
_leastN
.
+
apply
:
cmra_
unit
_leastN
.
Qed
.
Lemma
map_dom_op
m1
m2
:
dom
(
gset
K
)
(
m1
⋅
m2
)
≡
dom
_
m1
∪
dom
_
m2
.
Proof
.
...
...
@@ -268,14 +268,14 @@ Proof. apply map_insert_updateP'. Qed.
Lemma
map_singleton_update
i
(
x
y
:
A
)
:
x
~~>
y
→
{
[
i
:=
x
]
}
~~>
{
[
i
:=
y
]
}
.
Proof
.
apply
map_insert_update
.
Qed
.
Lemma
map_singleton_updateP_empty
`
{
Empty
A
,
!
CMRA
Ident
it
y
A
}
Lemma
map_singleton_updateP_empty
`
{
Empty
A
,
!
CMRA
Un
it
A
}
(
P
:
A
→
Prop
)
(
Q
:
gmap
K
A
→
Prop
)
i
:
∅
~~>:
P
→
(
∀
y
,
P
y
→
Q
{
[
i
:=
y
]
}
)
→
∅
~~>:
Q
.
Proof
.
intros
Hx
HQ
n
gf
Hg
.
destruct
(
Hx
n
(
from_option
∅
(
gf
!!
i
)))
as
(
y
&?&
Hy
).
{
move
:
(
Hg
i
).
rewrite
!
left_id
.
case
_
:
(
gf
!!
i
);
simpl
;
auto
using
cmra_
empty
_validN
.
}
case
_
:
(
gf
!!
i
);
simpl
;
auto
using
cmra_
unit
_validN
.
}
exists
{
[
i
:=
y
]
}
;
split
;
first
by
auto
.
intros
i
'
;
destruct
(
decide
(
i
'
=
i
))
as
[
->|
].
-
rewrite
lookup_op
lookup_singleton
.
...
...
@@ -283,7 +283,7 @@ Proof.
by
rewrite
right_id
.
-
move
:
(
Hg
i
'
).
by
rewrite
!
lookup_op
lookup_singleton_ne
// !left_id.
Qed
.
Lemma
map_singleton_updateP_empty
'
`
{
Empty
A
,
!
CMRA
Ident
it
y
A
}
(
P
:
A
→
Prop
)
i
:
Lemma
map_singleton_updateP_empty
'
`
{
Empty
A
,
!
CMRA
Un
it
A
}
(
P
:
A
→
Prop
)
i
:
∅
~~>:
P
→
∅
~~>:
λ
m
,
∃
y
,
m
=
{
[
i
:=
y
]
}
∧
P
y
.
Proof
.
eauto
using
map_singleton_updateP_empty
.
Qed
.
...
...
algebra/frac.v
View file @
732b3d49
...
...
@@ -172,7 +172,7 @@ Proof.
+
exfalso
;
inversion_clear
Hx
'
.
Qed
.
Canonical
Structure
fracR
:
cmraT
:=
CMRAT
frac_cofe_mixin
frac_cmra_mixin
.
Global
Instance
frac_cmra_
ident
it
y
:
CMRA
Ident
it
y
fracR
.
Global
Instance
frac_cmra_
un
it
:
CMRA
Un
it
fracR
.
Proof
.
split
.
done
.
by
intros
[].
apply
_.
Qed
.
Global
Instance
frac_cmra_discrete
:
CMRADiscrete
A
→
CMRADiscrete
fracR
.
Proof
.
...
...
@@ -238,7 +238,7 @@ Proof.
split
;
try
apply
_.
-
intros
n
[
p
a
|
];
destruct
1
;
split
;
auto
using
validN_preserving
.
-
intros
[
q1
a1
|
]
[
q2
a2
|
]
[[
q3
a3
|
]
Hx
];
inversion
Hx
;
setoid_subst
;
try
apply
:
cmra_
empty
_least
;
auto
.
inversion
Hx
;
setoid_subst
;
try
apply
:
cmra_
unit
_least
;
auto
.
destruct
(
included_preserving
f
a1
(
a1
⋅
a3
))
as
[
b
?
].
{
by
apply
cmra_included_l
.
}
by
exists
(
Frac
q3
b
);
constructor
.
...
...
algebra/iprod.v
View file @
732b3d49
...
...
@@ -160,11 +160,11 @@ Section iprod_cmra.
split_and
?
;
intros
x
;
apply
(
proj2_sig
(
g
x
)).
Qed
.
Canonical
Structure
iprodR
:
cmraT
:=
CMRAT
iprod_cofe_mixin
iprod_cmra_mixin
.
Global
Instance
iprod_cmra_
ident
it
y
`
{
∀
x
,
Empty
(
B
x
)
}
:
(
∀
x
,
CMRA
Ident
it
y
(
B
x
))
→
CMRA
Ident
it
y
iprodR
.
Global
Instance
iprod_cmra_
un
it
`
{
∀
x
,
Empty
(
B
x
)
}
:
(
∀
x
,
CMRA
Un
it
(
B
x
))
→
CMRA
Un
it
iprodR
.
Proof
.
intros
?
;
split
.
-
intros
x
;
apply
cmra_
empty
_valid
.
-
intros
x
;
apply
cmra_
unit
_valid
.
-
by
intros
f
x
;
rewrite
iprod_lookup_op
left_id
.
-
by
apply
_.
Qed
.
...
...
@@ -201,14 +201,14 @@ Section iprod_cmra.
Qed
.
(
**
Properties
of
iprod_singleton
.
*
)
Context
`
{
∀
x
,
Empty
(
B
x
),
∀
x
,
CMRA
Ident
it
y
(
B
x
)
}
.
Context
`
{
∀
x
,
Empty
(
B
x
),
∀
x
,
CMRA
Un
it
(
B
x
)
}
.
Lemma
iprod_singleton_validN
n
x
(
y
:
B
x
)
:
✓
{
n
}
iprod_singleton
x
y
↔
✓
{
n
}
y
.
Proof
.
split
;
[
by
move
=>/
(
_
x
);
rewrite
iprod_lookup_singleton
|
].
move
=>
Hx
x
'
;
destruct
(
decide
(
x
=
x
'
))
as
[
->|
];
rewrite
?
iprod_lookup_singleton
?
iprod_lookup_singleton_ne
//.
by
apply
cmra_
empty
_validN
.
by
apply
cmra_
unit
_validN
.
Qed
.
Lemma
iprod_core_singleton
x
(
y
:
B
x
)
:
...
...
@@ -216,7 +216,7 @@ Section iprod_cmra.
Proof
.
by
move
=>
x
'
;
destruct
(
decide
(
x
=
x
'
))
as
[
->|
];
rewrite
iprod_lookup_core
?
iprod_lookup_singleton
?
iprod_lookup_singleton_ne
// cmra_core_
empty
.
?
iprod_lookup_singleton_ne
// cmra_core_
unit
.
Qed
.
Lemma
iprod_op_singleton
(
x
:
A
)
(
y1
y2
:
B
x
)
:
...
...
algebra/option.v
View file @
732b3d49
...
...
@@ -119,7 +119,7 @@ Proof.
+
exists
(
None
,
None
);
repeat
constructor
.
Qed
.
Canonical
Structure
optionR
:=
CMRAT
option_cofe_mixin
option_cmra_mixin
.
Global
Instance
option_cmra_
ident
it
y
:
CMRA
Ident
it
y
optionR
.
Global
Instance
option_cmra_
un
it
:
CMRA
Un
it
optionR
.
Proof
.
split
.
done
.
by
intros
[].
by
inversion_clear
1.
Qed
.
Global
Instance
option_cmra_discrete
:
CMRADiscrete
A
→
CMRADiscrete
optionR
.
Proof
.
split
;
[
apply
_
|
].
by
intros
[
x
|
];
[
apply
(
cmra_discrete_valid
x
)
|
].
Qed
.
...
...
@@ -164,10 +164,10 @@ Lemma option_update x y : x ~~> y → Some x ~~> Some y.
Proof
.
rewrite
!
cmra_update_updateP
;
eauto
using
option_updateP
with
congruence
.
Qed
.
Lemma
option_update_None
`
{
Empty
A
,
!
CMRA
Ident
it
y
A
}
:
∅
~~>
Some
∅
.
Lemma
option_update_None
`
{
Empty
A
,
!
CMRA
Un
it
A
}
:
∅
~~>
Some
∅
.
Proof
.
intros
n
[
x
|
]
?
;
rewrite
/
op
/
cmra_op
/
validN
/
cmra_validN
/=
?
left_id
;
auto
using
cmra_
empty
_validN
.
auto
using
cmra_
unit
_validN
.
Qed
.
End
cmra
.
Arguments
optionR
:
clear
implicits
.
...
...
algebra/upred.v
View file @
732b3d49
...
...
@@ -499,11 +499,11 @@ Proof.
unseal
;
intros
Hab
Ha
;
split
=>
n
x
??
.
apply
H
Ψ
with
n
a
;
auto
.
by
symmetry
;
apply
Hab
with
x
.
by
apply
Ha
.
Qed
.
Lemma
eq_equiv
`
{
Empty
M
,
!
CMRA
Ident
it
y
M
}
{
A
:
cofeT
}
(
a
b
:
A
)
:
Lemma
eq_equiv
`
{
Empty
M
,
!
CMRA
Un
it
M
}
{
A
:
cofeT
}
(
a
b
:
A
)
:
True
⊑
(
a
≡
b
)
→
a
≡
b
.
Proof
.
unseal
=>
Hab
;
apply
equiv_dist
;
intros
n
;
apply
Hab
with
∅
;
last
done
.
apply
cmra_valid_validN
,
cmra_
empty
_valid
.
apply
cmra_valid_validN
,
cmra_
unit
_valid
.
Qed
.
Lemma
iff_equiv
P
Q
:
True
⊑
(
P
↔
Q
)
→
P
≡
Q
.
Proof
.
...
...
@@ -1002,7 +1002,7 @@ Lemma always_ownM (a : M) : core a ≡ a → (□ uPred_ownM a)%I ≡ uPred_ownM
Proof
.
by
intros
<-
;
rewrite
always_ownM_core
.
Qed
.
Lemma
ownM_something
:
True
⊑
∃
a
,
uPred_ownM
a
.
Proof
.
unseal
;
split
=>
n
x
??
.
by
exists
x
;
simpl
.
Qed
.
Lemma
ownM_empty
`
{
Empty
M
,
!
CMRA
Ident
it
y
M
}
:
True
⊑
uPred_ownM
∅
.
Lemma
ownM_empty
`
{
Empty
M
,
!
CMRA
Un
it
M
}
:
True
⊑
uPred_ownM
∅
.
Proof
.
unseal
;
split
=>
n
x
??
;
by
exists
x
;
rewrite
left_id
.
Qed
.
(
*
Valid
*
)
...
...
program_logic/auth.v
View file @
732b3d49
...
...
@@ -5,14 +5,14 @@ Import uPred.
(
*
The
CMRA
we
need
.
*
)
Class
authG
Λ
Σ
(
A
:
cmraT
)
`
{
Empty
A
}
:=
AuthG
{
auth_inG
:>
inG
Λ
Σ
(
authR
A
);
auth_identity
:>
CMRA
Ident
it
y
A
;
auth_identity
:>
CMRA
Un
it
A
;
auth_timeless
:>
CMRADiscrete
A
;
}
.
(
*
The
Functor
we
need
.
*
)
Definition
authGF
(
A
:
cmraT
)
:
gFunctor
:=
GFunctor
(
constRF
(
authR
A
)).
(
*
Show
and
register
that
they
match
.
*
)
Instance
authGF_inGF
(
A
:
cmraT
)
`
{
inGF
Λ
Σ
(
authGF
A
)
}
`
{
CMRA
Ident
it
y
A
,
CMRADiscrete
A
}
:
authG
Λ
Σ
A
.
`
{
CMRA
Un
it
A
,
CMRADiscrete
A
}
:
authG
Λ
Σ
A
.
Proof
.
split
;
try
apply
_.
apply
:
inGF_inG
.
Qed
.
Section
definitions
.
...
...
program_logic/ghost_ownership.v
View file @
732b3d49
...
...
@@ -19,10 +19,10 @@ Implicit Types a : A.
Instance
inG_empty
`
{
Empty
A
}
:
Empty
(
Σ
inG_id
(
iPreProp
Λ
(
globalF
Σ
)))
:=
cmra_transport
inG_prf
∅
.
Instance
inG_empty_spec
`
{
Empty
A
}
:
CMRA
Ident
it
y
A
→
CMRA
Ident
it
y
(
Σ
inG_id
(
iPreProp
Λ
(
globalF
Σ
))).
CMRA
Un
it
A
→
CMRA
Un
it
(
Σ
inG_id
(
iPreProp
Λ
(
globalF
Σ
))).
Proof
.
split
.
-
apply
cmra_transport_valid
,
cmra_
empty
_valid
.
-
apply
cmra_transport_valid
,
cmra_
unit
_valid
.
-
intros
x
;
rewrite
/
empty
/
inG_empty
;
destruct
inG_prf
.
by
rewrite
left_id
.
-
apply
_.
Qed
.
...
...
@@ -52,7 +52,7 @@ Lemma own_valid_r γ a : own γ a ⊑ (own γ a ★ ✓ a).
Proof
.
apply
:
uPred
.
always_entails_r
.
apply
own_valid
.
Qed
.
Lemma
own_valid_l
γ
a
:
own
γ
a
⊑
(
✓
a
★
own
γ
a
).
Proof
.
by
rewrite
comm
-
own_valid_r
.
Qed
.
Lemma
own_empty
`
{
CMRA
Ident
it
y
A
}
γ
:
True
⊑
own
γ
∅
.
Lemma
own_empty
`
{
CMRA
Un
it
A
}
γ
:
True
⊑
own
γ
∅
.
Proof
.
rewrite
ownG_empty
/
own
.
apply
equiv_spec
,
ownG_proper
.
(
*
FIXME
:
rewrite
to_globalF_empty
.
*
)
...
...
@@ -99,13 +99,13 @@ Proof.
by
apply
pvs_mono
,
exist_elim
=>
a
''
;
apply
const_elim_l
=>
->
.
Qed
.
Lemma
own_empty
`
{
Empty
A
,
!
CMRA
Ident
it
y
A
}
γ
E
:
Lemma
own_empty
`
{
Empty
A
,
!
CMRA
Un
it
A
}
γ
E
:
True
⊑
(
|={
E
}=>
own
γ
∅
).
Proof
.
rewrite
ownG_empty
/
own
.
apply
pvs_ownG_update
,
cmra_update_updateP
.
eapply
iprod_singleton_updateP_empty
;
first
by
eapply
map_singleton_updateP_empty
'
,
cmra_transport_updateP
'
,
cmra_update_updateP
,
cmra_update_
empty
.
cmra_update_updateP
,
cmra_update_
unit
.
naive_solver
.
Qed
.
End
global
.
program_logic/model.v
View file @
732b3d49
...
...
@@ -10,7 +10,7 @@ Structure iFunctor := IFunctor {
iFunctor_F
:>
rFunctor
;
iFunctor_contractive
:
rFunctorContractive
iFunctor_F
;
iFunctor_empty
(
A
:
cofeT
)
:
Empty
(
iFunctor_F
A
);
iFunctor_identity
(
A
:
cofeT
)
:
CMRA
Ident
it
y
(
iFunctor_F
A
);
iFunctor_identity
(
A
:
cofeT
)
:
CMRA
Un
it
(
iFunctor_F
A
);
}
.
Arguments
IFunctor
_
{
_
_
_
}
.
Existing
Instances
iFunctor_contractive
iFunctor_empty
iFunctor_identity
.
...
...
program_logic/ownership.v
View file @
732b3d49
...
...
@@ -28,7 +28,7 @@ Qed.
Lemma
always_ownI
i
P
:
(
□
ownI
i
P
)
%
I
≡
ownI
i
P
.
Proof
.
apply
uPred
.
always_ownM
.
by
rewrite
Res_core
!
cmra_core_
empty
map_core_singleton
.
by
rewrite
Res_core
!
cmra_core_
unit
map_core_singleton
.
Qed
.
Global
Instance
ownI_always_stable
i
P
:
AlwaysStable
(
ownI
i
P
).
Proof
.
by
rewrite
/
AlwaysStable
always_ownI
.
Qed
.
...
...
@@ -55,7 +55,7 @@ Proof. move=>a b [c H]. rewrite H ownG_op. eauto with I. Qed.
Lemma
always_ownG_core
m
:
(
□
ownG
(
core
m
))
%
I
≡
ownG
(
core
m
).
Proof
.
apply
uPred
.
always_ownM
.
by
rewrite
Res_core
!
cmra_core_
empty
-{
2
}
(
cmra_core_idemp
m
).
by
rewrite
Res_core
!
cmra_core_
unit
-{
2
}
(
cmra_core_idemp
m
).
Qed
.
Lemma
always_ownG
m
:
core
m
≡
m
→
(
□
ownG
m
)
%
I
≡
ownG
m
.
Proof
.
by
intros
<-
;
rewrite
always_ownG_core
.
Qed
.
...
...
@@ -82,17 +82,17 @@ Proof.
-
intros
[(
P
'
&
Hi
&
HP
)
_
];
rewrite
Hi
.
by
apply
Some_dist
,
symmetry
,
agree_valid_includedN
,
(
cmra_included_includedN
_
P
'
),
HP
;
apply
map_lookup_validN
with
(
wld
r
)
i
.
-
intros
?
;
split_and
?
;
try
apply
cmra_
empty
_leastN
;
eauto
.
-
intros
?
;
split_and
?
;
try
apply
cmra_
unit
_leastN
;
eauto
.
Qed
.
Lemma
ownP_spec
n
r
σ
:
✓
{
n
}
r
→
(
ownP
σ
)
n
r
↔
pst
r
≡
Excl
σ
.
Proof
.
intros
(
?&?&?
).
rewrite
/
ownP
;
uPred
.
unseal
.
rewrite
/
uPred_holds
/=
res_includedN
/=
Excl_includedN
//.
rewrite
(
timeless_iff
n
).
naive_solver
(
apply
cmra_
empty
_leastN
).
rewrite
(
timeless_iff
n
).
naive_solver
(
apply
cmra_
unit
_leastN
).
Qed
.
Lemma
ownG_spec
n
r
m
:
(
ownG
m
)
n
r
↔
m
≼
{
n
}
gst
r
.
Proof
.
rewrite
/
ownG
;
uPred
.
unseal
.
rewrite
/
uPred_holds
/=
res_includedN
;
naive_solver
(
apply
cmra_
empty
_leastN
).
rewrite
/
uPred_holds
/=
res_includedN
;
naive_solver
(
apply
cmra_
unit
_leastN
).
Qed
.
End
ownership
.
program_logic/resources.v
View file @
732b3d49
...
...
@@ -124,10 +124,10 @@ Proof.
by
exists
(
Res
w
σ
m
,
Res
w
'
σ'
m
'
).
Qed
.
Canonical
Structure
resR
:
cmraT
:=
CMRAT
res_cofe_mixin
res_cmra_mixin
.
Global
Instance
res_cmra_
ident
it
y
`
{
CMRA
Ident
it
y
M
}
:
CMRA
Ident
it
y
resR
.
Global
Instance
res_cmra_
un
it
`
{
CMRA
Un
it
M
}
:
CMRA
Un
it
resR
.
Proof
.
split
.
-
split_and
!
;
apply
cmra_
empty
_valid
.
-
split_and
!
;
apply
cmra_
unit
_valid
.
-
by
split
;
rewrite
/=
left_id
.
-
apply
_.
Qed
.
...
...
program_logic/viewshifts.v
View file @
732b3d49
...
...
@@ -117,7 +117,7 @@ Proof. by intros; apply vs_alt, own_updateP. Qed.
Lemma
vs_update
E
γ
a
a
'
:
a
~~>
a
'
→
own
γ
a
={
E
}=>
own
γ
a
'
.
Proof
.
by
intros
;
apply
vs_alt
,
own_update
.
Qed
.
Lemma
vs_own_empty
`
{
Empty
A
,
!
CMRA
Ident
it
y
A
}
E
γ
:
Lemma
vs_own_empty
`
{
Empty
A
,
!
CMRA
Un
it
A
}
E
γ
:
True
={
E
}=>
own
γ
∅
.
Proof
.
by
intros
;
eapply
vs_alt
,
own_empty
.
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