Skip to content
GitLab
Explore
Sign in
Primary navigation
Search or go to…
Project
I
Iris
Manage
Activity
Members
Labels
Plan
Issues
Issue boards
Milestones
Wiki
Code
Merge requests
Repository
Branches
Commits
Tags
Repository graph
Compare revisions
Snippets
Build
Pipelines
Jobs
Pipeline schedules
Artifacts
Deploy
Releases
Model registry
Operate
Environments
Monitor
Incidents
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
Rodolphe Lepigre
Iris
Commits
a20c86e3
Commit
a20c86e3
authored
8 years ago
by
Robbert Krebbers
Browse files
Options
Downloads
Patches
Plain Diff
More consistent names for (local) updates.
parent
aa6c43a2
No related branches found
No related tags found
No related merge requests found
Changes
4
Hide whitespace changes
Inline
Side-by-side
Showing
4 changed files
algebra/cmra.v
+4
-4
4 additions, 4 deletions
algebra/cmra.v
algebra/gmap.v
+9
-9
9 additions, 9 deletions
algebra/gmap.v
algebra/list.v
+4
-4
4 additions, 4 deletions
algebra/list.v
program_logic/ghost_ownership.v
+1
-1
1 addition, 1 deletion
program_logic/ghost_ownership.v
with
18 additions
and
18 deletions
algebra/cmra.v
+
4
−
4
View file @
a20c86e3
...
@@ -535,10 +535,10 @@ Proof.
...
@@ -535,10 +535,10 @@ Proof.
by
rewrite
cmra_valid_validN
equiv_dist
=>??
n
;
apply
(
local_updateN
L
)
.
by
rewrite
cmra_valid_validN
equiv_dist
=>??
n
;
apply
(
local_updateN
L
)
.
Qed
.
Qed
.
Global
Instance
local_update
_op
x
:
LocalUpdate
(
λ
_,
True
)
(
op
x
)
.
Global
Instance
op_
local_update
x
:
LocalUpdate
(
λ
_,
True
)
(
op
x
)
.
Proof
.
split
.
apply
_
.
by
intros
n
y1
y2
_
_;
rewrite
assoc
.
Qed
.
Proof
.
split
.
apply
_
.
by
intros
n
y1
y2
_
_;
rewrite
assoc
.
Qed
.
Global
Instance
local_update
_id
:
LocalUpdate
(
λ
_,
True
)
(
@
id
A
)
.
Global
Instance
id_
local_update
:
LocalUpdate
(
λ
_,
True
)
(
@
id
A
)
.
Proof
.
split
;
auto
with
typeclass_instances
.
Qed
.
Proof
.
split
;
auto
with
typeclass_instances
.
Qed
.
Global
Instance
exclusive_local_update
y
:
Global
Instance
exclusive_local_update
y
:
...
@@ -1218,14 +1218,14 @@ Section option.
...
@@ -1218,14 +1218,14 @@ Section option.
Proof
.
intros
.
destruct
mx
;
apply
_
.
Qed
.
Proof
.
intros
.
destruct
mx
;
apply
_
.
Qed
.
(** Updates *)
(** Updates *)
Global
Instance
option_local
_fmap
_update
L
Lv
:
Global
Instance
option_
fmap_
local_update
L
Lv
:
LocalUpdate
Lv
L
→
LocalUpdate
Lv
L
→
LocalUpdate
(
λ
mx
,
if
mx
is
Some
x
then
Lv
x
else
False
)
(
fmap
L
)
.
LocalUpdate
(
λ
mx
,
if
mx
is
Some
x
then
Lv
x
else
False
)
(
fmap
L
)
.
Proof
.
Proof
.
split
;
first
apply
_
.
split
;
first
apply
_
.
intros
n
[
x
|]
[
z
|];
constructor
;
by
eauto
using
(
local_updateN
L
)
.
intros
n
[
x
|]
[
z
|];
constructor
;
by
eauto
using
(
local_updateN
L
)
.
Qed
.
Qed
.
Global
Instance
option_
local_const
_update
Lv
y
:
Global
Instance
option_
const_local
_update
Lv
y
:
LocalUpdate
Lv
(
λ
_,
y
)
→
LocalUpdate
Lv
(
λ
_,
y
)
→
LocalUpdate
(
λ
mx
,
if
mx
is
Some
x
then
Lv
x
else
False
)
(
λ
_,
Some
y
)
.
LocalUpdate
(
λ
mx
,
if
mx
is
Some
x
then
Lv
x
else
False
)
(
λ
_,
Some
y
)
.
Proof
.
Proof
.
...
...
This diff is collapsed.
Click to expand it.
algebra/gmap.v
+
9
−
9
View file @
a20c86e3
...
@@ -283,7 +283,7 @@ Proof. apply insert_update. Qed.
...
@@ -283,7 +283,7 @@ Proof. apply insert_update. Qed.
Section
freshness
.
Section
freshness
.
Context
`{
Fresh
K
(
gset
K
),
!
FreshSpec
K
(
gset
K
)}
.
Context
`{
Fresh
K
(
gset
K
),
!
FreshSpec
K
(
gset
K
)}
.
Lemma
updateP_
alloc_
strong
(
Q
:
gmap
K
A
→
Prop
)
(
I
:
gset
K
)
m
x
:
Lemma
alloc_
updateP_strong
(
Q
:
gmap
K
A
→
Prop
)
(
I
:
gset
K
)
m
x
:
✓
x
→
(
∀
i
,
m
!!
i
=
None
→
i
∉
I
→
Q
(
<
[
i
:=
x
]
>
m
))
→
m
~~>:
Q
.
✓
x
→
(
∀
i
,
m
!!
i
=
None
→
i
∉
I
→
Q
(
<
[
i
:=
x
]
>
m
))
→
m
~~>:
Q
.
Proof
.
Proof
.
intros
?
HQ
.
apply
cmra_total_updateP
.
intros
?
HQ
.
apply
cmra_total_updateP
.
...
@@ -297,15 +297,15 @@ Proof.
...
@@ -297,15 +297,15 @@ Proof.
last
by
apply
not_elem_of_dom
;
rewrite
dom_op
not_elem_of_union
.
last
by
apply
not_elem_of_dom
;
rewrite
dom_op
not_elem_of_union
.
by
apply
insert_validN
;
[
apply
cmra_valid_validN
|]
.
by
apply
insert_validN
;
[
apply
cmra_valid_validN
|]
.
Qed
.
Qed
.
Lemma
updateP
_alloc
(
Q
:
gmap
K
A
→
Prop
)
m
x
:
Lemma
alloc_
updateP
(
Q
:
gmap
K
A
→
Prop
)
m
x
:
✓
x
→
(
∀
i
,
m
!!
i
=
None
→
Q
(
<
[
i
:=
x
]
>
m
))
→
m
~~>:
Q
.
✓
x
→
(
∀
i
,
m
!!
i
=
None
→
Q
(
<
[
i
:=
x
]
>
m
))
→
m
~~>:
Q
.
Proof
.
move
=>??
.
eapply
updateP_
alloc_
strong
with
(
I
:=
∅
);
by
eauto
.
Qed
.
Proof
.
move
=>??
.
eapply
alloc_
updateP_strong
with
(
I
:=
∅
);
by
eauto
.
Qed
.
Lemma
updateP_
alloc_
strong'
m
x
(
I
:
gset
K
)
:
Lemma
alloc_
updateP_strong'
m
x
(
I
:
gset
K
)
:
✓
x
→
m
~~>:
λ
m'
,
∃
i
,
i
∉
I
∧
m'
=
<
[
i
:=
x
]
>
m
∧
m
!!
i
=
None
.
✓
x
→
m
~~>:
λ
m'
,
∃
i
,
i
∉
I
∧
m'
=
<
[
i
:=
x
]
>
m
∧
m
!!
i
=
None
.
Proof
.
eauto
using
updateP_
alloc_
strong
.
Qed
.
Proof
.
eauto
using
alloc_
updateP_strong
.
Qed
.
Lemma
updateP
_alloc
'
m
x
:
Lemma
alloc_
updateP'
m
x
:
✓
x
→
m
~~>:
λ
m'
,
∃
i
,
m'
=
<
[
i
:=
x
]
>
m
∧
m
!!
i
=
None
.
✓
x
→
m
~~>:
λ
m'
,
∃
i
,
m'
=
<
[
i
:=
x
]
>
m
∧
m
!!
i
=
None
.
Proof
.
eauto
using
updateP
_alloc
.
Qed
.
Proof
.
eauto
using
alloc_
updateP
.
Qed
.
Lemma
singleton_updateP_unit
(
P
:
A
→
Prop
)
(
Q
:
gmap
K
A
→
Prop
)
u
i
:
Lemma
singleton_updateP_unit
(
P
:
A
→
Prop
)
(
Q
:
gmap
K
A
→
Prop
)
u
i
:
✓
u
→
LeftId
(
≡
)
u
(
⋅
)
→
✓
u
→
LeftId
(
≡
)
u
(
⋅
)
→
...
@@ -335,7 +335,7 @@ End freshness.
...
@@ -335,7 +335,7 @@ End freshness.
(* Allocation is a local update: Just use composition with a singleton map. *)
(* Allocation is a local update: Just use composition with a singleton map. *)
Global
Instance
gmap_
delete_update
:
Global
Instance
delete_
local_
update
:
LocalUpdate
(
λ
m
,
∃
x
,
m
!!
i
=
Some
x
∧
Exclusive
x
)
(
delete
i
)
.
LocalUpdate
(
λ
m
,
∃
x
,
m
!!
i
=
Some
x
∧
Exclusive
x
)
(
delete
i
)
.
Proof
.
Proof
.
split
;
first
apply
_
.
split
;
first
apply
_
.
...
@@ -348,7 +348,7 @@ Proof.
...
@@ -348,7 +348,7 @@ Proof.
Qed
.
Qed
.
(* Applying a local update at a position we own is a local update. *)
(* Applying a local update at a position we own is a local update. *)
Global
Instance
gmap_
alter_update
`{
!
LocalUpdate
Lv
L
}
i
:
Global
Instance
alter
_local
_update
`{
!
LocalUpdate
Lv
L
}
i
:
LocalUpdate
(
λ
m
,
∃
x
,
m
!!
i
=
Some
x
∧
Lv
x
)
(
alter
L
i
)
.
LocalUpdate
(
λ
m
,
∃
x
,
m
!!
i
=
Some
x
∧
Lv
x
)
(
alter
L
i
)
.
Proof
.
Proof
.
split
;
first
apply
_
.
split
;
first
apply
_
.
...
...
This diff is collapsed.
Click to expand it.
algebra/list.v
+
4
−
4
View file @
a20c86e3
...
@@ -316,7 +316,7 @@ Section properties.
...
@@ -316,7 +316,7 @@ Section properties.
Proof
.
by
rewrite
!
persistent_total
list_core_singletonM
=>
->
.
Qed
.
Proof
.
by
rewrite
!
persistent_total
list_core_singletonM
=>
->
.
Qed
.
(* Update *)
(* Update *)
Lemma
list_
updat
e_updateP
(
P
:
A
→
Prop
)
(
Q
:
list
A
→
Prop
)
l1
x
l2
:
Lemma
list_
middl
e_updateP
(
P
:
A
→
Prop
)
(
Q
:
list
A
→
Prop
)
l1
x
l2
:
x
~~>:
P
→
(
∀
y
,
P
y
→
Q
(
l1
++
y
::
l2
))
→
l1
++
x
::
l2
~~>:
Q
.
x
~~>:
P
→
(
∀
y
,
P
y
→
Q
(
l1
++
y
::
l2
))
→
l1
++
x
::
l2
~~>:
Q
.
Proof
.
Proof
.
intros
Hx
%
option_updateP'
HP
.
intros
Hx
%
option_updateP'
HP
.
...
@@ -332,13 +332,13 @@ Section properties.
...
@@ -332,13 +332,13 @@ Section properties.
rewrite
!
list_lookup_op
!
lookup_app_r
!
app_length
//=
;
lia
.
rewrite
!
list_lookup_op
!
lookup_app_r
!
app_length
//=
;
lia
.
Qed
.
Qed
.
Lemma
list_
updat
e_update
l1
l2
x
y
:
x
~~>
y
→
l1
++
x
::
l2
~~>
l1
++
y
::
l2
.
Lemma
list_
middl
e_update
l1
l2
x
y
:
x
~~>
y
→
l1
++
x
::
l2
~~>
l1
++
y
::
l2
.
Proof
.
Proof
.
rewrite
!
cmra_update_updateP
=>
H
;
eauto
using
list_
updat
e_updateP
with
subst
.
rewrite
!
cmra_update_updateP
=>
H
;
eauto
using
list_
middl
e_updateP
with
subst
.
Qed
.
Qed
.
(* Applying a local update at a position we own is a local update. *)
(* Applying a local update at a position we own is a local update. *)
Global
Instance
list_alter_update
`{
LocalUpdate
A
Lv
L
}
i
:
Global
Instance
list_alter_
local_
update
`{
LocalUpdate
A
Lv
L
}
i
:
LocalUpdate
(
λ
L
,
∃
x
,
L
!!
i
=
Some
x
∧
Lv
x
)
(
alter
L
i
)
.
LocalUpdate
(
λ
L
,
∃
x
,
L
!!
i
=
Some
x
∧
Lv
x
)
(
alter
L
i
)
.
Proof
.
Proof
.
split
;
[
apply
_|];
intros
n
l1
l2
(
x
&
Hi1
&
?)
Hm
;
apply
list_dist_lookup
=>
j
.
split
;
[
apply
_|];
intros
n
l1
l2
(
x
&
Hi1
&
?)
Hm
;
apply
list_dist_lookup
=>
j
.
...
...
This diff is collapsed.
Click to expand it.
program_logic/ghost_ownership.v
+
1
−
1
View file @
a20c86e3
...
@@ -49,7 +49,7 @@ Proof.
...
@@ -49,7 +49,7 @@ Proof.
rewrite
-
(
pvs_mono
_
_
(
∃
m
,
■
(
∃
γ
,
γ
∉
G
∧
m
=
to_globalF
γ
a
)
∧
ownG
m
)
%
I
)
.
rewrite
-
(
pvs_mono
_
_
(
∃
m
,
■
(
∃
γ
,
γ
∉
G
∧
m
=
to_globalF
γ
a
)
∧
ownG
m
)
%
I
)
.
-
rewrite
ownG_empty
.
-
rewrite
ownG_empty
.
eapply
pvs_ownG_updateP
,
(
iprod_singleton_updateP_empty
(
inG_id
i
));
eapply
pvs_ownG_updateP
,
(
iprod_singleton_updateP_empty
(
inG_id
i
));
first
(
eapply
updateP_
alloc_
strong'
,
cmra_transport_valid
,
Ha
);
first
(
eapply
alloc_
updateP_strong'
,
cmra_transport_valid
,
Ha
);
naive_solver
.
naive_solver
.
-
apply
exist_elim
=>
m
;
apply
const_elim_l
=>
-
[
γ
[
Hfresh
->
]]
.
-
apply
exist_elim
=>
m
;
apply
const_elim_l
=>
-
[
γ
[
Hfresh
->
]]
.
by
rewrite
-
(
exist_intro
γ
)
const_equiv
//
left_id
.
by
rewrite
-
(
exist_intro
γ
)
const_equiv
//
left_id
.
...
...
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