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
Marianna Rapoport
iris-coq
Commits
c2c84732
Commit
c2c84732
authored
Mar 08, 2016
by
Ralf Jung
Browse files
rename CMRA unit -> core
parent
361c9fbf
Changes
17
Hide whitespace changes
Inline
Side-by-side
algebra/agree.v
View file @
c2c84732
...
...
@@ -60,7 +60,7 @@ Program Instance agree_op : Op (agree A) := λ x y,
{|
agree_car
:
=
x
;
agree_is_valid
n
:
=
agree_is_valid
x
n
∧
agree_is_valid
y
n
∧
x
≡
{
n
}
≡
y
|}.
Next
Obligation
.
naive_solver
eauto
using
agree_valid_S
,
dist_S
.
Qed
.
Instance
agree_
unit
:
Unit
(
agree
A
)
:
=
id
.
Instance
agree_
core
:
Core
(
agree
A
)
:
=
id
.
Instance
agree_div
:
Div
(
agree
A
)
:
=
λ
x
y
,
x
.
Instance
:
Comm
(
≡
)
(@
op
(
agree
A
)
_
).
...
...
algebra/auth.v
View file @
c2c84732
...
...
@@ -85,8 +85,8 @@ Instance auth_validN : ValidN (auth A) := λ n x,
|
ExclBot
=>
False
end
.
Global
Arguments
auth_validN
_
!
_
/.
Instance
auth_
unit
:
Unit
(
auth
A
)
:
=
λ
x
,
Auth
(
unit
(
authoritative
x
))
(
unit
(
own
x
)).
Instance
auth_
core
:
Core
(
auth
A
)
:
=
λ
x
,
Auth
(
core
(
authoritative
x
))
(
core
(
own
x
)).
Instance
auth_op
:
Op
(
auth
A
)
:
=
λ
x
y
,
Auth
(
authoritative
x
⋅
authoritative
y
)
(
own
x
⋅
own
y
).
Instance
auth_div
:
Div
(
auth
A
)
:
=
λ
x
y
,
...
...
@@ -117,10 +117,10 @@ Proof.
-
intros
n
[[]
?]
?
;
naive_solver
eauto
using
cmra_includedN_S
,
cmra_validN_S
.
-
by
split
;
simpl
;
rewrite
assoc
.
-
by
split
;
simpl
;
rewrite
comm
.
-
by
split
;
simpl
;
rewrite
?cmra_
unit
_l
.
-
by
split
;
simpl
;
rewrite
?cmra_
unit
_idemp
.
-
by
split
;
simpl
;
rewrite
?cmra_
core
_l
.
-
by
split
;
simpl
;
rewrite
?cmra_
core
_idemp
.
-
intros
??
;
rewrite
!
auth_included
;
intros
[??].
by
split
;
simpl
;
apply
cmra_
unit
_preserving
.
by
split
;
simpl
;
apply
cmra_
core
_preserving
.
-
assert
(
∀
n
(
a
b1
b2
:
A
),
b1
⋅
b2
≼
{
n
}
a
→
b1
≼
{
n
}
a
).
{
intros
n
a
b1
b2
<-
;
apply
cmra_includedN_l
.
}
intros
n
[[
a1
|
|]
b1
]
[[
a2
|
|]
b2
]
;
...
...
algebra/cmra.v
View file @
c2c84732
From
algebra
Require
Export
cofe
.
Class
Unit
(
A
:
Type
)
:
=
unit
:
A
→
A
.
Instance
:
Params
(@
unit
)
2
.
Class
Core
(
A
:
Type
)
:
=
core
:
A
→
A
.
Instance
:
Params
(@
core
)
2
.
Class
Op
(
A
:
Type
)
:
=
op
:
A
→
A
→
A
.
Instance
:
Params
(@
op
)
2
.
...
...
@@ -34,10 +34,10 @@ Instance: Params (@includedN) 4.
Hint
Extern
0
(
_
≼
{
_
}
_
)
=>
reflexivity
.
Record
CMRAMixin
A
`
{
Dist
A
,
Equiv
A
,
Unit
A
,
Op
A
,
Valid
A
,
ValidN
A
,
Div
A
}
:
=
{
`
{
Dist
A
,
Equiv
A
,
Core
A
,
Op
A
,
Valid
A
,
ValidN
A
,
Div
A
}
:
=
{
(* setoids *)
mixin_cmra_op_ne
n
(
x
:
A
)
:
Proper
(
dist
n
==>
dist
n
)
(
op
x
)
;
mixin_cmra_
unit
_ne
n
:
Proper
(
dist
n
==>
dist
n
)
unit
;
mixin_cmra_
core
_ne
n
:
Proper
(
dist
n
==>
dist
n
)
core
;
mixin_cmra_validN_ne
n
:
Proper
(
dist
n
==>
impl
)
(
validN
n
)
;
mixin_cmra_div_ne
n
:
Proper
(
dist
n
==>
dist
n
==>
dist
n
)
div
;
(* valid *)
...
...
@@ -46,9 +46,9 @@ Record CMRAMixin A
(* monoid *)
mixin_cmra_assoc
:
Assoc
(
≡
)
(
⋅
)
;
mixin_cmra_comm
:
Comm
(
≡
)
(
⋅
)
;
mixin_cmra_
unit
_l
x
:
unit
x
⋅
x
≡
x
;
mixin_cmra_
unit
_idemp
x
:
unit
(
unit
x
)
≡
unit
x
;
mixin_cmra_
unit
_preserving
x
y
:
x
≼
y
→
unit
x
≼
unit
y
;
mixin_cmra_
core
_l
x
:
core
x
⋅
x
≡
x
;
mixin_cmra_
core
_idemp
x
:
core
(
core
x
)
≡
core
x
;
mixin_cmra_
core
_preserving
x
y
:
x
≼
y
→
core
x
≼
core
y
;
mixin_cmra_validN_op_l
n
x
y
:
✓
{
n
}
(
x
⋅
y
)
→
✓
{
n
}
x
;
mixin_cmra_op_div
x
y
:
x
≼
y
→
x
⋅
y
÷
x
≡
y
;
mixin_cmra_extend
n
x
y1
y2
:
...
...
@@ -62,7 +62,7 @@ Structure cmraT := CMRAT {
cmra_equiv
:
Equiv
cmra_car
;
cmra_dist
:
Dist
cmra_car
;
cmra_compl
:
Compl
cmra_car
;
cmra_
unit
:
Unit
cmra_car
;
cmra_
core
:
Core
cmra_car
;
cmra_op
:
Op
cmra_car
;
cmra_valid
:
Valid
cmra_car
;
cmra_validN
:
ValidN
cmra_car
;
...
...
@@ -75,7 +75,7 @@ Arguments cmra_car : simpl never.
Arguments
cmra_equiv
:
simpl
never
.
Arguments
cmra_dist
:
simpl
never
.
Arguments
cmra_compl
:
simpl
never
.
Arguments
cmra_
unit
:
simpl
never
.
Arguments
cmra_
core
:
simpl
never
.
Arguments
cmra_op
:
simpl
never
.
Arguments
cmra_valid
:
simpl
never
.
Arguments
cmra_validN
:
simpl
never
.
...
...
@@ -83,7 +83,7 @@ Arguments cmra_div : simpl never.
Arguments
cmra_cofe_mixin
:
simpl
never
.
Arguments
cmra_mixin
:
simpl
never
.
Add
Printing
Constructor
cmraT
.
Existing
Instances
cmra_
unit
cmra_op
cmra_valid
cmra_validN
cmra_div
.
Existing
Instances
cmra_
core
cmra_op
cmra_valid
cmra_validN
cmra_div
.
Coercion
cmra_cofeC
(
A
:
cmraT
)
:
cofeT
:
=
CofeT
(
cmra_cofe_mixin
A
).
Canonical
Structure
cmra_cofeC
.
...
...
@@ -93,8 +93,8 @@ Section cmra_mixin.
Implicit
Types
x
y
:
A
.
Global
Instance
cmra_op_ne
n
(
x
:
A
)
:
Proper
(
dist
n
==>
dist
n
)
(
op
x
).
Proof
.
apply
(
mixin_cmra_op_ne
_
(
cmra_mixin
A
)).
Qed
.
Global
Instance
cmra_
unit
_ne
n
:
Proper
(
dist
n
==>
dist
n
)
(@
unit
A
_
).
Proof
.
apply
(
mixin_cmra_
unit
_ne
_
(
cmra_mixin
A
)).
Qed
.
Global
Instance
cmra_
core
_ne
n
:
Proper
(
dist
n
==>
dist
n
)
(@
core
A
_
).
Proof
.
apply
(
mixin_cmra_
core
_ne
_
(
cmra_mixin
A
)).
Qed
.
Global
Instance
cmra_validN_ne
n
:
Proper
(
dist
n
==>
impl
)
(@
validN
A
_
n
).
Proof
.
apply
(
mixin_cmra_validN_ne
_
(
cmra_mixin
A
)).
Qed
.
Global
Instance
cmra_div_ne
n
:
...
...
@@ -108,12 +108,12 @@ Section cmra_mixin.
Proof
.
apply
(
mixin_cmra_assoc
_
(
cmra_mixin
A
)).
Qed
.
Global
Instance
cmra_comm
:
Comm
(
≡
)
(@
op
A
_
).
Proof
.
apply
(
mixin_cmra_comm
_
(
cmra_mixin
A
)).
Qed
.
Lemma
cmra_
unit
_l
x
:
unit
x
⋅
x
≡
x
.
Proof
.
apply
(
mixin_cmra_
unit
_l
_
(
cmra_mixin
A
)).
Qed
.
Lemma
cmra_
unit
_idemp
x
:
unit
(
unit
x
)
≡
unit
x
.
Proof
.
apply
(
mixin_cmra_
unit
_idemp
_
(
cmra_mixin
A
)).
Qed
.
Lemma
cmra_
unit
_preserving
x
y
:
x
≼
y
→
unit
x
≼
unit
y
.
Proof
.
apply
(
mixin_cmra_
unit
_preserving
_
(
cmra_mixin
A
)).
Qed
.
Lemma
cmra_
core
_l
x
:
core
x
⋅
x
≡
x
.
Proof
.
apply
(
mixin_cmra_
core
_l
_
(
cmra_mixin
A
)).
Qed
.
Lemma
cmra_
core
_idemp
x
:
core
(
core
x
)
≡
core
x
.
Proof
.
apply
(
mixin_cmra_
core
_idemp
_
(
cmra_mixin
A
)).
Qed
.
Lemma
cmra_
core
_preserving
x
y
:
x
≼
y
→
core
x
≼
core
y
.
Proof
.
apply
(
mixin_cmra_
core
_preserving
_
(
cmra_mixin
A
)).
Qed
.
Lemma
cmra_validN_op_l
n
x
y
:
✓
{
n
}
(
x
⋅
y
)
→
✓
{
n
}
x
.
Proof
.
apply
(
mixin_cmra_validN_op_l
_
(
cmra_mixin
A
)).
Qed
.
Lemma
cmra_op_div
x
y
:
x
≼
y
→
x
⋅
y
÷
x
≡
y
.
...
...
@@ -175,7 +175,7 @@ Implicit Types x y z : A.
Implicit
Types
xs
ys
zs
:
list
A
.
(** ** Setoids *)
Global
Instance
cmra_
unit
_proper
:
Proper
((
≡
)
==>
(
≡
))
(@
unit
A
_
).
Global
Instance
cmra_
core
_proper
:
Proper
((
≡
)
==>
(
≡
))
(@
core
A
_
).
Proof
.
apply
(
ne_proper
_
).
Qed
.
Global
Instance
cmra_op_ne'
n
:
Proper
(
dist
n
==>
dist
n
==>
dist
n
)
(@
op
A
_
).
Proof
.
...
...
@@ -236,15 +236,15 @@ Proof. rewrite (comm _ x); apply cmra_validN_op_l. Qed.
Lemma
cmra_valid_op_r
x
y
:
✓
(
x
⋅
y
)
→
✓
y
.
Proof
.
rewrite
!
cmra_valid_validN
;
eauto
using
cmra_validN_op_r
.
Qed
.
(** **
Units
*)
Lemma
cmra_
unit
_r
x
:
x
⋅
unit
x
≡
x
.
Proof
.
by
rewrite
(
comm
_
x
)
cmra_
unit
_l
.
Qed
.
Lemma
cmra_
unit_unit
x
:
unit
x
⋅
unit
x
≡
unit
x
.
Proof
.
by
rewrite
-{
2
}(
cmra_
unit
_idemp
x
)
cmra_
unit
_r
.
Qed
.
Lemma
cmra_
unit
_validN
n
x
:
✓
{
n
}
x
→
✓
{
n
}
unit
x
.
Proof
.
rewrite
-{
1
}(
cmra_
unit
_l
x
)
;
apply
cmra_validN_op_l
.
Qed
.
Lemma
cmra_
unit
_valid
x
:
✓
x
→
✓
unit
x
.
Proof
.
rewrite
-{
1
}(
cmra_
unit
_l
x
)
;
apply
cmra_valid_op_l
.
Qed
.
(** **
Core
*)
Lemma
cmra_
core
_r
x
:
x
⋅
core
x
≡
x
.
Proof
.
by
rewrite
(
comm
_
x
)
cmra_
core
_l
.
Qed
.
Lemma
cmra_
core_core
x
:
core
x
⋅
core
x
≡
core
x
.
Proof
.
by
rewrite
-{
2
}(
cmra_
core
_idemp
x
)
cmra_
core
_r
.
Qed
.
Lemma
cmra_
core
_validN
n
x
:
✓
{
n
}
x
→
✓
{
n
}
core
x
.
Proof
.
rewrite
-{
1
}(
cmra_
core
_l
x
)
;
apply
cmra_validN_op_l
.
Qed
.
Lemma
cmra_
core
_valid
x
:
✓
x
→
✓
core
x
.
Proof
.
rewrite
-{
1
}(
cmra_
core
_l
x
)
;
apply
cmra_valid_op_l
.
Qed
.
(** ** Div *)
Lemma
cmra_op_div'
n
x
y
:
x
≼
{
n
}
y
→
x
⋅
y
÷
x
≡
{
n
}
≡
y
.
...
...
@@ -260,7 +260,7 @@ Qed.
Global
Instance
cmra_includedN_preorder
n
:
PreOrder
(@
includedN
A
_
_
n
).
Proof
.
split
.
-
by
intros
x
;
exists
(
unit
x
)
;
rewrite
cmra_
unit
_r
.
-
by
intros
x
;
exists
(
core
x
)
;
rewrite
cmra_
core
_r
.
-
intros
x
y
z
[
z1
Hy
]
[
z2
Hz
]
;
exists
(
z1
⋅
z2
).
by
rewrite
assoc
-
Hy
-
Hz
.
Qed
.
...
...
@@ -288,13 +288,13 @@ Proof. rewrite (comm op); apply cmra_includedN_l. Qed.
Lemma
cmra_included_r
x
y
:
y
≼
x
⋅
y
.
Proof
.
rewrite
(
comm
op
)
;
apply
cmra_included_l
.
Qed
.
Lemma
cmra_
unit
_preservingN
n
x
y
:
x
≼
{
n
}
y
→
unit
x
≼
{
n
}
unit
y
.
Lemma
cmra_
core
_preservingN
n
x
y
:
x
≼
{
n
}
y
→
core
x
≼
{
n
}
core
y
.
Proof
.
intros
[
z
->].
apply
cmra_included_includedN
,
cmra_
unit
_preserving
,
cmra_included_l
.
apply
cmra_included_includedN
,
cmra_
core
_preserving
,
cmra_included_l
.
Qed
.
Lemma
cmra_included_
unit
x
:
unit
x
≼
x
.
Proof
.
by
exists
x
;
rewrite
cmra_
unit
_l
.
Qed
.
Lemma
cmra_included_
core
x
:
core
x
≼
x
.
Proof
.
by
exists
x
;
rewrite
cmra_
core
_l
.
Qed
.
Lemma
cmra_preservingN_l
n
x
y
z
:
x
≼
{
n
}
y
→
z
⋅
x
≼
{
n
}
z
⋅
y
.
Proof
.
by
intros
[
z1
Hz1
]
;
exists
z1
;
rewrite
Hz1
(
assoc
op
).
Qed
.
Lemma
cmra_preserving_l
x
y
z
:
x
≼
y
→
z
⋅
x
≼
z
⋅
y
.
...
...
@@ -358,8 +358,8 @@ Section identity.
Proof
.
by
exists
x
;
rewrite
left_id
.
Qed
.
Global
Instance
cmra_empty_right_id
:
RightId
(
≡
)
∅
(
⋅
).
Proof
.
by
intros
x
;
rewrite
(
comm
op
)
left_id
.
Qed
.
Lemma
cmra_
unit
_empty
:
unit
∅
≡
∅
.
Proof
.
by
rewrite
-{
2
}(
cmra_
unit
_l
∅
)
right_id
.
Qed
.
Lemma
cmra_
core
_empty
:
core
∅
≡
∅
.
Proof
.
by
rewrite
-{
2
}(
cmra_
core
_l
∅
)
right_id
.
Qed
.
End
identity
.
(** ** Local updates *)
...
...
@@ -468,7 +468,7 @@ Section cmra_transport.
Proof
.
by
intros
???
;
destruct
H
.
Qed
.
Lemma
cmra_transport_op
x
y
:
T
(
x
⋅
y
)
=
T
x
⋅
T
y
.
Proof
.
by
destruct
H
.
Qed
.
Lemma
cmra_transport_
unit
x
:
T
(
unit
x
)
=
unit
(
T
x
).
Lemma
cmra_transport_
core
x
:
T
(
core
x
)
=
core
(
T
x
).
Proof
.
by
destruct
H
.
Qed
.
Lemma
cmra_transport_validN
n
x
:
✓
{
n
}
T
x
↔
✓
{
n
}
x
.
Proof
.
by
destruct
H
.
Qed
.
...
...
@@ -486,25 +486,25 @@ End cmra_transport.
(** * Instances *)
(** ** Discrete CMRA *)
Class
RA
A
`
{
Equiv
A
,
Unit
A
,
Op
A
,
Valid
A
,
Div
A
}
:
=
{
Class
RA
A
`
{
Equiv
A
,
Core
A
,
Op
A
,
Valid
A
,
Div
A
}
:
=
{
(* setoids *)
ra_op_ne
(
x
:
A
)
:
Proper
((
≡
)
==>
(
≡
))
(
op
x
)
;
ra_
unit
_ne
:
>
Proper
((
≡
)
==>
(
≡
))
unit
;
ra_
core
_ne
:
>
Proper
((
≡
)
==>
(
≡
))
core
;
ra_validN_ne
:
>
Proper
((
≡
)
==>
impl
)
valid
;
ra_div_ne
:
>
Proper
((
≡
)
==>
(
≡
)
==>
(
≡
))
div
;
(* monoid *)
ra_assoc
:
>
Assoc
(
≡
)
(
⋅
)
;
ra_comm
:
>
Comm
(
≡
)
(
⋅
)
;
ra_
unit
_l
x
:
unit
x
⋅
x
≡
x
;
ra_
unit
_idemp
x
:
unit
(
unit
x
)
≡
unit
x
;
ra_
unit
_preserving
x
y
:
x
≼
y
→
unit
x
≼
unit
y
;
ra_
core
_l
x
:
core
x
⋅
x
≡
x
;
ra_
core
_idemp
x
:
core
(
core
x
)
≡
core
x
;
ra_
core
_preserving
x
y
:
x
≼
y
→
core
x
≼
core
y
;
ra_valid_op_l
x
y
:
✓
(
x
⋅
y
)
→
✓
x
;
ra_op_div
x
y
:
x
≼
y
→
x
⋅
y
÷
x
≡
y
}.
Section
discrete
.
Context
{
A
:
cofeT
}
`
{
Discrete
A
}.
Context
`
{
Unit
A
,
Op
A
,
Valid
A
,
Div
A
}
(
ra
:
RA
A
).
Context
`
{
Core
A
,
Op
A
,
Valid
A
,
Div
A
}
(
ra
:
RA
A
).
Instance
discrete_validN
:
ValidN
A
:
=
λ
n
x
,
✓
x
.
Definition
discrete_cmra_mixin
:
CMRAMixin
A
.
...
...
@@ -523,7 +523,7 @@ End discrete.
(** ** CMRA for the unit type *)
Section
unit
.
Instance
unit_valid
:
Valid
()
:
=
λ
x
,
True
.
Instance
unit_
unit
:
Unit
()
:
=
λ
x
,
x
.
Instance
unit_
core
:
Core
()
:
=
λ
x
,
x
.
Instance
unit_op
:
Op
()
:
=
λ
x
y
,
().
Instance
unit_div
:
Div
()
:
=
λ
x
y
,
().
Global
Instance
unit_empty
:
Empty
()
:
=
().
...
...
@@ -541,7 +541,7 @@ Section prod.
Context
{
A
B
:
cmraT
}.
Instance
prod_op
:
Op
(
A
*
B
)
:
=
λ
x
y
,
(
x
.
1
⋅
y
.
1
,
x
.
2
⋅
y
.
2
).
Global
Instance
prod_empty
`
{
Empty
A
,
Empty
B
}
:
Empty
(
A
*
B
)
:
=
(
∅
,
∅
).
Instance
prod_
unit
:
Unit
(
A
*
B
)
:
=
λ
x
,
(
unit
(
x
.
1
),
unit
(
x
.
2
)).
Instance
prod_
core
:
Core
(
A
*
B
)
:
=
λ
x
,
(
core
(
x
.
1
),
core
(
x
.
2
)).
Instance
prod_valid
:
Valid
(
A
*
B
)
:
=
λ
x
,
✓
x
.
1
∧
✓
x
.
2
.
Instance
prod_validN
:
ValidN
(
A
*
B
)
:
=
λ
n
x
,
✓
{
n
}
x
.
1
∧
✓
{
n
}
x
.
2
.
Instance
prod_div
:
Div
(
A
*
B
)
:
=
λ
x
y
,
(
x
.
1
÷
y
.
1
,
x
.
2
÷
y
.
2
).
...
...
@@ -569,10 +569,10 @@ Section prod.
-
by
intros
n
x
[??]
;
split
;
apply
cmra_validN_S
.
-
by
split
;
rewrite
/=
assoc
.
-
by
split
;
rewrite
/=
comm
.
-
by
split
;
rewrite
/=
cmra_
unit
_l
.
-
by
split
;
rewrite
/=
cmra_
unit
_idemp
.
-
by
split
;
rewrite
/=
cmra_
core
_l
.
-
by
split
;
rewrite
/=
cmra_
core
_idemp
.
-
intros
x
y
;
rewrite
!
prod_included
.
by
intros
[??]
;
split
;
apply
cmra_
unit
_preserving
.
by
intros
[??]
;
split
;
apply
cmra_
core
_preserving
.
-
intros
n
x
y
[??]
;
split
;
simpl
in
*
;
eauto
using
cmra_validN_op_l
.
-
intros
x
y
;
rewrite
prod_included
;
intros
[??].
by
split
;
apply
cmra_op_div
.
...
...
algebra/dec_agree.v
View file @
c2c84732
...
...
@@ -2,7 +2,7 @@ From algebra Require Export cmra.
Local
Arguments
validN
_
_
_
!
_
/.
Local
Arguments
valid
_
_
!
_
/.
Local
Arguments
op
_
_
_
!
_
/.
Local
Arguments
unit
_
_
!
_
/.
Local
Arguments
core
_
_
!
_
/.
(* This is isomorphic to option, but has a very different RA structure. *)
Inductive
dec_agree
(
A
:
Type
)
:
Type
:
=
...
...
@@ -26,7 +26,7 @@ Instance dec_agree_op : Op (dec_agree A) := λ x y,
|
DecAgree
a
,
DecAgree
b
=>
if
decide
(
a
=
b
)
then
DecAgree
a
else
DecAgreeBot
|
_
,
_
=>
DecAgreeBot
end
.
Instance
dec_agree_
unit
:
Unit
(
dec_agree
A
)
:
=
id
.
Instance
dec_agree_
core
:
Core
(
dec_agree
A
)
:
=
id
.
Instance
dec_agree_div
:
Div
(
dec_agree
A
)
:
=
λ
x
y
,
x
.
Definition
dec_agree_ra
:
RA
(
dec_agree
A
).
...
...
algebra/dra.v
View file @
c2c84732
...
...
@@ -18,17 +18,17 @@ Definition dra_included `{Equiv A, Valid A, Disjoint A, Op A} := λ x y,
Instance
:
Params
(@
dra_included
)
4
.
Local
Infix
"≼"
:
=
dra_included
.
Class
DRA
A
`
{
Equiv
A
,
Valid
A
,
Unit
A
,
Disjoint
A
,
Op
A
,
Div
A
}
:
=
{
Class
DRA
A
`
{
Equiv
A
,
Valid
A
,
Core
A
,
Disjoint
A
,
Op
A
,
Div
A
}
:
=
{
(* setoids *)
dra_equivalence
:
>
Equivalence
((
≡
)
:
relation
A
)
;
dra_op_proper
:
>
Proper
((
≡
)
==>
(
≡
)
==>
(
≡
))
(
⋅
)
;
dra_
unit
_proper
:
>
Proper
((
≡
)
==>
(
≡
))
unit
;
dra_
core
_proper
:
>
Proper
((
≡
)
==>
(
≡
))
core
;
dra_valid_proper
:
>
Proper
((
≡
)
==>
impl
)
valid
;
dra_disjoint_proper
:
>
∀
x
,
Proper
((
≡
)
==>
impl
)
(
disjoint
x
)
;
dra_div_proper
:
>
Proper
((
≡
)
==>
(
≡
)
==>
(
≡
))
div
;
(* validity *)
dra_op_valid
x
y
:
✓
x
→
✓
y
→
x
⊥
y
→
✓
(
x
⋅
y
)
;
dra_
unit
_valid
x
:
✓
x
→
✓
unit
x
;
dra_
core
_valid
x
:
✓
x
→
✓
core
x
;
dra_div_valid
x
y
:
✓
x
→
✓
y
→
x
≼
y
→
✓
(
y
÷
x
)
;
(* monoid *)
dra_assoc
:
>
Assoc
(
≡
)
(
⋅
)
;
...
...
@@ -36,10 +36,10 @@ Class DRA A `{Equiv A, Valid A, Unit A, Disjoint A, Op A, Div A} := {
dra_disjoint_move_l
x
y
z
:
✓
x
→
✓
y
→
✓
z
→
x
⊥
y
→
x
⋅
y
⊥
z
→
x
⊥
y
⋅
z
;
dra_symmetric
:
>
Symmetric
(@
disjoint
A
_
)
;
dra_comm
x
y
:
✓
x
→
✓
y
→
x
⊥
y
→
x
⋅
y
≡
y
⋅
x
;
dra_
unit
_disjoint_l
x
:
✓
x
→
unit
x
⊥
x
;
dra_
unit
_l
x
:
✓
x
→
unit
x
⋅
x
≡
x
;
dra_
unit
_idemp
x
:
✓
x
→
unit
(
unit
x
)
≡
unit
x
;
dra_
unit
_preserving
x
y
:
✓
x
→
✓
y
→
x
≼
y
→
unit
x
≼
unit
y
;
dra_
core
_disjoint_l
x
:
✓
x
→
core
x
⊥
x
;
dra_
core
_l
x
:
✓
x
→
core
x
⋅
x
≡
x
;
dra_
core
_idemp
x
:
✓
x
→
core
(
core
x
)
≡
core
x
;
dra_
core
_preserving
x
y
:
✓
x
→
✓
y
→
x
≼
y
→
core
x
≼
core
y
;
dra_disjoint_div
x
y
:
✓
x
→
✓
y
→
x
≼
y
→
x
⊥
y
÷
x
;
dra_op_div
x
y
:
✓
x
→
✓
y
→
x
≼
y
→
x
⋅
y
÷
x
≡
y
}.
...
...
@@ -88,9 +88,9 @@ Hint Unfold dra_included.
Lemma
validity_valid_car_valid
(
z
:
T
)
:
✓
z
→
✓
validity_car
z
.
Proof
.
apply
validity_prf
.
Qed
.
Hint
Resolve
validity_valid_car_valid
.
Program
Instance
validity_
unit
:
Unit
T
:
=
λ
x
,
Validity
(
unit
(
validity_car
x
))
(
✓
x
)
_
.
Solve
Obligations
with
naive_solver
auto
using
dra_
unit
_valid
.
Program
Instance
validity_
core
:
Core
T
:
=
λ
x
,
Validity
(
core
(
validity_car
x
))
(
✓
x
)
_
.
Solve
Obligations
with
naive_solver
auto
using
dra_
core
_valid
.
Program
Instance
validity_op
:
Op
T
:
=
λ
x
y
,
Validity
(
validity_car
x
⋅
validity_car
y
)
(
✓
x
∧
✓
y
∧
validity_car
x
⊥
validity_car
y
)
_
.
...
...
@@ -118,14 +118,14 @@ Proof.
|
by
intros
;
rewrite
assoc
].
-
intros
[
x
px
?]
[
y
py
?]
;
split
;
naive_solver
eauto
using
dra_comm
.
-
intros
[
x
px
?]
;
split
;
naive_solver
eauto
using
dra_
unit
_l
,
dra_
unit
_disjoint_l
.
-
intros
[
x
px
?]
;
split
;
naive_solver
eauto
using
dra_
unit
_idemp
.
-
intros
x
y
Hxy
;
exists
(
unit
y
÷
unit
x
).
naive_solver
eauto
using
dra_
core
_l
,
dra_
core
_disjoint_l
.
-
intros
[
x
px
?]
;
split
;
naive_solver
eauto
using
dra_
core
_idemp
.
-
intros
x
y
Hxy
;
exists
(
core
y
÷
core
x
).
destruct
x
as
[
x
px
?],
y
as
[
y
py
?],
Hxy
as
[[
z
pz
?]
[??]]
;
simpl
in
*.
assert
(
py
→
unit
x
≼
unit
y
)
by
intuition
eauto
10
using
dra_
unit
_preserving
.
assert
(
py
→
core
x
≼
core
y
)
by
intuition
eauto
10
using
dra_
core
_preserving
.
constructor
;
[|
symmetry
]
;
simpl
in
*
;
intuition
eauto
using
dra_op_div
,
dra_disjoint_div
,
dra_
unit
_valid
.
intuition
eauto
using
dra_op_div
,
dra_disjoint_div
,
dra_
core
_valid
.
-
by
intros
[
x
px
?]
[
y
py
?]
(?&?&?).
-
intros
[
x
px
?]
[
y
py
?]
[[
z
pz
?]
[??]]
;
split
;
simpl
in
*
;
intuition
eauto
10
using
dra_disjoint_div
,
dra_op_div
.
...
...
algebra/excl.v
View file @
c2c84732
...
...
@@ -91,7 +91,7 @@ Instance excl_valid : Valid (excl A) := λ x,
Instance
excl_validN
:
ValidN
(
excl
A
)
:
=
λ
n
x
,
match
x
with
Excl
_
|
ExclUnit
=>
True
|
ExclBot
=>
False
end
.
Global
Instance
excl_empty
:
Empty
(
excl
A
)
:
=
ExclUnit
.
Instance
excl_
unit
:
Unit
(
excl
A
)
:
=
λ
_
,
∅
.
Instance
excl_
core
:
Core
(
excl
A
)
:
=
λ
_
,
∅
.
Instance
excl_op
:
Op
(
excl
A
)
:
=
λ
x
y
,
match
x
,
y
with
|
Excl
a
,
ExclUnit
|
ExclUnit
,
Excl
a
=>
Excl
a
...
...
algebra/fin_maps.v
View file @
c2c84732
...
...
@@ -93,7 +93,7 @@ Context `{Countable K} {A : cmraT}.
Implicit
Types
m
:
gmap
K
A
.
Instance
map_op
:
Op
(
gmap
K
A
)
:
=
merge
op
.
Instance
map_
unit
:
Unit
(
gmap
K
A
)
:
=
fmap
unit
.
Instance
map_
core
:
Core
(
gmap
K
A
)
:
=
fmap
core
.
Instance
map_valid
:
Valid
(
gmap
K
A
)
:
=
λ
m
,
∀
i
,
✓
(
m
!!
i
).
Instance
map_validN
:
ValidN
(
gmap
K
A
)
:
=
λ
n
m
,
∀
i
,
✓
{
n
}
(
m
!!
i
).
Instance
map_div
:
Div
(
gmap
K
A
)
:
=
merge
div
.
...
...
@@ -102,7 +102,7 @@ Lemma lookup_op m1 m2 i : (m1 ⋅ m2) !! i = m1 !! i ⋅ m2 !! i.
Proof
.
by
apply
lookup_merge
.
Qed
.
Lemma
lookup_div
m1
m2
i
:
(
m1
÷
m2
)
!!
i
=
m1
!!
i
÷
m2
!!
i
.
Proof
.
by
apply
lookup_merge
.
Qed
.
Lemma
lookup_
unit
m
i
:
unit
m
!!
i
=
unit
(
m
!!
i
).
Lemma
lookup_
core
m
i
:
core
m
!!
i
=
core
(
m
!!
i
).
Proof
.
by
apply
lookup_fmap
.
Qed
.
Lemma
map_included_spec
(
m1
m2
:
gmap
K
A
)
:
m1
≼
m2
↔
∀
i
,
m1
!!
i
≼
m2
!!
i
.
...
...
@@ -125,7 +125,7 @@ Definition map_cmra_mixin : CMRAMixin (gmap K A).
Proof
.
split
.
-
by
intros
n
m1
m2
m3
Hm
i
;
rewrite
!
lookup_op
(
Hm
i
).
-
by
intros
n
m1
m2
Hm
i
;
rewrite
!
lookup_
unit
(
Hm
i
).
-
by
intros
n
m1
m2
Hm
i
;
rewrite
!
lookup_
core
(
Hm
i
).
-
by
intros
n
m1
m2
Hm
?
i
;
rewrite
-(
Hm
i
).
-
by
intros
n
m1
m1'
Hm1
m2
m2'
Hm2
i
;
rewrite
!
lookup_div
(
Hm1
i
)
(
Hm2
i
).
-
intros
m
;
split
.
...
...
@@ -134,10 +134,10 @@ Proof.
-
intros
n
m
Hm
i
;
apply
cmra_validN_S
,
Hm
.
-
by
intros
m1
m2
m3
i
;
rewrite
!
lookup_op
assoc
.
-
by
intros
m1
m2
i
;
rewrite
!
lookup_op
comm
.
-
by
intros
m
i
;
rewrite
lookup_op
!
lookup_
unit
cmra_
unit
_l
.
-
by
intros
m
i
;
rewrite
!
lookup_
unit
cmra_
unit
_idemp
.
-
by
intros
m
i
;
rewrite
lookup_op
!
lookup_
core
cmra_
core
_l
.
-
by
intros
m
i
;
rewrite
!
lookup_
core
cmra_
core
_idemp
.
-
intros
x
y
;
rewrite
!
map_included_spec
;
intros
Hm
i
.
by
rewrite
!
lookup_
unit
;
apply
cmra_
unit
_preserving
.
by
rewrite
!
lookup_
core
;
apply
cmra_
core
_preserving
.
-
intros
n
m1
m2
Hm
i
;
apply
cmra_validN_op_l
with
(
m2
!!
i
).
by
rewrite
-
lookup_op
.
-
intros
x
y
;
rewrite
map_included_spec
=>
?
i
.
...
...
@@ -201,21 +201,21 @@ Lemma map_singleton_valid i x : ✓ ({[ i := x ]} : gmap K A) ↔ ✓ x.
Proof
.
rewrite
!
cmra_valid_validN
.
by
setoid_rewrite
map_singleton_validN
.
Qed
.
Lemma
map_insert_singleton_opN
n
m
i
x
:
m
!!
i
=
None
∨
m
!!
i
≡
{
n
}
≡
Some
(
unit
x
)
→
<[
i
:
=
x
]>
m
≡
{
n
}
≡
{[
i
:
=
x
]}
⋅
m
.
m
!!
i
=
None
∨
m
!!
i
≡
{
n
}
≡
Some
(
core
x
)
→
<[
i
:
=
x
]>
m
≡
{
n
}
≡
{[
i
:
=
x
]}
⋅
m
.
Proof
.
intros
Hi
j
;
destruct
(
decide
(
i
=
j
))
as
[->|]
;
[|
by
rewrite
lookup_op
lookup_insert_ne
//
lookup_singleton_ne
//
left_id
].
rewrite
lookup_op
lookup_insert
lookup_singleton
.
by
destruct
Hi
as
[->|
->]
;
constructor
;
rewrite
?cmra_
unit
_r
.
by
destruct
Hi
as
[->|
->]
;
constructor
;
rewrite
?cmra_
core
_r
.
Qed
.
Lemma
map_insert_singleton_op
m
i
x
:
m
!!
i
=
None
∨
m
!!
i
≡
Some
(
unit
x
)
→
<[
i
:
=
x
]>
m
≡
{[
i
:
=
x
]}
⋅
m
.
m
!!
i
=
None
∨
m
!!
i
≡
Some
(
core
x
)
→
<[
i
:
=
x
]>
m
≡
{[
i
:
=
x
]}
⋅
m
.
Proof
.
rewrite
!
equiv_dist
;
naive_solver
eauto
using
map_insert_singleton_opN
.
Qed
.
Lemma
map_
unit
_singleton
(
i
:
K
)
(
x
:
A
)
:
unit
({[
i
:
=
x
]}
:
gmap
K
A
)
=
{[
i
:
=
unit
x
]}.
Lemma
map_
core
_singleton
(
i
:
K
)
(
x
:
A
)
:
core
({[
i
:
=
x
]}
:
gmap
K
A
)
=
{[
i
:
=
core
x
]}.
Proof
.
apply
map_fmap_singleton
.
Qed
.
Lemma
map_op_singleton
(
i
:
K
)
(
x
y
:
A
)
:
{[
i
:
=
x
]}
⋅
{[
i
:
=
y
]}
=
({[
i
:
=
x
⋅
y
]}
:
gmap
K
A
).
...
...
@@ -315,7 +315,7 @@ End freshness.
(* Allocation is a local update: Just use composition with a singleton map. *)
(* Deallocation is *not* a local update. The trouble is that if we
own {[ i ↦ x ]}, then the frame could always own "
unit
x", and prevent
own {[ i ↦ x ]}, then the frame could always own "
core
x", and prevent
deallocation. *)
(* Applying a local update at a position we own is a local update. *)
...
...
algebra/frac.v
View file @
c2c84732
...
...
@@ -122,7 +122,7 @@ Instance frac_valid : Valid (frac A) := λ x,
Instance
frac_validN
:
ValidN
(
frac
A
)
:
=
λ
n
x
,
match
x
with
Frac
q
a
=>
(
q
≤
1
)%
Qc
∧
✓
{
n
}
a
|
FracUnit
=>
True
end
.
Global
Instance
frac_empty
:
Empty
(
frac
A
)
:
=
FracUnit
.
Instance
frac_
unit
:
Unit
(
frac
A
)
:
=
λ
_
,
∅
.
Instance
frac_
core
:
Core
(
frac
A
)
:
=
λ
_
,
∅
.
Instance
frac_op
:
Op
(
frac
A
)
:
=
λ
x
y
,
match
x
,
y
with
|
Frac
q1
a
,
Frac
q2
b
=>
Frac
(
q1
+
q2
)
(
a
⋅
b
)
...
...
@@ -225,7 +225,7 @@ Qed.
Lemma
frac_update
(
a1
a2
:
A
)
p
:
a1
~~>
a2
→
Frac
p
a1
~~>
Frac
p
a2
.
Proof
.
intros
Ha
n
[
q
b
|]
[??]
;
split
;
auto
.
apply
cmra_validN_op_l
with
(
unit
a1
),
Ha
.
by
rewrite
cmra_
unit
_r
.
apply
cmra_validN_op_l
with
(
core
a1
),
Ha
.
by
rewrite
cmra_
core
_r
.
Qed
.
End
cmra
.
...
...
algebra/iprod.v
View file @
c2c84732
...
...
@@ -117,13 +117,13 @@ Section iprod_cmra.
Implicit
Types
f
g
:
iprod
B
.
Instance
iprod_op
:
Op
(
iprod
B
)
:
=
λ
f
g
x
,
f
x
⋅
g
x
.
Instance
iprod_
unit
:
Unit
(
iprod
B
)
:
=
λ
f
x
,
unit
(
f
x
).
Instance
iprod_
core
:
Core
(
iprod
B
)
:
=
λ
f
x
,
core
(
f
x
).
Instance
iprod_valid
:
Valid
(
iprod
B
)
:
=
λ
f
,
∀
x
,
✓
f
x
.
Instance
iprod_validN
:
ValidN
(
iprod
B
)
:
=
λ
n
f
,
∀
x
,
✓
{
n
}
f
x
.
Instance
iprod_div
:
Div
(
iprod
B
)
:
=
λ
f
g
x
,
f
x
÷
g
x
.
Definition
iprod_lookup_op
f
g
x
:
(
f
⋅
g
)
x
=
f
x
⋅
g
x
:
=
eq_refl
.
Definition
iprod_lookup_
unit
f
x
:
(
unit
f
)
x
=
unit
(
f
x
)
:
=
eq_refl
.
Definition
iprod_lookup_
core
f
x
:
(
core
f
)
x
=
core
(
f
x
)
:
=
eq_refl
.
Definition
iprod_lookup_div
f
g
x
:
(
f
÷
g
)
x
=
f
x
÷
g
x
:
=
eq_refl
.
Lemma
iprod_included_spec
(
f
g
:
iprod
B
)
:
f
≼
g
↔
∀
x
,
f
x
≼
g
x
.
...
...
@@ -138,7 +138,7 @@ Section iprod_cmra.
Proof
.
split
.
-
by
intros
n
f1
f2
f3
Hf
x
;
rewrite
iprod_lookup_op
(
Hf
x
).
-
by
intros
n
f1
f2
Hf
x
;
rewrite
iprod_lookup_
unit
(
Hf
x
).
-
by
intros
n
f1
f2
Hf
x
;
rewrite
iprod_lookup_
core
(
Hf
x
).
-
by
intros
n
f1
f2
Hf
?
x
;
rewrite
-(
Hf
x
).
-
by
intros
n
f
f'
Hf
g
g'
Hg
i
;
rewrite
iprod_lookup_div
(
Hf
i
)
(
Hg
i
).
-
intros
g
;
split
.
...
...
@@ -147,10 +147,10 @@ Section iprod_cmra.
-
intros
n
f
Hf
x
;
apply
cmra_validN_S
,
Hf
.
-
by
intros
f1
f2
f3
x
;
rewrite
iprod_lookup_op
assoc
.
-
by
intros
f1
f2
x
;
rewrite
iprod_lookup_op
comm
.
-
by
intros
f
x
;
rewrite
iprod_lookup_op
iprod_lookup_
unit
cmra_
unit
_l
.
-
by
intros
f
x
;
rewrite
iprod_lookup_
unit
cmra_
unit
_idemp
.
-
by
intros
f
x
;
rewrite
iprod_lookup_op
iprod_lookup_
core
cmra_
core
_l
.
-
by
intros
f
x
;
rewrite
iprod_lookup_
core
cmra_
core
_idemp
.
-
intros
f1
f2
;
rewrite
!
iprod_included_spec
=>
Hf
x
.
by
rewrite
iprod_lookup_
unit
;
apply
cmra_
unit
_preserving
,
Hf
.
by
rewrite
iprod_lookup_
core
;
apply
cmra_
core
_preserving
,
Hf
.
-
intros
n
f1
f2
Hf
x
;
apply
cmra_validN_op_l
with
(
f2
x
),
Hf
.
-
intros
f1
f2
;
rewrite
iprod_included_spec
=>
Hf
x
.
by
rewrite
iprod_lookup_op
iprod_lookup_div
cmra_op_div
;
try
apply
Hf
.
...
...
@@ -211,12 +211,12 @@ Section iprod_cmra.
by
apply
cmra_empty_validN
.
Qed
.
Lemma
iprod_
unit
_singleton
x
(
y
:
B
x
)
:
unit
(
iprod_singleton
x
y
)
≡
iprod_singleton
x
(
unit
y
).
Lemma
iprod_
core
_singleton
x
(
y
:
B
x
)
:
core
(
iprod_singleton
x
y
)
≡
iprod_singleton
x
(
core
y
).
Proof
.
by
move
=>
x'
;
destruct
(
decide
(
x
=
x'
))
as
[->|]
;
rewrite
iprod_lookup_
unit
?iprod_lookup_singleton
?iprod_lookup_singleton_ne
//
cmra_
unit
_empty
.
rewrite
iprod_lookup_
core
?iprod_lookup_singleton
?iprod_lookup_singleton_ne
//
cmra_
core
_empty
.
Qed
.
Lemma
iprod_op_singleton
(
x
:
A
)
(
y1
y2
:
B
x
)
:
...
...
algebra/option.v