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-coq
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
Joshua Yanovski
iris-coq
Commits
a20c86e3
Commit
a20c86e3
authored
Jun 16, 2016
by
Robbert Krebbers
Browse files
Options
Browse Files
Download
Email Patches
Plain Diff
More consistent names for (local) updates.
parent
aa6c43a2
Changes
4
Hide whitespace changes
Inline
Side-by-side
Showing
4 changed files
with
18 additions
and
18 deletions
+18
-18
algebra/cmra.v
algebra/cmra.v
+4
-4
algebra/gmap.v
algebra/gmap.v
+9
-9
algebra/list.v
algebra/list.v
+4
-4
program_logic/ghost_ownership.v
program_logic/ghost_ownership.v
+1
-1
No files found.
algebra/cmra.v
View file @
a20c86e3
...
...
@@ -535,10 +535,10 @@ Proof.
by
rewrite
cmra_valid_validN
equiv_dist
=>??
n
;
apply
(
local_updateN
L
).
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
.
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
.
Global
Instance
exclusive_local_update
y
:
...
...
@@ -1218,14 +1218,14 @@ Section option.
Proof
.
intros
.
destruct
mx
;
apply
_.
Qed
.
(
**
Updates
*
)
Global
Instance
option_
local_fmap
_update
L
Lv
:
Global
Instance
option_
fmap_local
_update
L
Lv
:
LocalUpdate
Lv
L
→
LocalUpdate
(
λ
mx
,
if
mx
is
Some
x
then
Lv
x
else
False
)
(
fmap
L
).
Proof
.
split
;
first
apply
_.
intros
n
[
x
|
]
[
z
|
];
constructor
;
by
eauto
using
(
local_updateN
L
).
Qed
.
Global
Instance
option_
local_const
_update
Lv
y
:
Global
Instance
option_
const_local
_update
Lv
y
:
LocalUpdate
Lv
(
λ
_
,
y
)
→
LocalUpdate
(
λ
mx
,
if
mx
is
Some
x
then
Lv
x
else
False
)
(
λ
_
,
Some
y
).
Proof
.
...
...
algebra/gmap.v
View file @
a20c86e3
...
...
@@ -283,7 +283,7 @@ Proof. apply insert_update. Qed.
Section
freshness
.
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
.
Proof
.
intros
?
HQ
.
apply
cmra_total_updateP
.
...
...
@@ -297,15 +297,15 @@ Proof.
last
by
apply
not_elem_of_dom
;
rewrite
dom_op
not_elem_of_union
.
by
apply
insert_validN
;
[
apply
cmra_valid_validN
|
].
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
.
Proof
.
move
=>??
.
eapply
updateP_alloc
_strong
with
(
I
:=
∅
);
by
eauto
.
Qed
.
Lemma
updateP_alloc
_strong
'
m
x
(
I
:
gset
K
)
:
Proof
.
move
=>??
.
eapply
alloc_updateP
_strong
with
(
I
:=
∅
);
by
eauto
.
Qed
.
Lemma
alloc_updateP
_strong
'
m
x
(
I
:
gset
K
)
:
✓
x
→
m
~~>:
λ
m
'
,
∃
i
,
i
∉
I
∧
m
'
=
<
[
i
:=
x
]
>
m
∧
m
!!
i
=
None
.
Proof
.
eauto
using
updateP_alloc
_strong
.
Qed
.
Lemma
updateP_alloc
'
m
x
:
Proof
.
eauto
using
alloc_updateP
_strong
.
Qed
.
Lemma
alloc_updateP
'
m
x
:
✓
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
:
✓
u
→
LeftId
(
≡
)
u
(
⋅
)
→
...
...
@@ -335,7 +335,7 @@ End freshness.
(
*
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
).
Proof
.
split
;
first
apply
_.
...
...
@@ -348,7 +348,7 @@ Proof.
Qed
.
(
*
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
).
Proof
.
split
;
first
apply
_.
...
...
algebra/list.v
View file @
a20c86e3
...
...
@@ -316,7 +316,7 @@ Section properties.
Proof
.
by
rewrite
!
persistent_total
list_core_singletonM
=>
->
.
Qed
.
(
*
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
.
Proof
.
intros
Hx
%
option_updateP
'
HP
.
...
...
@@ -332,13 +332,13 @@ Section properties.
rewrite
!
list_lookup_op
!
lookup_app_r
!
app_length
//=; lia.
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
.
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
.
(
*
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
).
Proof
.
split
;
[
apply
_
|
];
intros
n
l1
l2
(
x
&
Hi1
&?
)
Hm
;
apply
list_dist_lookup
=>
j
.
...
...
program_logic/ghost_ownership.v
View file @
a20c86e3
...
...
@@ -49,7 +49,7 @@ Proof.
rewrite
-
(
pvs_mono
_
_
(
∃
m
,
■
(
∃
γ
,
γ
∉
G
∧
m
=
to_globalF
γ
a
)
∧
ownG
m
)
%
I
).
-
rewrite
ownG_empty
.
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
.
-
apply
exist_elim
=>
m
;
apply
const_elim_l
=>-
[
γ
[
Hfresh
->
]].
by
rewrite
-
(
exist_intro
γ
)
const_equiv
// left_id.
...
...
Write
Preview
Markdown
is supported
0%
Try again
or
attach a new file
.
Attach a 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