Skip to content
GitLab
Projects
Groups
Snippets
Help
Loading...
Help
Help
Support
Community forum
Keyboard shortcuts
?
Submit feedback
Contribute to GitLab
Sign in / Register
Toggle navigation
I
Iris
Project overview
Project overview
Details
Activity
Releases
Repository
Repository
Files
Commits
Branches
Tags
Contributors
Graph
Compare
Issues
0
Issues
0
List
Boards
Labels
Service Desk
Milestones
Merge Requests
0
Merge Requests
0
CI / CD
CI / CD
Pipelines
Jobs
Schedules
Operations
Operations
Incidents
Environments
Analytics
Analytics
CI / CD
Repository
Value Stream
Wiki
Wiki
Snippets
Snippets
Members
Members
Collapse sidebar
Close sidebar
Activity
Graph
Create a new issue
Jobs
Commits
Issue Boards
Open sidebar
George Pirlea
Iris
Commits
9df894ee
Commit
9df894ee
authored
Jan 13, 2016
by
Robbert Krebbers
Browse files
Options
Browse Files
Download
Email Patches
Plain Diff
Use ssreflect for modures.
parent
3967e715
Changes
13
Show whitespace changes
Inline
Side-by-side
Showing
13 changed files
with
193 additions
and
196 deletions
+193
-196
channel/heap_lang.v
channel/heap_lang.v
+0
-1
modures/agree.v
modures/agree.v
+9
-8
modures/auth.v
modures/auth.v
+4
-4
modures/base.v
modures/base.v
+3
-0
modures/cmra.v
modures/cmra.v
+13
-13
modures/cofe.v
modures/cofe.v
+11
-12
modures/cofe_solver.v
modures/cofe_solver.v
+32
-33
modures/dra.v
modures/dra.v
+10
-11
modures/excl.v
modures/excl.v
+4
-2
modures/fin_maps.v
modures/fin_maps.v
+28
-28
modures/logic.v
modures/logic.v
+49
-55
modures/ra.v
modures/ra.v
+23
-22
modures/sts.v
modures/sts.v
+7
-7
No files found.
channel/heap_lang.v
View file @
9df894ee
Require
Import
mathcomp
.
ssreflect
.
ssreflect
.
Require
Import
Autosubst
.
Autosubst
.
Require
Import
prelude
.
option
prelude
.
gmap
iris
.
language
.
...
...
modures/agree.v
View file @
9df894ee
...
...
@@ -71,7 +71,7 @@ Proof.
eauto
using
agree_valid_le
.
Qed
.
Instance
:
Proper
(
dist
n
==>
dist
n
==>
dist
n
)
op
.
Proof
.
by
intros
n
x1
x2
Hx
y1
y2
Hy
;
rewrite
Hy
,!(
commutative
_
_
y2
),
Hx
.
Qed
.
Proof
.
by
intros
n
x1
x2
Hx
y1
y2
Hy
;
rewrite
Hy
!(
commutative
_
_
y2
)
Hx
.
Qed
.
Instance
:
Proper
((
≡
)
==>
(
≡
)
==>
(
≡
))
op
:
=
ne_proper_2
_
.
Instance
:
Associative
(
≡
)
(@
op
(
agree
A
)
_
).
Proof
.
...
...
@@ -82,19 +82,20 @@ Qed.
Lemma
agree_includedN
(
x
y
:
agree
A
)
n
:
x
≼
{
n
}
y
↔
y
={
n
}=
x
⋅
y
.
Proof
.
split
;
[|
by
intros
?
;
exists
y
].
by
intros
[
z
Hz
]
;
rewrite
Hz
,
(
associative
_
),
agree_idempotent
.
by
intros
[
z
Hz
]
;
rewrite
Hz
(
associative
_
)
agree_idempotent
.
Qed
.
Global
Instance
agree_cmra
:
CMRA
(
agree
A
).
Proof
.
split
;
try
(
apply
_
||
done
).
*
intros
n
x
y
Hxy
[?
Hx
]
;
split
;
[
by
apply
Hxy
|
intros
n'
?].
rewrite
<-(
proj2
Hxy
n'
),
(
Hx
n'
)
by
eauto
using
agree_valid_le
.
rewrite
-(
proj2
Hxy
n'
)
1
?(
Hx
n'
)
;
eauto
using
agree_valid_le
.
by
apply
dist_le
with
n
;
try
apply
Hxy
.
*
by
intros
n
x1
x2
Hx
y1
y2
Hy
.
*
intros
x
;
split
;
[
apply
agree_valid_0
|].
by
intros
n'
;
rewrite
Nat
.
le_0_r
;
intros
->.
*
intros
n
x
[?
Hx
]
;
split
;
[
by
apply
agree_valid_S
|
intros
n'
?].
rewrite
(
Hx
n'
)
by
auto
;
symmetry
;
apply
dist_le
with
n
;
try
apply
Hx
;
auto
.
rewrite
(
Hx
n'
)
;
last
auto
.
symmetry
;
apply
dist_le
with
n
;
try
apply
Hx
;
auto
.
*
intros
x
;
apply
agree_idempotent
.
*
by
intros
x
y
n
[(?&?&?)
?].
*
by
intros
x
y
n
;
rewrite
agree_includedN
.
...
...
@@ -106,7 +107,7 @@ Global Instance agree_extend : CMRAExtend (agree A).
Proof
.
intros
n
x
y1
y2
?
Hx
;
exists
(
x
,
x
)
;
simpl
;
split
.
*
by
rewrite
agree_idempotent
.
*
by
rewrite
Hx
,
(
agree_op_inv
x
y1
y2
),
agree_idempotent
by
done
.
*
by
rewrite
Hx
(
agree_op_inv
x
y1
y2
)
//
agree_idempotent
.
Qed
.
Program
Definition
to_agree
(
x
:
A
)
:
agree
A
:
=
{|
agree_car
n
:
=
x
;
agree_is_valid
n
:
=
True
|}.
...
...
@@ -120,7 +121,7 @@ Proof. by intros ? [? Hx]; apply Hx. Qed.
Lemma
agree_to_agree_inj
(
x
y
:
agree
A
)
a
n
:
✓
{
n
}
x
→
x
={
n
}=
to_agree
a
⋅
y
→
x
n
={
n
}=
a
.
Proof
.
by
intros
;
transitivity
((
to_agree
a
⋅
y
)
n
)
;
[
by
apply
agree_car_ne
|]
.
by
intros
;
transitivity
((
to_agree
a
⋅
y
)
n
)
;
first
apply
agree_car_ne
.
Qed
.
End
agree
.
...
...
@@ -141,7 +142,7 @@ Section agree_map.
Proof
.
split
;
[|
by
intros
n
x
[?
Hx
]
;
split
;
simpl
;
[|
by
intros
n'
?
;
rewrite
Hx
]].
intros
x
y
n
;
rewrite
!
agree_includedN
;
intros
Hy
;
rewrite
Hy
.
split
;
[
split
;
simpl
;
try
tauto
|
done
]
.
split
;
last
done
;
split
;
simpl
;
last
tauto
.
by
intros
(?&?&
Hxy
)
;
repeat
split
;
intros
;
try
apply
Hxy
;
try
apply
Hf
;
eauto
using
@
agree_valid_le
.
Qed
.
...
...
@@ -157,6 +158,6 @@ Definition agreeRA_map {A B} (f : A -n> B) : agreeRA A -n> agreeRA B :=
CofeMor
(
agree_map
f
:
agreeRA
A
→
agreeRA
B
).
Instance
agreeRA_map_ne
A
B
n
:
Proper
(
dist
n
==>
dist
n
)
(@
agreeRA_map
A
B
).
Proof
.
intros
f
g
Hfg
x
;
split
;
simpl
;
intros
;
[
done
|]
.
intros
f
g
Hfg
x
;
split
;
simpl
;
intros
;
first
done
.
by
apply
dist_le
with
n
;
try
apply
Hfg
.
Qed
.
modures/auth.v
View file @
9df894ee
...
...
@@ -86,15 +86,15 @@ Instance auth_cmra `{CMRA A} : CMRA (auth A).
Proof
.
split
.
*
apply
_
.
*
by
intros
n
x
y1
y2
[
Hy
Hy'
]
;
split
;
simpl
;
rewrite
?Hy
,
?Hy'
.
*
by
intros
n
y1
y2
[
Hy
Hy'
]
;
split
;
simpl
;
rewrite
?Hy
,
?Hy'
.
*
by
intros
n
x
y1
y2
[
Hy
Hy'
]
;
split
;
simpl
;
rewrite
?Hy
?Hy'
.
*
by
intros
n
y1
y2
[
Hy
Hy'
]
;
split
;
simpl
;
rewrite
?Hy
?Hy'
.
*
intros
n
[
x
a
]
[
y
b
]
[
Hx
Ha
]
;
simpl
in
*
;
destruct
Hx
as
[[][]|
|
|]
;
intros
?
;
cofe_subst
;
auto
.
*
by
intros
n
x1
x2
[
Hx
Hx'
]
y1
y2
[
Hy
Hy'
]
;
split
;
simpl
;
rewrite
?Hy
,
?Hy'
,
?Hx
,
?Hx'
.
split
;
simpl
;
rewrite
?Hy
?Hy'
?Hx
?Hx'
.
*
by
intros
[[]
?]
;
simpl
.
*
intros
n
[[]
?]
?
;
naive_solver
eauto
using
cmra_included_S
,
cmra_valid_S
.
*
destruct
x
as
[[
a
|
|]
b
]
;
simpl
;
rewrite
?cmra_included_includedN
,
*
destruct
x
as
[[
a
|
|]
b
]
;
simpl
;
rewrite
?cmra_included_includedN
?cmra_valid_validN
;
[
naive_solver
|
naive_solver
|].
split
;
[
done
|
intros
Hn
;
discriminate
(
Hn
1
)].
*
by
split
;
simpl
;
rewrite
(
associative
_
).
...
...
modures/base.v
0 → 100644
View file @
9df894ee
Require
Export
mathcomp
.
ssreflect
.
ssreflect
.
Require
Export
prelude
.
prelude
.
Global
Set
Bullet
Behavior
"Strict Subproofs"
.
modures/cmra.v
View file @
9df894ee
...
...
@@ -110,7 +110,7 @@ Global Instance cmra_ra : RA A.
Proof
.
split
;
try
by
(
destruct
cmra
;
eauto
using
ne_proper
,
ne_proper_2
with
typeclass_instances
).
*
by
intros
x1
x2
Hx
;
rewrite
!
cmra_valid_validN
;
intros
?
n
;
rewrite
<
-
Hx
.
*
by
intros
x1
x2
Hx
;
rewrite
!
cmra_valid_validN
;
intros
?
n
;
rewrite
-
Hx
.
*
intros
x
y
;
rewrite
!
cmra_included_includedN
.
eauto
using
cmra_unit_preserving
.
*
intros
x
y
;
rewrite
!
cmra_valid_validN
;
intros
?
n
.
...
...
@@ -125,10 +125,10 @@ Proof. induction 2; eauto using cmra_valid_S. Qed.
Global
Instance
ra_op_ne
n
:
Proper
(
dist
n
==>
dist
n
==>
dist
n
)
(
⋅
).
Proof
.
intros
x1
x2
Hx
y1
y2
Hy
.
by
rewrite
Hy
,
(
commutative
_
x1
),
Hx
,
(
commutative
_
y2
).
by
rewrite
Hy
(
commutative
_
x1
)
Hx
(
commutative
_
y2
).
Qed
.
Lemma
cmra_unit_valid
x
n
:
✓
{
n
}
x
→
✓
{
n
}
(
unit
x
).
Proof
.
rewrite
<-(
cmra_unit_l
x
)
at
1
;
apply
cmra_valid_op_l
.
Qed
.
Proof
.
rewrite
-{
1
}(
cmra_unit_l
x
)
;
apply
cmra_valid_op_l
.
Qed
.
(** * Timeless *)
Lemma
cmra_timeless_included_l
`
{!
CMRAExtend
A
}
x
y
:
...
...
@@ -136,7 +136,7 @@ Lemma cmra_timeless_included_l `{!CMRAExtend A} x y :
Proof
.
intros
??
[
x'
?].
destruct
(
cmra_extend_op
1
y
x
x'
)
as
([
z
z'
]&
Hy
&
Hz
&
Hz'
)
;
auto
;
simpl
in
*.
by
exists
z'
;
rewrite
Hy
,
(
timeless
x
z
)
by
done
.
by
exists
z'
;
rewrite
Hy
(
timeless
x
z
)
.
Qed
.
Lemma
cmra_timeless_included_r
n
x
y
:
Timeless
y
→
x
≼
{
1
}
y
→
x
≼
{
n
}
y
.
Proof
.
intros
?
[
x'
?].
exists
x'
.
by
apply
equiv_dist
,
(
timeless
y
).
Qed
.
...
...
@@ -145,8 +145,8 @@ Lemma cmra_op_timeless `{!CMRAExtend A} x1 x2 :
Proof
.
intros
???
z
Hz
.
destruct
(
cmra_extend_op
1
z
x1
x2
)
as
([
y1
y2
]&
Hz'
&?&?)
;
auto
;
simpl
in
*.
{
by
rewrite
<
-
?Hz
;
apply
cmra_valid_validN
.
}
by
rewrite
Hz'
,
(
timeless
x1
y1
),
(
timeless
x2
y2
).
{
by
rewrite
-
?Hz
;
apply
cmra_valid_validN
.
}
by
rewrite
Hz'
(
timeless
x1
y1
)
//
(
timeless
x2
y2
).
Qed
.
(** * Included *)
...
...
@@ -176,17 +176,17 @@ Proof.
split
.
*
by
intros
x
;
exists
(
unit
x
)
;
rewrite
ra_unit_r
.
*
intros
x
y
z
[
z1
Hy
]
[
z2
Hz
]
;
exists
(
z1
⋅
z2
).
by
rewrite
(
associative
_
)
,
<-
Hy
,
<
-
Hz
.
by
rewrite
(
associative
_
)
-
Hy
-
Hz
.
Qed
.
Lemma
cmra_valid_includedN
x
y
n
:
✓
{
n
}
y
→
x
≼
{
n
}
y
→
✓
{
n
}
x
.
Proof
.
intros
Hyv
[
z
Hy
]
;
re
write
Hy
in
Hyv
;
eauto
using
cmra_valid_op_l
.
Qed
.
Proof
.
intros
Hyv
[
z
Hy
]
;
re
vert
Hyv
;
rewrite
Hy
;
apply
cmra_valid_op_l
.
Qed
.
Lemma
cmra_valid_included
x
y
n
:
✓
{
n
}
y
→
x
≼
y
→
✓
{
n
}
x
.
Proof
.
rewrite
cmra_included_includedN
;
eauto
using
cmra_valid_includedN
.
Qed
.
Lemma
cmra_included_dist_l
x1
x2
x1'
n
:
x1
≼
x2
→
x1'
={
n
}=
x1
→
∃
x2'
,
x1'
≼
x2'
∧
x2'
={
n
}=
x2
.
Proof
.
intros
[
z
Hx2
]
Hx1
;
exists
(
x1'
⋅
z
)
;
split
;
auto
using
ra_included_l
.
by
rewrite
Hx1
,
Hx2
.
by
rewrite
Hx1
Hx2
.
Qed
.
(** * Properties of [(⇝)] relation *)
...
...
@@ -281,11 +281,11 @@ Qed.
Instance
prod_cmra
`
{
CMRA
A
,
CMRA
B
}
:
CMRA
(
A
*
B
).
Proof
.
split
;
try
apply
_
.
*
by
intros
n
x
y1
y2
[
Hy1
Hy2
]
;
split
;
simpl
in
*
;
rewrite
?Hy1
,
?Hy2
.
*
by
intros
n
y1
y2
[
Hy1
Hy2
]
;
split
;
simpl
in
*
;
rewrite
?Hy1
,
?Hy2
.
*
by
intros
n
y1
y2
[
Hy1
Hy2
]
[??]
;
split
;
simpl
in
*
;
rewrite
<-
?Hy1
,
<
-
?Hy2
.
*
by
intros
n
x
y1
y2
[
Hy1
Hy2
]
;
split
;
simpl
in
*
;
rewrite
?Hy1
?Hy2
.
*
by
intros
n
y1
y2
[
Hy1
Hy2
]
;
split
;
simpl
in
*
;
rewrite
?Hy1
?Hy2
.
*
by
intros
n
y1
y2
[
Hy1
Hy2
]
[??]
;
split
;
simpl
in
*
;
rewrite
-
?Hy1
-
?Hy2
.
*
by
intros
n
x1
x2
[
Hx1
Hx2
]
y1
y2
[
Hy1
Hy2
]
;
split
;
simpl
in
*
;
rewrite
?Hx1
,
?Hx2
,
?Hy1
,
?Hy2
.
split
;
simpl
in
*
;
rewrite
?Hx1
?Hx2
?Hy1
?Hy2
.
*
split
;
apply
cmra_valid_0
.
*
by
intros
n
x
[??]
;
split
;
apply
cmra_valid_S
.
*
intros
x
;
split
;
[
by
intros
[??]
n
;
split
;
apply
cmra_valid_validN
|].
...
...
modures/cofe.v
View file @
9df894ee
Require
Export
prelude
.
prelud
e
.
Require
Export
modures
.
bas
e
.
(** Unbundeled version *)
Class
Dist
A
:
=
dist
:
nat
→
relation
A
.
...
...
@@ -75,8 +75,7 @@ Section cofe.
Qed
.
Global
Instance
dist_proper
n
:
Proper
((
≡
)
==>
(
≡
)
==>
iff
)
(
dist
n
).
Proof
.
intros
x1
x2
Hx
y1
y2
Hy
.
by
rewrite
equiv_dist
in
Hx
,
Hy
;
rewrite
(
Hx
n
),
(
Hy
n
).
by
move
=>
x1
x2
/
equiv_dist
Hx
y1
y2
/
equiv_dist
Hy
;
rewrite
(
Hx
n
)
(
Hy
n
).
Qed
.
Global
Instance
dist_proper_2
n
x
:
Proper
((
≡
)
==>
iff
)
(
dist
n
x
).
Proof
.
by
apply
dist_proper
.
Qed
.
...
...
@@ -90,10 +89,10 @@ Section cofe.
Proper
((
≡
)
==>
(
≡
)
==>
(
≡
))
f
|
100
.
Proof
.
unfold
Proper
,
respectful
;
setoid_rewrite
equiv_dist
.
by
intros
x1
x2
Hx
y1
y2
Hy
n
;
rewrite
(
Hx
n
)
,
(
Hy
n
).
by
intros
x1
x2
Hx
y1
y2
Hy
n
;
rewrite
(
Hx
n
)
(
Hy
n
).
Qed
.
Lemma
compl_ne
(
c1
c2
:
chain
A
)
n
:
c1
n
={
n
}=
c2
n
→
compl
c1
={
n
}=
compl
c2
.
Proof
.
intros
.
by
rewrite
(
conv_compl
c1
n
)
,
(
conv_compl
c2
n
).
Qed
.
Proof
.
intros
.
by
rewrite
(
conv_compl
c1
n
)
(
conv_compl
c2
n
).
Qed
.
Lemma
compl_ext
(
c1
c2
:
chain
A
)
:
(
∀
i
,
c1
i
≡
c2
i
)
→
compl
c1
≡
compl
c2
.
Proof
.
setoid_rewrite
equiv_dist
;
naive_solver
eauto
using
compl_ne
.
Qed
.
Global
Instance
contractive_ne
`
{
Cofe
B
}
(
f
:
A
→
B
)
`
{!
Contractive
f
}
n
:
...
...
@@ -118,7 +117,7 @@ Program Definition fixpoint_chain `{Cofe A, Inhabited A} (f : A → A)
`
{!
Contractive
f
}
:
chain
A
:
=
{|
chain_car
i
:
=
Nat
.
iter
i
f
inhabitant
|}.
Next
Obligation
.
intros
A
????
f
?
x
n
;
induction
n
as
[|
n
IH
]
;
intros
i
?
;
[
done
|].
destruct
i
as
[|
i
]
;
simpl
;
try
lia
;
apply
contractive
,
IH
;
auto
with
lia
.
destruct
i
as
[|
i
]
;
simpl
;
first
lia
;
apply
contractive
,
IH
;
auto
with
lia
.
Qed
.
Program
Definition
fixpoint
`
{
Cofe
A
,
Inhabited
A
}
(
f
:
A
→
A
)
`
{!
Contractive
f
}
:
A
:
=
compl
(
fixpoint_chain
f
).
...
...
@@ -129,14 +128,14 @@ Section fixpoint.
Proof
.
apply
equiv_dist
;
intros
n
;
unfold
fixpoint
.
rewrite
(
conv_compl
(
fixpoint_chain
f
)
n
).
by
rewrite
(
chain_cauchy
(
fixpoint_chain
f
)
n
(
S
n
))
at
1
by
lia
.
by
rewrite
{
1
}(
chain_cauchy
(
fixpoint_chain
f
)
n
(
S
n
))
;
last
lia
.
Qed
.
Lemma
fixpoint_ne
(
g
:
A
→
A
)
`
{!
Contractive
g
}
n
:
(
∀
z
,
f
z
={
n
}=
g
z
)
→
fixpoint
f
={
n
}=
fixpoint
g
.
Proof
.
intros
Hfg
;
unfold
fixpoint
.
rewrite
(
conv_compl
(
fixpoint_chain
f
)
n
)
,
(
conv_compl
(
fixpoint_chain
g
)
n
).
induction
n
as
[|
n
IH
]
;
simpl
in
*
;
[
done
|]
.
rewrite
(
conv_compl
(
fixpoint_chain
f
)
n
)
(
conv_compl
(
fixpoint_chain
g
)
n
).
induction
n
as
[|
n
IH
]
;
simpl
in
*
;
first
done
.
rewrite
Hfg
;
apply
contractive
,
IH
;
auto
using
dist_S
.
Qed
.
Lemma
fixpoint_proper
(
g
:
A
→
A
)
`
{!
Contractive
g
}
:
...
...
@@ -166,8 +165,8 @@ Program Instance cofe_mor_compl (A B : cofeT) : Compl (cofeMor A B) := λ c,
{|
cofe_mor_car
x
:
=
compl
(
fun_chain
c
x
)
|}.
Next
Obligation
.
intros
A
B
c
n
x
y
Hxy
.
rewrite
(
conv_compl
(
fun_chain
c
x
)
n
)
,
(
conv_compl
(
fun_chain
c
y
)
n
)
.
simpl
;
rewrite
Hxy
;
apply
(
chain_cauchy
c
)
;
lia
.
rewrite
(
conv_compl
(
fun_chain
c
x
)
n
)
(
conv_compl
(
fun_chain
c
y
)
n
)
/=
Hxy
.
apply
(
chain_cauchy
c
)
;
lia
.
Qed
.
Instance
cofe_mor_cofe
(
A
B
:
cofeT
)
:
Cofe
(
cofeMor
A
B
).
Proof
.
...
...
@@ -204,7 +203,7 @@ Instance: Params (@ccompose) 3.
Infix
"◎"
:
=
ccompose
(
at
level
40
,
left
associativity
).
Lemma
ccompose_ne
{
A
B
C
}
(
f1
f2
:
B
-
n
>
C
)
(
g1
g2
:
A
-
n
>
B
)
n
:
f1
={
n
}=
f2
→
g1
={
n
}=
g2
→
f1
◎
g1
={
n
}=
f2
◎
g2
.
Proof
.
by
intros
Hf
Hg
x
;
simpl
;
rewrite
(
Hg
x
),
(
Hf
(
g2
x
)).
Qed
.
Proof
.
by
intros
Hf
Hg
x
;
rewrite
/=
(
Hg
x
)
(
Hf
(
g2
x
)).
Qed
.
(** Pre-composition as a functor *)
Local
Instance
ccompose_l_ne'
{
A
B
C
}
(
f
:
B
-
n
>
A
)
n
:
...
...
modures/cofe_solver.v
View file @
9df894ee
...
...
@@ -17,7 +17,7 @@ Lemma map_ext {A1 A2 B1 B2 : cofeT}
(
f
:
A2
-
n
>
A1
)
(
f'
:
A2
-
n
>
A1
)
(
g
:
B1
-
n
>
B2
)
(
g'
:
B1
-
n
>
B2
)
x
x'
:
(
∀
x
,
f
x
≡
f'
x
)
→
(
∀
y
,
g
y
≡
g'
y
)
→
x
≡
x'
→
map
(
f
,
g
)
x
≡
map
(
f'
,
g'
)
x'
.
Proof
.
by
rewrite
<-!
cofe_mor_ext
;
intros
Hf
Hg
Hx
;
rewrite
Hf
,
Hg
,
Hx
.
Qed
.
Proof
.
by
rewrite
-!
cofe_mor_ext
;
intros
->
->
->
.
Qed
.
Fixpoint
A
(
k
:
nat
)
:
cofeT
:
=
match
k
with
0
=>
CofeT
unit
|
S
k
=>
F
(
A
k
)
(
A
k
)
end
.
...
...
@@ -30,13 +30,13 @@ Definition g_S k (x : A (S (S k))) : g x = map (f,g) x := eq_refl.
Lemma
gf
{
k
}
(
x
:
A
k
)
:
g
(
f
x
)
≡
x
.
Proof
.
induction
k
as
[|
k
IH
]
;
simpl
in
*
;
[
by
destruct
x
|].
rewrite
<-
map_comp
;
rewrite
<-(
map_id
_
_
x
)
at
2
;
by
apply
map_ext
.
rewrite
-
map_comp
-{
2
}(
map_id
_
_
x
)
;
by
apply
map_ext
.
Qed
.
Lemma
fg
{
n
k
}
(
x
:
A
(
S
k
))
:
n
≤
k
→
f
(
g
x
)
={
n
}=
x
.
Proof
.
intros
Hnk
;
apply
dist_le
with
k
;
auto
;
clear
Hnk
.
induction
k
as
[|
k
IH
]
;
simpl
;
[
apply
dist_0
|].
rewrite
<-(
map_id
_
_
x
)
at
2
;
rewrite
<
-
map_comp
;
by
apply
map_contractive
.
rewrite
-{
2
}(
map_id
_
_
x
)
-
map_comp
;
by
apply
map_contractive
.
Qed
.
Arguments
A
_
:
simpl
never
.
Arguments
f
{
_
}
:
simpl
never
.
...
...
@@ -56,8 +56,7 @@ Program Instance tower_compl : Compl tower := λ c,
Next
Obligation
.
intros
c
k
;
apply
equiv_dist
;
intros
n
.
rewrite
(
conv_compl
(
tower_chain
c
k
)
n
).
rewrite
(
conv_compl
(
tower_chain
c
(
S
k
))
n
)
;
simpl
.
by
rewrite
(
g_tower
(
c
n
)
k
).
by
rewrite
(
conv_compl
(
tower_chain
c
(
S
k
))
n
)
/=
(
g_tower
(
c
n
)
k
).
Qed
.
Instance
tower_cofe
:
Cofe
tower
.
Proof
.
...
...
@@ -69,9 +68,9 @@ Proof.
+
by
intros
X
Y
?
n
.
+
by
intros
X
Y
Z
??
n
;
transitivity
(
Y
n
).
*
intros
k
X
Y
HXY
n
;
apply
dist_S
.
by
rewrite
<-(
g_tower
X
),
(
HXY
(
S
n
)),
g_tower
.
by
rewrite
-(
g_tower
X
)
(
HXY
(
S
n
))
g_tower
.
*
intros
X
Y
k
;
apply
dist_0
.
*
intros
c
n
k
;
simpl
;
rewrite
(
conv_compl
(
tower_chain
c
k
)
n
).
*
intros
c
n
k
;
rewrite
/=
(
conv_compl
(
tower_chain
c
k
)
n
).
apply
(
chain_cauchy
c
)
;
lia
.
Qed
.
Definition
T
:
cofeT
:
=
CofeT
tower
.
...
...
@@ -81,16 +80,16 @@ Fixpoint ff {k} (i : nat) : A k -n> A (i + k) :=
Fixpoint
gg
{
k
}
(
i
:
nat
)
:
A
(
i
+
k
)
-
n
>
A
k
:
=
match
i
with
0
=>
cid
|
S
i
=>
gg
i
◎
g
end
.
Lemma
ggff
{
k
i
}
(
x
:
A
k
)
:
gg
i
(
ff
i
x
)
≡
x
.
Proof
.
induction
i
as
[|
i
IH
]
;
simpl
;
[
done
|
by
rewrite
(
gf
(
ff
i
x
))
,
IH
].
Qed
.
Proof
.
induction
i
as
[|
i
IH
]
;
simpl
;
[
done
|
by
rewrite
(
gf
(
ff
i
x
))
IH
].
Qed
.
Lemma
f_tower
{
n
k
}
(
X
:
tower
)
:
n
≤
k
→
f
(
X
k
)
={
n
}=
X
(
S
k
).
Proof
.
intros
.
by
rewrite
<-(
fg
(
X
(
S
k
))),
<
-(
g_tower
X
).
Qed
.
Proof
.
intros
.
by
rewrite
-(
fg
(
X
(
S
k
)))
//
-(
g_tower
X
).
Qed
.
Lemma
ff_tower
{
n
}
k
i
(
X
:
tower
)
:
n
≤
k
→
ff
i
(
X
k
)
={
n
}=
X
(
i
+
k
).
Proof
.
intros
;
induction
i
as
[|
i
IH
]
;
simpl
;
[
done
|].
by
rewrite
IH
,
(
f_tower
X
)
by
lia
.
by
rewrite
IH
(
f_tower
X
)
;
last
lia
.
Qed
.
Lemma
gg_tower
k
i
(
X
:
tower
)
:
gg
i
(
X
(
i
+
k
))
≡
X
k
.
Proof
.
by
induction
i
as
[|
i
IH
]
;
simpl
;
[
done
|
rewrite
g_tower
,
IH
].
Qed
.
Proof
.
by
induction
i
as
[|
i
IH
]
;
simpl
;
[
done
|
rewrite
g_tower
IH
].
Qed
.
Instance
tower_car_ne
n
k
:
Proper
(
dist
n
==>
dist
n
)
(
λ
X
,
tower_car
X
k
).
Proof
.
by
intros
X
Y
HX
.
Qed
.
...
...
@@ -114,14 +113,14 @@ Lemma gg_gg {k i i1 i2 j} (H1 : k = i + j) (H2 : k = i2 + (i1 + j)) (x : A k) :
Proof
.
assert
(
i
=
i2
+
i1
)
by
lia
;
simplify_equality'
.
revert
j
x
H1
.
induction
i2
as
[|
i2
IH
]
;
intros
j
X
H1
;
simplify_equality'
;
[
by
rewrite
coerce_id
|
by
rewrite
g_coerce
,
IH
].
[
by
rewrite
coerce_id
|
by
rewrite
g_coerce
IH
].
Qed
.
Lemma
ff_ff
{
k
i
i1
i2
j
}
(
H1
:
i
+
k
=
j
)
(
H2
:
i1
+
(
i2
+
k
)
=
j
)
(
x
:
A
k
)
:
coerce
H1
(
ff
i
x
)
=
coerce
H2
(
ff
i1
(
ff
i2
x
)).
Proof
.
assert
(
i
=
i1
+
i2
)
by
lia
;
simplify_equality'
.
induction
i1
as
[|
i1
IH
]
;
simplify_equality'
;
[
by
rewrite
coerce_id
|
by
rewrite
coerce_f
,
IH
].
[
by
rewrite
coerce_id
|
by
rewrite
coerce_f
IH
].
Qed
.
Definition
embed'
{
k
}
(
i
:
nat
)
:
A
k
-
n
>
A
i
:
=
...
...
@@ -135,10 +134,10 @@ Proof.
*
symmetry
;
by
erewrite
(@
gg_gg
_
_
1
(
k
-
S
i
))
;
simpl
.
*
exfalso
;
lia
.
*
assert
(
i
=
k
)
by
lia
;
subst
.
rewrite
(
ff_ff
_
(
eq_refl
(
1
+
(
0
+
k
))))
;
simpl
;
rewrite
gf
.
rewrite
(
ff_ff
_
(
eq_refl
(
1
+
(
0
+
k
))))
/=
gf
.
by
rewrite
(
gg_gg
_
(
eq_refl
(
0
+
(
0
+
k
)))).
*
assert
(
H
:
1
+
((
i
-
k
)
+
k
)
=
S
i
)
by
lia
;
rewrite
(
ff_ff
_
H
)
;
simpl
.
rewrite
<-(
gf
(
ff
(
i
-
k
)
x
))
at
2
;
rewrite
g_coerce
.
*
assert
(
H
:
1
+
((
i
-
k
)
+
k
)
=
S
i
)
by
lia
.
rewrite
(
ff_ff
_
H
)
/=
-{
2
}(
gf
(
ff
(
i
-
k
)
x
))
g_coerce
.
by
erewrite
coerce_proper
by
done
.
Qed
.
Program
Definition
embed_inf
(
k
:
nat
)
(
x
:
A
k
)
:
T
:
=
...
...
@@ -152,21 +151,21 @@ Proof.
rewrite
equiv_dist
;
intros
n
i
.
unfold
embed_inf
,
embed
;
simpl
;
unfold
embed'
.
destruct
(
le_lt_dec
i
(
S
k
)),
(
le_lt_dec
i
k
)
;
simpl
.
*
assert
(
H
:
S
k
=
S
(
k
-
i
)
+
(
0
+
i
))
by
lia
;
rewrite
(
gg_gg
_
H
)
;
simpl
.
*
assert
(
H
:
S
k
=
S
(
k
-
i
)
+
(
0
+
i
))
by
lia
;
rewrite
(
gg_gg
_
H
)
/=
.
by
erewrite
g_coerce
,
gf
,
coerce_proper
by
done
.
*
assert
(
S
k
=
0
+
(
0
+
i
))
as
H
by
lia
.
rewrite
(
gg_gg
_
H
)
;
simplify_equality'
.
by
rewrite
(
ff_ff
_
(
eq_refl
(
1
+
(
0
+
k
)))).
*
exfalso
;
lia
.
*
assert
(
H
:
(
i
-
S
k
)
+
(
1
+
k
)
=
i
)
by
lia
;
rewrite
(
ff_ff
_
H
)
;
simpl
.
*
assert
(
H
:
(
i
-
S
k
)
+
(
1
+
k
)
=
i
)
by
lia
;
rewrite
(
ff_ff
_
H
)
/=
.
by
erewrite
coerce_proper
by
done
.
Qed
.
Lemma
embed_tower
j
n
(
X
:
T
)
:
n
≤
j
→
embed
j
(
X
j
)
={
n
}=
X
.
Proof
.
intros
Hn
i
;
simpl
;
unfold
embed'
;
destruct
(
le_lt_dec
i
j
)
as
[
H
|
H
]
;
simpl
.
*
rewrite
<
-(
gg_tower
i
(
j
-
i
)
X
).
*
rewrite
-(
gg_tower
i
(
j
-
i
)
X
).
apply
(
_
:
Proper
(
_
==>
_
)
(
gg
_
))
;
by
destruct
(
eq_sym
_
).
*
rewrite
(
ff_tower
j
(
i
-
j
)
X
)
by
lia
;
by
destruct
(
Nat
.
sub_add
_
_
_
).
*
rewrite
(
ff_tower
j
(
i
-
j
)
X
)
;
last
lia
.
by
destruct
(
Nat
.
sub_add
_
_
_
).
Qed
.
Program
Definition
unfold_chain
(
X
:
T
)
:
chain
(
F
T
T
)
:
=
...
...
@@ -175,14 +174,14 @@ Next Obligation.
intros
X
n
i
Hi
.
assert
(
∃
k
,
i
=
k
+
n
)
as
[
k
?]
by
(
exists
(
i
-
n
)
;
lia
)
;
subst
;
clear
Hi
.
induction
k
as
[|
k
Hk
]
;
simpl
;
[
done
|].
rewrite
Hk
,
(
f_tower
X
),
f_S
,
<-
map_comp
by
lia
.
rewrite
Hk
(
f_tower
X
)
;
last
lia
;
rewrite
f_S
-
map_comp
.
apply
dist_S
,
map_contractive
.
split
;
intros
Y
;
symmetry
;
apply
equiv_dist
;
[
apply
g_tower
|
apply
embed_f
].
Qed
.
Definition
unfold'
(
X
:
T
)
:
F
T
T
:
=
compl
(
unfold_chain
X
).
Instance
unfold_ne
:
Proper
(
dist
n
==>
dist
n
)
unfold'
.
Proof
.
by
intros
n
X
Y
HXY
;
unfold
unfold'
;
apply
compl_ne
;
simpl
;
rewrite
(
HXY
n
).
by
intros
n
X
Y
HXY
;
unfold
unfold'
;
apply
compl_ne
;
rewrite
/=
(
HXY
n
).
Qed
.
Definition
unfold
:
T
-
n
>
F
T
T
:
=
CofeMor
unfold'
.
...
...
@@ -190,7 +189,7 @@ Program Definition fold' (X : F T T) : T :=
{|
tower_car
n
:
=
g
(
map
(
embed
n
,
project
n
)
X
)
|}.
Next
Obligation
.
intros
X
k
;
apply
(
_
:
Proper
((
≡
)
==>
(
≡
))
g
).
rewrite
<
-(
map_comp
_
_
_
_
_
_
(
embed
(
S
k
))
f
(
project
(
S
k
))
g
).
rewrite
-(
map_comp
_
_
_
_
_
_
(
embed
(
S
k
))
f
(
project
(
S
k
))
g
).
apply
map_ext
;
[
apply
embed_f
|
intros
Y
;
apply
g_tower
|
done
].
Qed
.
Instance
fold_ne
:
Proper
(
dist
n
==>
dist
n
)
fold'
.
...
...
@@ -202,14 +201,14 @@ Proof.
assert
(
map_ff_gg
:
∀
i
k
(
x
:
A
(
S
i
+
k
))
(
H
:
S
i
+
k
=
i
+
S
k
),
map
(
ff
i
,
gg
i
)
x
≡
gg
i
(
coerce
H
x
)).
{
intros
i
;
induction
i
as
[|
i
IH
]
;
intros
k
x
H
;
simpl
.
{
by
rewrite
coerce_id
,
map_id
.
}
rewrite
map_comp
,
g_coerce
;
apply
IH
.
}
{
by
rewrite
coerce_id
map_id
.
}
rewrite
map_comp
g_coerce
;
apply
IH
.
}
rewrite
equiv_dist
;
intros
n
k
;
unfold
unfold
,
fold
;
simpl
.
rewrite
<-
g_tower
,
<
-(
gg_tower
_
n
)
;
apply
(
_
:
Proper
(
_
==>
_
)
g
).
rewrite
-
g_tower
-(
gg_tower
_
n
)
;
apply
(
_
:
Proper
(
_
==>
_
)
g
).
transitivity
(
map
(
ff
n
,
gg
n
)
(
X
(
S
(
n
+
k
)))).
{
unfold
unfold'
;
rewrite
(
conv_compl
(
unfold_chain
X
)
n
).
rewrite
(
chain_cauchy
(
unfold_chain
X
)
n
(
n
+
k
))
by
lia
;
simpl
.
rewrite
(
f_tower
X
)
,
<-
map_comp
by
lia
.
rewrite
(
chain_cauchy
(
unfold_chain
X
)
n
(
n
+
k
))
/=
;
last
lia
.
rewrite
(
f_tower
X
)
;
last
lia
;
rewrite
-
map_comp
.
apply
dist_S
.
apply
map_contractive
;
split
;
intros
x
;
simpl
;
unfold
embed'
.
*
destruct
(
le_lt_dec
_
_
)
;
simpl
.
{
assert
(
n
=
0
)
by
lia
;
subst
.
apply
dist_0
.
}
...
...
@@ -221,12 +220,12 @@ Proof.
Qed
.
Definition
unfold_fold
X
:
unfold
(
fold
X
)
≡
X
.
Proof
.
rewrite
equiv_dist
;
intros
n
;
unfold
unfold
;
simpl
.
unfold
unfold'
;
rewrite
(
conv_compl
(
unfold_chain
(
fold
X
))
n
)
;
simpl
.
rewrite
(
fg
(
map
(
embed
_
,
project
n
)
X
))
,
<-
map_comp
by
lia
;
rewrite
<-(
map_id
_
_
X
)
at
2
.
rewrite
equiv_dist
;
intros
n
.
rewrite
/(
unfold
)
/=
/(
unfold'
)
(
conv_compl
(
unfold_chain
(
fold
X
))
n
)
/=
.
rewrite
(
fg
(
map
(
embed
_
,
project
n
)
X
))
;
last
lia
.
rewrite
-
map_comp
-{
2
}(
map_id
_
_
X
)
.
apply
dist_S
,
map_contractive
;
split
;
intros
Y
i
;
apply
embed_tower
;
lia
.
Qed
.
End
solver
.
Global
Opaque
cofe_solver
.
T
cofe_solver
.
fold
cofe_solver
.
unfold
.
Global
Opaque
T
fold
unfold
.
modures/dra.v
View file @
9df894ee
...
...
@@ -67,19 +67,17 @@ Hint Immediate dra_op_proper : typeclass_instances.
Instance
:
Proper
((
≡
)
==>
(
≡
)
==>
iff
)
(
⊥
).
Proof
.
intros
x1
x2
Hx
y1
y2
Hy
;
split
.
*
by
rewrite
Hy
,
(
symmetry_iff
(
⊥
)
x1
),
(
symmetry_iff
(
⊥
)
x2
),
Hx
.
*
by
rewrite
<-
Hy
,
(
symmetry_iff
(
⊥
)
x2
),
(
symmetry_iff
(
⊥
)
x1
),
<
-
Hx
.
*
by
rewrite
Hy
(
symmetry_iff
(
⊥
)
x1
)
(
symmetry_iff
(
⊥
)
x2
)
Hx
.
*
by
rewrite
-
Hy
(
symmetry_iff
(
⊥
)
x2
)
(
symmetry_iff
(
⊥
)
x1
)
-
Hx
.
Qed
.
Lemma
dra_disjoint_rl
x
y
z
:
✓
x
→
✓
y
→
✓
z
→
y
⊥
z
→
x
⊥
y
⋅
z
→
x
⊥
y
.
Proof
.
intros
???.
rewrite
!(
symmetry_iff
_
x
).
by
apply
dra_disjoint_ll
.
Qed
.
Lemma
dra_disjoint_lr
x
y
z
:
✓
x
→
✓
y
→
✓
z
→
x
⊥
y
→
x
⋅
y
⊥
z
→
y
⊥
z
.
Proof
.
intros
????.
rewrite
dra_commutative
by
done
.
by
apply
dra_disjoint_ll
.
Qed
.
Proof
.
intros
????.
rewrite
dra_commutative
//.
by
apply
dra_disjoint_ll
.
Qed
.
Lemma
dra_disjoint_move_r
x
y
z
:
✓
x
→
✓
y
→
✓
z
→
y
⊥
z
→
x
⊥
y
⋅
z
→
x
⋅
y
⊥
z
.
Proof
.
intros
;
symmetry
;
rewrite
dra_commutative
by
eauto
using
dra_disjoint_rl
.
intros
;
symmetry
;
rewrite
dra_commutative
;
eauto
using
dra_disjoint_rl
.
apply
dra_disjoint_move_l
;
auto
;
by
rewrite
dra_commutative
.
Qed
.
Hint
Immediate
dra_disjoint_move_l
dra_disjoint_move_r
.
...
...
@@ -100,20 +98,21 @@ Program Instance validity_minus : Minus T := λ x y,
Validity
(
validity_car
x
⩪
validity_car
y
)
(
✓
x
∧
✓
y
∧
validity_car
y
≼
validity_car
x
)
_
.
Solve
Obligations
with
naive_solver
auto
using
dra_minus_valid
.
Instance
validity_ra
:
RA
T
.
Proof
.
split
.
*
apply
_
.
*
intros
???
[?
Heq
]
;
split
;
simpl
;
[|
by
intros
(?&?&?)
;
rewrite
Heq
].
split
;
intros
(?&?&?)
;
split_ands'
;
first
[
by
rewrite
?Heq
by
tauto
|
by
rewrite
<-
?Heq
by
tauto
|
tauto
].
first
[
rewrite
?Heq
;
tauto
|
rewrite
-
?Heq
;
tauto
|
tauto
].
*
by
intros
??
[?
Heq
]
;
split
;
[
done
|]
;
simpl
;
intros
?
;
rewrite
Heq
.
*
by
intros
??
Heq
?
;
rewrite
<-
Heq
.
*
by
intros
??
->
?
.
*
intros
x1
x2
[?
Hx
]
y1
y2
[?
Hy
]
;
split
;
simpl
;
[|
by
intros
(?&?&?)
;
rewrite
Hx
,
Hy
].
split
;
simpl
;
[|
by
intros
(?&?&?)
;
rewrite
Hx
//
Hy
].
split
;
intros
(?&?&
z
&?&?)
;
split_ands'
;
try
tauto
.
+
exists
z
.
by
rewrite
<-
Hy
,
<
-
Hx
.
+
exists
z
.
by
rewrite
Hx
,
Hy
by
tauto
.
+
exists
z
.
by
rewrite
-
Hy
//
-
Hx
.
+
exists
z
.
by
rewrite
Hx
?Hy
;
tauto
.
*
intros
[
x
px
?]
[
y
py
?]
[
z
pz
?]
;
split
;
simpl
;
[
intuition
eauto
2
using
dra_disjoint_lr
,
dra_disjoint_rl
|
intros
;
apply
(
associative
_
)].
...
...
modures/excl.v
View file @
9df894ee
...
...
@@ -61,13 +61,15 @@ Proof.
feed
inversion
(
chain_cauchy
c
1
n
)
;
auto
with
lia
;
constructor
.
destruct
(
c
1
)
;
simplify_equality'
.
Qed
.
Instance
Excl_timeless
`
{
Equiv
A
,
Dist
A
}
(
x
:
excl
A
)
:
Timeless
x
→
Timeless
(
Excl
x
).
Instance
Excl_timeless
`
{
Equiv
A
,
Dist
A
}
(
x
:
A
)
:
Timeless
x
→
Timeless
(
Excl
x
).
Proof
.
by
inversion_clear
2
;
constructor
;
apply
(
timeless
_
).
Qed
.
Instance
ExclUnit_timeless
`
{
Equiv
A
,
Dist
A
}
:
Timeless
(@
ExclUnit
A
).
Proof
.
by
inversion_clear
1
;
constructor
.
Qed
.
Instance
ExclBot_timeless
`
{
Equiv
A
,
Dist
A
}
:
Timeless
(@
ExclBot
A
).
Proof
.
by
inversion_clear
1
;
constructor
.
Qed
.
Instance
excl_timeless
`
{
Equiv
A
,
Dist
A
}
:
(
∀
x
:
A
,
Timeless
x
)
→
∀
x
:
excl
A
,
Timeless
x
.
Proof
.
intros
?
[]
;
apply
_
.
Qed
.
(* CMRA *)
Instance
excl_valid
{
A
}
:
Valid
(
excl
A
)
:
=
λ
x
,
...
...
modures/fin_maps.v
View file @
9df894ee
...
...
@@ -56,15 +56,16 @@ Global Instance map_lookup_timeless `{Cofe A} (m : gmap K A) i :
Timeless
m
→
Timeless
(
m
!!
i
).
Proof
.
intros
?
[
x
|]
Hx
;
[|
by
symmetry
;
apply
(
timeless
_
)].
rewrite
(
timeless
m
(<[
i
:
=
x
]>
m
)),
lookup_insert
;
auto
.
by
symmetry
in
Hx
;
inversion
Hx
;
cofe_subst
;
rewrite
insert_id
.
assert
(
m
={
1
}=
<[
i
:
=
x
]>
m
)
by
(
by
symmetry
in
Hx
;
inversion
Hx
;
cofe_subst
;
rewrite
insert_id
).
by
rewrite
(
timeless
m
(<[
i
:
=
x
]>
m
))
//
lookup_insert
.
Qed
.
Global
Instance
map_ra_insert_timeless
`
{
Cofe
A
}
(
m
:
gmap
K
A
)
i
x
:
Timeless
x
→
Timeless
m
→
Timeless
(<[
i
:
=
x
]>
m
).
Proof
.
intros
??
m'
Hm
j
;
destruct
(
decide
(
i
=
j
))
;
simplify_map_equality
.