Skip to content
GitLab
Explore
Sign in
Primary navigation
Search or go to…
Project
Iris
Manage
Activity
Members
Labels
Plan
Issues
Issue boards
Milestones
Wiki
Code
Merge requests
Repository
Branches
Commits
Tags
Repository graph
Compare revisions
Build
Pipelines
Jobs
Pipeline schedules
Artifacts
Deploy
Releases
Package registry
Model registry
Operate
Terraform modules
Monitor
Service Desk
Analyze
Value stream analytics
Contributor analytics
CI/CD analytics
Repository analytics
Model experiments
Help
Help
Support
GitLab documentation
Compare GitLab plans
Community forum
Contribute to GitLab
Provide feedback
Terms and privacy
Keyboard shortcuts
?
Snippets
Groups
Projects
Show more breadcrumbs
Pierre Roux
Iris
Commits
12796181
Commit
12796181
authored
4 years ago
by
Robbert Krebbers
Browse files
Options
Downloads
Patches
Plain Diff
Add documentation for auth & view.
parent
f98235de
No related branches found
No related tags found
No related merge requests found
Changes
2
Hide whitespace changes
Inline
Side-by-side
Showing
2 changed files
theories/algebra/auth.v
+15
-13
15 additions, 13 deletions
theories/algebra/auth.v
theories/algebra/view.v
+45
-16
45 additions, 16 deletions
theories/algebra/view.v
with
60 additions
and
29 deletions
theories/algebra/auth.v
+
15
−
13
View file @
12796181
...
@@ -3,23 +3,25 @@ From iris.algebra Require Import proofmode_classes.
...
@@ -3,23 +3,25 @@ From iris.algebra Require Import proofmode_classes.
From
iris
.
base_logic
Require
Import
base_logic
.
From
iris
.
base_logic
Require
Import
base_logic
.
From
iris
Require
Import
options
.
From
iris
Require
Import
options
.
(** The authoritative camera with fractional authoritative elements. The camera
(** The authoritative camera with fractional authoritative elements *)
[auth] has 3 types of elements: the fractional authoritative element [●{q} a],
(** The authoritative camera has 2 types of elements: the authoritative
the full authoritative element [● a ≡ ●{1} a], and the fragment [◯ b] (of which
element [●{q} a] and the fragment [◯ b] (of which there can be several). To
there can be several). Updates are only possible with the full authoritative
enable sharing of the authoritative element [●{q} a], it is equiped with a
element [● a], while fractional authoritative elements have agreement, i.e.,
fraction [q]. Updates are only possible with the full authoritative element
[✓ (●{p1} a1 ⋅ ●{p2} a2) → a1 ≡ a2]. *)
[● a] (syntax for [●{1} a]]), while fractional authoritative elements have
agreement, i.e., [✓ (●{p1} a1 ⋅ ●{p2} a2) → a1 ≡ a2]. *)
(** * Definition of the view relation *)
(** * Definition of the view relation *)
(** The authoritative camera is obtained by instantiating the view camera. *)
Definition
auth_view_rel_raw
{
A
:
ucmraT
}
(
n
:
nat
)
(
a
b
:
A
)
:
Prop
:=
Definition
auth_view_rel_raw
{
A
:
ucmraT
}
(
n
:
nat
)
(
a
b
:
A
)
:
Prop
:=
b
≼
{
n
}
a
∧
✓
{
n
}
a
.
b
≼
{
n
}
a
∧
✓
{
n
}
a
.
Lemma
auth_view_rel_raw_mono
(
A
:
ucmraT
)
n1
n2
(
a1
a2
b1
b2
:
A
)
:
Lemma
auth_view_rel_raw_mono
(
A
:
ucmraT
)
n1
n2
(
a1
a2
b1
b2
:
A
)
:
auth_view_rel_raw
n1
a1
b1
→
a1
≡
{
n
1
}
≡
a2
→
b2
≼
{
n
1
}
b1
→
n2
≤
n1
→
auth_view_rel_raw
n1
a1
b1
→
a1
≡
{
n
2
}
≡
a2
→
b2
≼
{
n
2
}
b1
→
n2
≤
n1
→
auth_view_rel_raw
n2
a2
b2
.
auth_view_rel_raw
n2
a2
b2
.
Proof
.
Proof
.
intros
[??]
Ha12
??
.
split
.
intros
[??]
Ha12
??
.
split
.
-
apply
cmra_includedN_le
with
n
1
;
[
|
done
]
.
rewrite
-
Ha12
.
by
trans
b
1
.
-
trans
b
1
;
[
done
|
]
.
rewrite
-
Ha12
.
by
apply
cmra_includedN_le
with
n
1
.
-
apply
cmra_validN_le
with
n1
;
[|
lia
]
.
by
rewrite
-
Ha12
.
-
rewrite
-
Ha12
.
by
apply
cmra_validN_le
with
n1
.
Qed
.
Qed
.
Lemma
auth_view_rel_raw_valid
(
A
:
ucmraT
)
n
(
a
b
:
A
)
:
Lemma
auth_view_rel_raw_valid
(
A
:
ucmraT
)
n
(
a
b
:
A
)
:
auth_view_rel_raw
n
a
b
→
✓
{
n
}
b
.
auth_view_rel_raw
n
a
b
→
✓
{
n
}
b
.
...
@@ -38,7 +40,7 @@ Proof.
...
@@ -38,7 +40,7 @@ Proof.
-
by
apply
cmra_discrete_valid_iff_0
.
-
by
apply
cmra_discrete_valid_iff_0
.
Qed
.
Qed
.
(** * Definition and operations on
au
th *)
(** * Definition and operations on th
e authoritative camera
*)
(** The type [auth] is not defined as a [Definition], but as a [Notation].
(** The type [auth] is not defined as a [Definition], but as a [Notation].
This way, one can use [auth A] with [A : Type] instead of [A : ucmraT], and let
This way, one can use [auth A] with [A : Type] instead of [A : ucmraT], and let
canonical structure search determine the corresponding camera instance. *)
canonical structure search determine the corresponding camera instance. *)
...
@@ -59,7 +61,7 @@ Notation "◯ a" := (auth_frag a) (at level 20).
...
@@ -59,7 +61,7 @@ Notation "◯ a" := (auth_frag a) (at level 20).
Notation
"●{ q } a"
:=
(
auth_auth
q
a
)
(
at
level
20
,
format
"●{ q } a"
)
.
Notation
"●{ q } a"
:=
(
auth_auth
q
a
)
(
at
level
20
,
format
"●{ q } a"
)
.
Notation
"● a"
:=
(
auth_auth
1
a
)
(
at
level
20
)
.
Notation
"● a"
:=
(
auth_auth
1
a
)
(
at
level
20
)
.
(** * Laws *)
(** * Laws
of the authoritative camera
*)
Section
auth
.
Section
auth
.
Context
{
A
:
ucmraT
}
.
Context
{
A
:
ucmraT
}
.
Implicit
Types
a
b
:
A
.
Implicit
Types
a
b
:
A
.
...
@@ -210,8 +212,8 @@ Section auth.
...
@@ -210,8 +212,8 @@ Section auth.
◯
b1
≼
●
{
p
}
a
⋅
◯
b2
↔
b1
≼
b2
.
◯
b1
≼
●
{
p
}
a
⋅
◯
b2
↔
b1
≼
b2
.
Proof
.
apply
view_frag_included
.
Qed
.
Proof
.
apply
view_frag_included
.
Qed
.
(** The weaker [
view
_both_included] lemmas below are a consequence of the
(** The weaker [
auth
_both_included] lemmas below are a consequence of the
[
view
_auth_included] and [
view
_frag_included] lemmas above. *)
[
auth
_auth_included] and [
auth
_frag_included] lemmas above. *)
Lemma
auth_both_includedN
n
p1
p2
a1
a2
b1
b2
:
Lemma
auth_both_includedN
n
p1
p2
a1
a2
b1
b2
:
●
{
p1
}
a1
⋅
◯
b1
≼
{
n
}
●
{
p2
}
a2
⋅
◯
b2
↔
(
p1
≤
p2
)
%
Qc
∧
a1
≡
{
n
}
≡
a2
∧
b1
≼
{
n
}
b2
.
●
{
p1
}
a1
⋅
◯
b1
≼
{
n
}
●
{
p2
}
a2
⋅
◯
b2
↔
(
p1
≤
p2
)
%
Qc
∧
a1
≡
{
n
}
≡
a2
∧
b1
≼
{
n
}
b2
.
Proof
.
apply
view_both_includedN
.
Qed
.
Proof
.
apply
view_both_includedN
.
Qed
.
...
...
This diff is collapsed.
Click to expand it.
theories/algebra/view.v
+
45
−
16
View file @
12796181
...
@@ -4,23 +4,52 @@ From iris.algebra Require Import proofmode_classes.
...
@@ -4,23 +4,52 @@ From iris.algebra Require Import proofmode_classes.
From
iris
.
base_logic
Require
Import
base_logic
.
From
iris
.
base_logic
Require
Import
base_logic
.
From
iris
Require
Import
options
.
From
iris
Require
Import
options
.
(** The view camera with fractional authoritative parts. The camera [view] has
(** The view camera with fractional authoritative elements *)
3 types of elements: the fractional authoritative [●V{q} a], the full
(** The view camera, which is reminiscent of the views framework, is used to
authoritative [●V a ≡ ●V{1} a], and the non-authoritative [◯V a]. Updates are
provide a logical/"small-footprint" "view" of some "large-footprint" piece of
only possible with the full authoritative element [●V a], while fractional
data, which can be shared in the separation logic sense, i.e., different parts
authoritative elements have agreement: [✓ (●V{p} a ⋅ ●V{q} b) → a ≡ b]. *)
of the data can be separately owned by different functions or threads. This is
achieved using the two elements of the view camera:
- The authoritative element [●V a], which describes the data under consideration.
- The fragment [◯V b], which provides a logical view of the data [a].
To enable sharing of the fragments, the type of fragments is equipped with a
camera structure so ownership of fragments can be split. Concretely, fragments
enjoy the rule [◯V (b1 ⋅ b2) = ◯V b1 ⋅ ◯V b2].
To enable sharing of the authoritative element [●V{q} a], it is equipped with a
fraction [q]. Updates are only possible with the full authoritative element
[●V a] (syntax for [●V{1} a]]), while fractional authoritative elements have
agreement, i.e., [✓ (●V{p1} a1 ⋅ ●V{p2} a2) → a1 ≡ a2]. *)
(** * The view relation *)
(** * The view relation *)
(** The view camera is parametrized by a (step-indexed) relation
(** To relate the authoritative element [a] to its possible fragments [b], the
[view_rel n a b] that relates the authoritative element [a] to the composition
view camera is parametrized by a (step-indexed) relation [view_rel n a b]. This
of all fragments [b]. This relation should be a.) down-closed w.r.t. the
relation should be a.) closed under smaller step-indexes [n], b.) non-expansive
step-indexed [n], b.) non-expansive w.r.t. the argument [a], c.) monotone w.r.t.
w.r.t. the argument [a], c.) closed under smaller [b] (which implies
the argument [b], d.) ensure validity of the argument [b]. *)
non-expansiveness w.r.t. [b]), and d.) ensure validity of the argument [b].
Note 1: Instead of requiring both a step-indexed and a non-step-indexed version
of the relation (like cameras do for validity), we use [∀ n, view_rel n] as the
non-step-indexed version. This is anyway necessary when using [≼{n}] as the
relation (like the authoritative camera does) as its non-step-indexed version
is not equivalent to [∀ n, x ≼{n} y].
Note 2: The view relation is defined as a canonical structure so that given a
relation [nat → A → B → Prop], the instance with the laws can be inferred. We do
not use type classes for this purpose because cameras themselves are represented
using canonical structures. It has proven fragile for a canonical structure
instance to take a type class as a parameter (in this case, [viewR] would need
to take a class with the view relation laws). *)
Structure
view_rel
(
A
:
ofeT
)
(
B
:
ucmraT
)
:=
ViewRel
{
Structure
view_rel
(
A
:
ofeT
)
(
B
:
ucmraT
)
:=
ViewRel
{
view_rel_holds
:>
nat
→
A
→
B
→
Prop
;
view_rel_holds
:>
nat
→
A
→
B
→
Prop
;
view_rel_mono
n1
n2
a1
a2
b1
b2
:
view_rel_mono
n1
n2
a1
a2
b1
b2
:
view_rel_holds
n1
a1
b1
→
view_rel_holds
n1
a1
b1
→
a1
≡
{
n1
}
≡
a2
→
b2
≼
{
n1
}
b1
→
n2
≤
n1
→
view_rel_holds
n2
a2
b2
;
a1
≡
{
n2
}
≡
a2
→
b2
≼
{
n2
}
b1
→
n2
≤
n1
→
view_rel_holds
n2
a2
b2
;
view_rel_validN
n
a
b
:
view_rel_validN
n
a
b
:
view_rel_holds
n
a
b
→
✓
{
n
}
b
view_rel_holds
n
a
b
→
✓
{
n
}
b
}
.
}
.
...
@@ -64,7 +93,7 @@ Notation "●V{ q } a" := (view_auth q a) (at level 20, format "●V{ q } a").
...
@@ -64,7 +93,7 @@ Notation "●V{ q } a" := (view_auth q a) (at level 20, format "●V{ q } a").
Notation
"●V a"
:=
(
view_auth
1
a
)
(
at
level
20
)
.
Notation
"●V a"
:=
(
view_auth
1
a
)
(
at
level
20
)
.
Notation
"◯V a"
:=
(
view_frag
a
)
(
at
level
20
)
.
Notation
"◯V a"
:=
(
view_frag
a
)
(
at
level
20
)
.
(** *
Of
e *)
(** *
The OFE structur
e *)
Section
ofe
.
Section
ofe
.
Context
{
A
B
:
ofeT
}
(
rel
:
nat
→
A
→
B
→
Prop
)
.
Context
{
A
B
:
ofeT
}
(
rel
:
nat
→
A
→
B
→
Prop
)
.
Implicit
Types
a
:
A
.
Implicit
Types
a
:
A
.
...
@@ -111,7 +140,7 @@ Section ofe.
...
@@ -111,7 +140,7 @@ Section ofe.
Proof
.
by
uPred
.
unseal
.
Qed
.
Proof
.
by
uPred
.
unseal
.
Qed
.
End
ofe
.
End
ofe
.
(** *
Camera
*)
(** *
The camera structure
*)
Section
cmra
.
Section
cmra
.
Context
{
A
B
}
(
rel
:
view_rel
A
B
)
.
Context
{
A
B
}
(
rel
:
view_rel
A
B
)
.
Implicit
Types
a
:
A
.
Implicit
Types
a
:
A
.
...
@@ -161,14 +190,14 @@ Section cmra.
...
@@ -161,14 +190,14 @@ Section cmra.
Instance
view_op
:
Op
(
view
rel
)
:=
λ
x
y
,
Instance
view_op
:
Op
(
view
rel
)
:=
λ
x
y
,
View
(
view_auth_proj
x
⋅
view_auth_proj
y
)
(
view_frag_proj
x
⋅
view_frag_proj
y
)
.
View
(
view_auth_proj
x
⋅
view_auth_proj
y
)
(
view_frag_proj
x
⋅
view_frag_proj
y
)
.
Definition
view_valid_eq
:
Local
Definition
view_valid_eq
:
valid
=
λ
x
,
valid
=
λ
x
,
match
view_auth_proj
x
with
match
view_auth_proj
x
with
|
Some
(
q
,
ag
)
=>
|
Some
(
q
,
ag
)
=>
✓
q
∧
(
∀
n
,
∃
a
,
ag
≡
{
n
}
≡
to_agree
a
∧
rel
n
a
(
view_frag_proj
x
))
✓
q
∧
(
∀
n
,
∃
a
,
ag
≡
{
n
}
≡
to_agree
a
∧
rel
n
a
(
view_frag_proj
x
))
|
None
=>
✓
view_frag_proj
x
|
None
=>
✓
view_frag_proj
x
end
:=
eq_refl
_
.
end
:=
eq_refl
_
.
Definition
view_validN_eq
:
Local
Definition
view_validN_eq
:
validN
=
λ
n
x
,
validN
=
λ
n
x
,
match
view_auth_proj
x
with
match
view_auth_proj
x
with
|
Some
(
q
,
ag
)
=>
✓
{
n
}
q
∧
∃
a
,
ag
≡
{
n
}
≡
to_agree
a
∧
rel
n
a
(
view_frag_proj
x
)
|
Some
(
q
,
ag
)
=>
✓
{
n
}
q
∧
∃
a
,
ag
≡
{
n
}
≡
to_agree
a
∧
rel
n
a
(
view_frag_proj
x
)
...
@@ -473,7 +502,7 @@ Section cmra.
...
@@ -473,7 +502,7 @@ Section cmra.
Qed
.
Qed
.
End
cmra
.
End
cmra
.
(** *
F
unctor *)
(** *
Utilities to construct f
unctor
s
*)
(** Due to the dependent type [rel] in [view] we cannot actually define
(** Due to the dependent type [rel] in [view] we cannot actually define
instances of the functor structures [rFunctor] and [urFunctor]. Functors can
instances of the functor structures [rFunctor] and [urFunctor]. Functors can
only be defined for instances of [view], like [auth]. To make it more convenient
only be defined for instances of [view], like [auth]. To make it more convenient
...
...
This diff is collapsed.
Click to expand it.
Preview
0%
Loading
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!
Save comment
Cancel
Please
register
or
sign in
to comment