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
Jonas Kastberg
iris
Commits
c311eeca
Commit
c311eeca
authored
Sep 07, 2017
by
Robbert Krebbers
Browse files
Rename `Persistent` → `CoreId` (camera elements whose core is itself).
parent
a3028b92
Changes
16
Hide whitespace changes
Inline
Side-by-side
theories/algebra/agree.v
View file @
c311eeca
...
...
@@ -146,7 +146,7 @@ Canonical Structure agreeR : cmraT := CmraT (agree A) agree_cmra_mixin.
Global
Instance
agree_cmra_total
:
CmraTotal
agreeR
.
Proof
.
rewrite
/
CmraTotal
;
eauto
.
Qed
.
Global
Instance
agree_
persistent
(
x
:
agree
A
)
:
Persistent
x
.
Global
Instance
agree_
core_id
(
x
:
agree
A
)
:
CoreId
x
.
Proof
.
by
constructor
.
Qed
.
Global
Instance
agree_cmra_discrete
:
OfeDiscrete
A
→
CmraDiscrete
agreeR
.
...
...
theories/algebra/auth.v
View file @
c311eeca
...
...
@@ -183,12 +183,12 @@ Proof.
split
;
simpl
.
-
rewrite
auth_valid_eq
/=.
apply
ucmra_unit_valid
.
-
by
intros
x
;
constructor
;
rewrite
/=
left_id
.
-
do
2
constructor
;
simpl
;
apply
(
persistent
_core
_
).
-
do
2
constructor
;
simpl
;
apply
(
core_id
_core
_
).
Qed
.
Canonical
Structure
authUR
:
=
UcmraT
(
auth
A
)
auth_ucmra_mixin
.
Global
Instance
auth_frag_
persistent
a
:
Persistent
a
→
Persistent
(
◯
a
).
Proof
.
do
2
constructor
;
simpl
;
auto
.
by
apply
persistent
_core
.
Qed
.
Global
Instance
auth_frag_
core_id
a
:
CoreId
a
→
CoreId
(
◯
a
).
Proof
.
do
2
constructor
;
simpl
;
auto
.
by
apply
core_id
_core
.
Qed
.
(** Internalized properties *)
Lemma
auth_equivI
{
M
}
(
x
y
:
auth
A
)
:
...
...
theories/algebra/cmra.v
View file @
c311eeca
...
...
@@ -142,11 +142,11 @@ Definition opM {A : cmraT} (x : A) (my : option A) :=
match
my
with
Some
y
=>
x
⋅
y
|
None
=>
x
end
.
Infix
"⋅?"
:
=
opM
(
at
level
50
,
left
associativity
)
:
C_scope
.
(** *
Persistent
elements *)
Class
Persistent
{
A
:
cmraT
}
(
x
:
A
)
:
=
persistent
:
pcore
x
≡
Some
x
.
Arguments
persistent
{
_
}
_
{
_
}.
Hint
Mode
Persistent
+
!
:
typeclass_instances
.
Instance
:
Params
(@
Persistent
)
1
.
(** *
CoreId
elements *)
Class
CoreId
{
A
:
cmraT
}
(
x
:
A
)
:
=
core_id
:
pcore
x
≡
Some
x
.
Arguments
core_id
{
_
}
_
{
_
}.
Hint
Mode
CoreId
+
!
:
typeclass_instances
.
Instance
:
Params
(@
CoreId
)
1
.
(** * Exclusive elements (i.e., elements that cannot have a frame). *)
Class
Exclusive
{
A
:
cmraT
}
(
x
:
A
)
:
=
exclusive0_l
y
:
✓
{
0
}
(
x
⋅
y
)
→
False
.
...
...
@@ -316,7 +316,7 @@ Proof. destruct 2; by ofe_subst. Qed.
Global
Instance
cmra_opM_proper
:
Proper
((
≡
)
==>
(
≡
)
==>
(
≡
))
(@
opM
A
).
Proof
.
destruct
2
;
by
setoid_subst
.
Qed
.
Global
Instance
Persistent
_proper
:
Proper
((
≡
)
==>
iff
)
(@
Persistent
A
).
Global
Instance
CoreId
_proper
:
Proper
((
≡
)
==>
iff
)
(@
CoreId
A
).
Proof
.
solve_proper
.
Qed
.
Global
Instance
Exclusive_proper
:
Proper
((
≡
)
==>
iff
)
(@
Exclusive
A
).
Proof
.
intros
x
y
Hxy
.
rewrite
/
Exclusive
.
by
setoid_rewrite
Hxy
.
Qed
.
...
...
@@ -361,8 +361,8 @@ Proof.
intros
Hv
Hx
%
cmra_pcore_l
.
move
:
Hv
;
rewrite
-
Hx
.
apply
cmra_valid_op_l
.
Qed
.
(** **
Persistent
elements *)
Lemma
persistent_dup
x
`
{!
Persistent
x
}
:
x
≡
x
⋅
x
.
(** **
CoreId
elements *)
Lemma
core_id_dup
x
`
{!
CoreId
x
}
:
x
≡
x
⋅
x
.
Proof
.
by
apply
cmra_pcore_dup'
with
x
.
Qed
.
(** ** Exclusive elements *)
...
...
@@ -498,19 +498,19 @@ Section total_core.
Lemma
cmra_core_valid
x
:
✓
x
→
✓
core
x
.
Proof
.
rewrite
-{
1
}(
cmra_core_l
x
)
;
apply
cmra_valid_op_l
.
Qed
.
Lemma
persistent
_total
x
:
Persistent
x
↔
core
x
≡
x
.
Lemma
core_id
_total
x
:
CoreId
x
↔
core
x
≡
x
.
Proof
.
split
;
[
intros
;
by
rewrite
/
core
/=
(
persistent
x
)|].
rewrite
/
Persistent
/
core
/=.
split
;
[
intros
;
by
rewrite
/
core
/=
(
core_id
x
)|].
rewrite
/
CoreId
/
core
/=.
destruct
(
cmra_total
x
)
as
[?
->].
by
constructor
.
Qed
.
Lemma
persistent
_core
x
`
{!
Persistent
x
}
:
core
x
≡
x
.
Proof
.
by
apply
persistent
_total
.
Qed
.
Lemma
core_id
_core
x
`
{!
CoreId
x
}
:
core
x
≡
x
.
Proof
.
by
apply
core_id
_total
.
Qed
.
Global
Instance
cmra_core_
persistent
x
:
Persistent
(
core
x
).
Global
Instance
cmra_core_
core_id
x
:
CoreId
(
core
x
).
Proof
.
destruct
(
cmra_total
x
)
as
[
cx
Hcx
].
rewrite
/
Persistent
/
core
/=
Hcx
/=.
eauto
using
cmra_pcore_idemp
.
rewrite
/
CoreId
/
core
/=
Hcx
/=.
eauto
using
cmra_pcore_idemp
.
Qed
.
Lemma
cmra_included_core
x
:
core
x
≼
x
.
...
...
@@ -624,13 +624,13 @@ Section ucmra.
Proof
.
by
exists
x
;
rewrite
left_id
.
Qed
.
Global
Instance
ucmra_unit_right_id
:
RightId
(
≡
)
ε
(@
op
A
_
).
Proof
.
by
intros
x
;
rewrite
(
comm
op
)
left_id
.
Qed
.
Global
Instance
ucmra_unit_
persistent
:
Persistent
(
ε
:
A
).
Global
Instance
ucmra_unit_
core_id
:
CoreId
(
ε
:
A
).
Proof
.
apply
ucmra_pcore_unit
.
Qed
.
Global
Instance
cmra_unit_cmra_total
:
CmraTotal
A
.
Proof
.
intros
x
.
destruct
(
cmra_pcore_mono'
ε
x
ε
)
as
(
cx
&->&?)
;
eauto
using
ucmra_unit_least
,
(
persistent
(
ε
:
A
)).
eauto
using
ucmra_unit_least
,
(
core_id
(
ε
:
A
)).
Qed
.
Global
Instance
empty_cancelable
:
Cancelable
(
ε
:
A
).
Proof
.
intros
???.
by
rewrite
!
left_id
.
Qed
.
...
...
@@ -666,9 +666,9 @@ Section cmra_leibniz.
Lemma
cmra_pcore_dup_L
x
cx
:
pcore
x
=
Some
cx
→
cx
=
cx
⋅
cx
.
Proof
.
unfold_leibniz
.
apply
cmra_pcore_dup'
.
Qed
.
(** **
Persistent
elements *)
Lemma
persistent
_dup_L
x
`
{!
Persistent
x
}
:
x
=
x
⋅
x
.
Proof
.
unfold_leibniz
.
by
apply
persistent
_dup
.
Qed
.
(** **
CoreId
elements *)
Lemma
core_id
_dup_L
x
`
{!
CoreId
x
}
:
x
=
x
⋅
x
.
Proof
.
unfold_leibniz
.
by
apply
core_id
_dup
.
Qed
.
(** ** Total core *)
Section
total_core
.
...
...
@@ -682,10 +682,10 @@ Section cmra_leibniz.
Proof
.
unfold_leibniz
.
apply
cmra_core_idemp
.
Qed
.
Lemma
cmra_core_dup_L
x
:
core
x
=
core
x
⋅
core
x
.
Proof
.
unfold_leibniz
.
apply
cmra_core_dup
.
Qed
.
Lemma
persistent
_total_L
x
:
Persistent
x
↔
core
x
=
x
.
Proof
.
unfold_leibniz
.
apply
persistent
_total
.
Qed
.
Lemma
persistent
_core_L
x
`
{!
Persistent
x
}
:
core
x
=
x
.
Proof
.
by
apply
persistent
_total_L
.
Qed
.
Lemma
core_id
_total_L
x
:
CoreId
x
↔
core
x
=
x
.
Proof
.
unfold_leibniz
.
apply
core_id
_total
.
Qed
.
Lemma
core_id
_core_L
x
`
{!
CoreId
x
}
:
core
x
=
x
.
Proof
.
by
apply
core_id
_total_L
.
Qed
.
End
total_core
.
End
cmra_leibniz
.
...
...
@@ -847,7 +847,7 @@ Section cmra_transport.
Proof
.
by
destruct
H
.
Qed
.
Global
Instance
cmra_transport_discrete
x
:
Discrete
x
→
Discrete
(
T
x
).
Proof
.
by
destruct
H
.
Qed
.
Global
Instance
cmra_transport_
persistent
x
:
Persistent
x
→
Persistent
(
T
x
).
Global
Instance
cmra_transport_
core_id
x
:
CoreId
x
→
CoreId
(
T
x
).
Proof
.
by
destruct
H
.
Qed
.
End
cmra_transport
.
...
...
@@ -938,7 +938,7 @@ Section unit.
Global
Instance
unit_cmra_discrete
:
CmraDiscrete
unitR
.
Proof
.
done
.
Qed
.
Global
Instance
unit_
persistent
(
x
:
())
:
Persistent
x
.
Global
Instance
unit_
core_id
(
x
:
())
:
CoreId
x
.
Proof
.
by
constructor
.
Qed
.
Global
Instance
unit_cancelable
(
x
:
())
:
Cancelable
x
.
Proof
.
by
constructor
.
Qed
.
...
...
@@ -1008,7 +1008,7 @@ Section mnat.
Proof
.
split
;
apply
_
||
done
.
Qed
.
Canonical
Structure
mnatUR
:
ucmraT
:
=
UcmraT
mnat
mnat_ucmra_mixin
.
Global
Instance
mnat_
persistent
(
x
:
mnat
)
:
Persistent
x
.
Global
Instance
mnat_
core_id
(
x
:
mnat
)
:
CoreId
x
.
Proof
.
by
constructor
.
Qed
.
End
mnat
.
...
...
@@ -1122,9 +1122,9 @@ Section prod.
CmraDiscrete
A
→
CmraDiscrete
B
→
CmraDiscrete
prodR
.
Proof
.
split
.
apply
_
.
by
intros
?
[]
;
split
;
apply
cmra_discrete_valid
.
Qed
.
Global
Instance
pair_
persistent
x
y
:
Persistent
x
→
Persistent
y
→
Persistent
(
x
,
y
).
Proof
.
by
rewrite
/
Persistent
prod_pcore_Some'
.
Qed
.
Global
Instance
pair_
core_id
x
y
:
CoreId
x
→
CoreId
y
→
CoreId
(
x
,
y
).
Proof
.
by
rewrite
/
CoreId
prod_pcore_Some'
.
Qed
.
Global
Instance
pair_exclusive_l
x
y
:
Exclusive
x
→
Exclusive
(
x
,
y
).
Proof
.
by
intros
?[][?%
exclusive0_l
].
Qed
.
...
...
@@ -1152,7 +1152,7 @@ Section prod_unit.
split
.
-
split
;
apply
ucmra_unit_valid
.
-
by
split
;
rewrite
/=
left_id
.
-
rewrite
prod_pcore_Some'
;
split
;
apply
(
persistent
_
).
-
rewrite
prod_pcore_Some'
;
split
;
apply
(
core_id
_
).
Qed
.
Canonical
Structure
prodUR
:
=
UcmraT
(
prod
A
B
)
prod_ucmra_mixin
.
...
...
@@ -1328,10 +1328,10 @@ Section option.
Lemma
op_is_Some
mx
my
:
is_Some
(
mx
⋅
my
)
↔
is_Some
mx
∨
is_Some
my
.
Proof
.
rewrite
-!
not_eq_None_Some
op_None
.
destruct
mx
,
my
;
naive_solver
.
Qed
.
Global
Instance
Some_
persistent
(
x
:
A
)
:
Persistent
x
→
Persistent
(
Some
x
).
Global
Instance
Some_
core_id
(
x
:
A
)
:
CoreId
x
→
CoreId
(
Some
x
).
Proof
.
by
constructor
.
Qed
.
Global
Instance
option_
persistent
(
mx
:
option
A
)
:
(
∀
x
:
A
,
Persistent
x
)
→
Persistent
mx
.
Global
Instance
option_
core_id
(
mx
:
option
A
)
:
(
∀
x
:
A
,
CoreId
x
)
→
CoreId
mx
.
Proof
.
intros
.
destruct
mx
;
apply
_
.
Qed
.
Lemma
exclusiveN_Some_l
n
x
`
{!
Exclusive
x
}
my
:
...
...
theories/algebra/csum.v
View file @
c311eeca
...
...
@@ -256,10 +256,10 @@ Proof.
by
move
=>[
a
|
b
|]
HH
/=
;
try
apply
cmra_discrete_valid
.
Qed
.
Global
Instance
Cinl_
persistent
a
:
Persistent
a
→
Persistent
(
Cinl
a
).
Proof
.
rewrite
/
Persistent
/=.
inversion_clear
1
;
by
repeat
constructor
.
Qed
.
Global
Instance
Cinr_
persistent
b
:
Persistent
b
→
Persistent
(
Cinr
b
).
Proof
.
rewrite
/
Persistent
/=.
inversion_clear
1
;
by
repeat
constructor
.
Qed
.
Global
Instance
Cinl_
core_id
a
:
CoreId
a
→
CoreId
(
Cinl
a
).
Proof
.
rewrite
/
CoreId
/=.
inversion_clear
1
;
by
repeat
constructor
.
Qed
.
Global
Instance
Cinr_
core_id
b
:
CoreId
b
→
CoreId
(
Cinr
b
).
Proof
.
rewrite
/
CoreId
/=.
inversion_clear
1
;
by
repeat
constructor
.
Qed
.
Global
Instance
Cinl_exclusive
a
:
Exclusive
a
→
Exclusive
(
Cinl
a
).
Proof
.
by
move
=>
H
[]?
=>[/
H
||].
Qed
.
...
...
theories/algebra/deprecated.v
View file @
c311eeca
...
...
@@ -56,7 +56,7 @@ Global Instance dec_agree_cmra_total : CmraTotal dec_agreeR.
Proof
.
intros
x
.
by
exists
x
.
Qed
.
(* Some properties of this CMRA *)
Global
Instance
dec_agree_
persistent
(
x
:
dec_agreeR
)
:
Persistent
x
.
Global
Instance
dec_agree_
core_id
(
x
:
dec_agreeR
)
:
CoreId
x
.
Proof
.
by
constructor
.
Qed
.
Lemma
dec_agree_ne
a
b
:
a
≠
b
→
DecAgree
a
⋅
DecAgree
b
=
DecAgreeBot
.
...
...
theories/algebra/frac_auth.v
View file @
c311eeca
...
...
@@ -94,10 +94,10 @@ Section frac_auth.
Proof
.
by
rewrite
/
IsOp'
/
IsOp
=>
/
leibniz_equiv_iff
->
->.
Qed
.
Global
Instance
is_op_frac_auth_persistent
(
q
q1
q2
:
frac
)
(
a
:
A
)
:
Persistent
a
→
IsOp
q
q1
q2
→
IsOp'
(
◯
!{
q
}
a
)
(
◯
!{
q1
}
a
)
(
◯
!{
q2
}
a
).
CoreId
a
→
IsOp
q
q1
q2
→
IsOp'
(
◯
!{
q
}
a
)
(
◯
!{
q1
}
a
)
(
◯
!{
q2
}
a
).
Proof
.
rewrite
/
IsOp'
/
IsOp
=>
?
/
leibniz_equiv_iff
->.
by
rewrite
-
frag_auth_op
-
persistent
_dup
.
by
rewrite
-
frag_auth_op
-
core_id
_dup
.
Qed
.
Lemma
frac_auth_update
q
a
b
a'
b'
:
...
...
theories/algebra/gmap.v
View file @
c311eeca
...
...
@@ -250,14 +250,14 @@ Lemma op_singleton (i : K) (x y : A) :
{[
i
:
=
x
]}
⋅
{[
i
:
=
y
]}
=
({[
i
:
=
x
⋅
y
]}
:
gmap
K
A
).
Proof
.
by
apply
(
merge_singleton
_
_
_
x
y
).
Qed
.
Global
Instance
gmap_
persistent
m
:
(
∀
x
:
A
,
Persistent
x
)
→
Persistent
m
.
Global
Instance
gmap_
core_id
m
:
(
∀
x
:
A
,
CoreId
x
)
→
CoreId
m
.
Proof
.
intros
;
apply
persistent
_total
=>
i
.
rewrite
lookup_core
.
apply
(
persistent
_core
_
).
intros
;
apply
core_id
_total
=>
i
.
rewrite
lookup_core
.
apply
(
core_id
_core
_
).
Qed
.
Global
Instance
gmap_singleton_
persistent
i
(
x
:
A
)
:
Persistent
x
→
Persistent
{[
i
:
=
x
]}.
Proof
.
intros
.
by
apply
persistent
_total
,
core_singleton'
.
Qed
.
Global
Instance
gmap_singleton_
core_id
i
(
x
:
A
)
:
CoreId
x
→
CoreId
{[
i
:
=
x
]}.
Proof
.
intros
.
by
apply
core_id
_total
,
core_singleton'
.
Qed
.
Lemma
singleton_includedN
n
m
i
x
:
{[
i
:
=
x
]}
≼
{
n
}
m
↔
∃
y
,
m
!!
i
≡
{
n
}
≡
Some
y
∧
Some
x
≼
{
n
}
Some
y
.
...
...
theories/algebra/gset.v
View file @
c311eeca
...
...
@@ -58,8 +58,8 @@ Section gset.
split
.
done
.
rewrite
gset_op_union
.
set_solver
.
Qed
.
Global
Instance
gset_
persistent
X
:
Persistent
X
.
Proof
.
by
apply
persistent
_total
;
rewrite
gset_core_self
.
Qed
.
Global
Instance
gset_
core_id
X
:
CoreId
X
.
Proof
.
by
apply
core_id
_total
;
rewrite
gset_core_self
.
Qed
.
End
gset
.
Arguments
gsetC
_
{
_
_
}.
...
...
theories/algebra/iprod.v
View file @
c311eeca
...
...
@@ -136,7 +136,7 @@ Section iprod_cmra.
split
.
-
intros
x
;
apply
ucmra_unit_valid
.
-
by
intros
f
x
;
rewrite
iprod_lookup_op
left_id
.
-
constructor
=>
x
.
apply
persistent
_core
,
_
.
-
constructor
=>
x
.
apply
core_id
_core
,
_
.
Qed
.
Canonical
Structure
iprodUR
:
=
UcmraT
(
iprod
B
)
iprod_ucmra_mixin
.
...
...
@@ -215,12 +215,12 @@ Section iprod_singleton.
Proof
.
move
=>
x'
;
destruct
(
decide
(
x
=
x'
))
as
[->|]
;
by
rewrite
iprod_lookup_core
?iprod_lookup_singleton
?iprod_lookup_singleton_ne
//
(
persistent
_core
∅
).
?iprod_lookup_singleton_ne
//
(
core_id
_core
∅
).
Qed
.
Global
Instance
iprod_singleton_
persistent
x
(
y
:
B
x
)
:
Persistent
y
→
Persistent
(
iprod_singleton
x
y
).
Proof
.
by
rewrite
!
persistent
_total
iprod_core_singleton
=>
->.
Qed
.
Global
Instance
iprod_singleton_
core_id
x
(
y
:
B
x
)
:
CoreId
y
→
CoreId
(
iprod_singleton
x
y
).
Proof
.
by
rewrite
!
core_id
_total
iprod_core_singleton
=>
->.
Qed
.
Lemma
iprod_op_singleton
(
x
:
A
)
(
y1
y2
:
B
x
)
:
iprod_singleton
x
y1
⋅
iprod_singleton
x
y2
≡
iprod_singleton
x
(
y1
⋅
y2
).
...
...
theories/algebra/list.v
View file @
c311eeca
...
...
@@ -237,10 +237,10 @@ Section cmra.
by
apply
cmra_discrete_valid
.
Qed
.
Global
Instance
list_
persistent
l
:
(
∀
x
:
A
,
Persistent
x
)
→
Persistent
l
.
Global
Instance
list_
core_id
l
:
(
∀
x
:
A
,
CoreId
x
)
→
CoreId
l
.
Proof
.
intros
?
;
constructor
;
apply
list_equiv_lookup
=>
i
.
by
rewrite
list_lookup_core
(
persistent
_core
(
l
!!
i
)).
by
rewrite
list_lookup_core
(
core_id
_core
(
l
!!
i
)).
Qed
.
(** Internalized properties *)
...
...
@@ -329,7 +329,7 @@ Section properties.
Lemma
list_core_singletonM
i
(
x
:
A
)
:
core
{[
i
:
=
x
]}
≡
{[
i
:
=
core
x
]}.
Proof
.
rewrite
/
singletonM
/
list_singletonM
.
by
rewrite
{
1
}/
core
/=
fmap_app
fmap_replicate
(
persistent
_core
∅
).
by
rewrite
{
1
}/
core
/=
fmap_app
fmap_replicate
(
core_id
_core
∅
).
Qed
.
Lemma
list_op_singletonM
i
(
x
y
:
A
)
:
{[
i
:
=
x
]}
⋅
{[
i
:
=
y
]}
≡
{[
i
:
=
x
⋅
y
]}.
...
...
@@ -342,9 +342,9 @@ Section properties.
Proof
.
rewrite
/
singletonM
/
list_singletonM
/=.
induction
i
;
f_equal
/=
;
auto
.
Qed
.
Global
Instance
list_singleton_
persistent
i
(
x
:
A
)
:
Persistent
x
→
Persistent
{[
i
:
=
x
]}.
Proof
.
by
rewrite
!
persistent
_total
list_core_singletonM
=>
->.
Qed
.
Global
Instance
list_singleton_
core_id
i
(
x
:
A
)
:
CoreId
x
→
CoreId
{[
i
:
=
x
]}.
Proof
.
by
rewrite
!
core_id
_total
list_core_singletonM
=>
->.
Qed
.
(* Update *)
Lemma
list_singleton_updateP
(
P
:
A
→
Prop
)
(
Q
:
list
A
→
Prop
)
x
:
...
...
theories/algebra/sts.v
View file @
c311eeca
...
...
@@ -391,7 +391,7 @@ Proof.
eapply
closed_steps
,
Hsup
;
first
done
.
set_solver
+
Hs'
.
Qed
.
Global
Instance
sts_frag_
peristent
S
:
Persistent
(
sts_frag
S
∅
).
Global
Instance
sts_frag_
core_id
S
:
CoreId
(
sts_frag
S
∅
).
Proof
.
constructor
;
split
=>
//=
[[??]].
by
rewrite
/
dra
.
dra_core
/=
sts
.
up_closed
.
Qed
.
...
...
theories/base_logic/derived.v
View file @
c311eeca
...
...
@@ -753,10 +753,10 @@ Lemma except_0_frame_r P Q : ◇ P ∗ Q ⊢ ◇ (P ∗ Q).
Proof
.
by
rewrite
{
1
}(
except_0_intro
Q
)
except_0_sep
.
Qed
.
(* Own and valid derived *)
Lemma
always_ownM
(
a
:
M
)
:
Persistent
a
→
□
uPred_ownM
a
⊣
⊢
uPred_ownM
a
.
Lemma
always_ownM
(
a
:
M
)
:
CoreId
a
→
□
uPred_ownM
a
⊣
⊢
uPred_ownM
a
.
Proof
.
intros
;
apply
(
anti_symm
_
)
;
first
by
apply
:
always_elim
.
by
rewrite
{
1
}
always_ownM_core
persistent
_core
.
by
rewrite
{
1
}
always_ownM_core
core_id
_core
.
Qed
.
Lemma
ownM_invalid
(
a
:
M
)
:
¬
✓
{
0
}
a
→
uPred_ownM
a
⊢
False
.
Proof
.
by
intros
;
rewrite
ownM_valid
cmra_valid_elim
.
Qed
.
...
...
@@ -931,7 +931,7 @@ Global Instance later_persistent P : PersistentP P → PersistentP (▷ P).
Proof
.
by
intros
;
rewrite
/
PersistentP
always_later
;
apply
later_mono
.
Qed
.
Global
Instance
laterN_persistent
n
P
:
PersistentP
P
→
PersistentP
(
▷
^
n
P
).
Proof
.
induction
n
;
apply
_
.
Qed
.
Global
Instance
ownM_persistent
:
Persistent
a
→
PersistentP
(@
uPred_ownM
M
a
).
Global
Instance
ownM_persistent
:
CoreId
a
→
PersistentP
(@
uPred_ownM
M
a
).
Proof
.
intros
.
by
rewrite
/
PersistentP
always_ownM
.
Qed
.
Global
Instance
from_option_persistent
{
A
}
P
(
Ψ
:
A
→
uPred
M
)
(
mx
:
option
A
)
:
(
∀
x
,
PersistentP
(
Ψ
x
))
→
PersistentP
P
→
PersistentP
(
from_option
Ψ
P
mx
).
...
...
theories/base_logic/lib/auth.v
View file @
c311eeca
...
...
@@ -32,7 +32,7 @@ Section definitions.
Proof
.
solve_proper
.
Qed
.
Global
Instance
auth_own_timeless
a
:
TimelessP
(
auth_own
a
).
Proof
.
apply
_
.
Qed
.
Global
Instance
auth_own_
persistent
a
:
Persistent
a
→
PersistentP
(
auth_own
a
).
Global
Instance
auth_own_
core_id
a
:
CoreId
a
→
PersistentP
(
auth_own
a
).
Proof
.
apply
_
.
Qed
.
Global
Instance
auth_inv_ne
n
:
...
...
@@ -78,7 +78,7 @@ Section auth.
FromAnd
false
(
auth_own
γ
a
)
(
auth_own
γ
b1
)
(
auth_own
γ
b2
)
|
90
.
Proof
.
rewrite
/
IsOp
/
FromAnd
=>
->.
by
rewrite
auth_own_op
.
Qed
.
Global
Instance
from_and_auth_own_persistent
γ
a
b1
b2
:
IsOp
a
b1
b2
→
Or
(
Persistent
b1
)
(
Persistent
b2
)
→
IsOp
a
b1
b2
→
Or
(
CoreId
b1
)
(
CoreId
b2
)
→
FromAnd
true
(
auth_own
γ
a
)
(
auth_own
γ
b1
)
(
auth_own
γ
b2
)
|
91
.
Proof
.
intros
?
Hper
;
apply
mk_from_and_persistent
;
[
destruct
Hper
;
apply
_
|].
...
...
theories/base_logic/lib/own.v
View file @
c311eeca
...
...
@@ -108,7 +108,7 @@ Proof. by rewrite comm -own_valid_r. Qed.
Global
Instance
own_timeless
γ
a
:
Discrete
a
→
TimelessP
(
own
γ
a
).
Proof
.
rewrite
!
own_eq
/
own_def
;
apply
_
.
Qed
.
Global
Instance
own_core_persistent
γ
a
:
Persistent
a
→
PersistentP
(
own
γ
a
).
Global
Instance
own_core_persistent
γ
a
:
CoreId
a
→
PersistentP
(
own
γ
a
).
Proof
.
rewrite
!
own_eq
/
own_def
;
apply
_
.
Qed
.
(** ** Allocation *)
...
...
@@ -193,7 +193,7 @@ Section proofmode_classes.
IsOp
a
b1
b2
→
FromAnd
false
(
own
γ
a
)
(
own
γ
b1
)
(
own
γ
b2
).
Proof
.
intros
.
by
rewrite
/
FromAnd
-
own_op
-
is_op
.
Qed
.
Global
Instance
from_and_own_persistent
γ
a
b1
b2
:
IsOp
a
b1
b2
→
Or
(
Persistent
b1
)
(
Persistent
b2
)
→
IsOp
a
b1
b2
→
Or
(
CoreId
b1
)
(
CoreId
b2
)
→
FromAnd
true
(
own
γ
a
)
(
own
γ
b1
)
(
own
γ
b2
).
Proof
.
intros
?
Hper
;
apply
mk_from_and_persistent
;
[
destruct
Hper
;
apply
_
|].
...
...
theories/proofmode/class_instances.v
View file @
c311eeca
...
...
@@ -370,7 +370,7 @@ Global Instance from_sep_ownM (a b1 b2 : M) :
FromAnd
false
(
uPred_ownM
a
)
(
uPred_ownM
b1
)
(
uPred_ownM
b2
).
Proof
.
intros
.
by
rewrite
/
FromAnd
-
ownM_op
-
is_op
.
Qed
.
Global
Instance
from_sep_ownM_persistent
(
a
b1
b2
:
M
)
:
IsOp
a
b1
b2
→
Or
(
Persistent
b1
)
(
Persistent
b2
)
→
IsOp
a
b1
b2
→
Or
(
CoreId
b1
)
(
CoreId
b2
)
→
FromAnd
true
(
uPred_ownM
a
)
(
uPred_ownM
b1
)
(
uPred_ownM
b2
).
Proof
.
intros
?
Hper
;
apply
mk_from_and_persistent
;
[
destruct
Hper
;
apply
_
|].
...
...
@@ -406,11 +406,11 @@ Global Instance is_op_pair {A B : cmraT} (a b1 b2 : A) (a' b1' b2' : B) :
IsOp
a
b1
b2
→
IsOp
a'
b1'
b2'
→
IsOp'
(
a
,
a'
)
(
b1
,
b1'
)
(
b2
,
b2'
).
Proof
.
by
constructor
.
Qed
.
Global
Instance
is_op_pair_persistent_l
{
A
B
:
cmraT
}
(
a
:
A
)
(
a'
b1'
b2'
:
B
)
:
Persistent
a
→
IsOp
a'
b1'
b2'
→
IsOp'
(
a
,
a'
)
(
a
,
b1'
)
(
a
,
b2'
).
Proof
.
constructor
=>
//=.
by
rewrite
-
persistent
_dup
.
Qed
.
CoreId
a
→
IsOp
a'
b1'
b2'
→
IsOp'
(
a
,
a'
)
(
a
,
b1'
)
(
a
,
b2'
).
Proof
.
constructor
=>
//=.
by
rewrite
-
core_id
_dup
.
Qed
.
Global
Instance
is_op_pair_persistent_r
{
A
B
:
cmraT
}
(
a
b1
b2
:
A
)
(
a'
:
B
)
:
Persistent
a'
→
IsOp
a
b1
b2
→
IsOp'
(
a
,
a'
)
(
b1
,
a'
)
(
b2
,
a'
).
Proof
.
constructor
=>
//=.
by
rewrite
-
persistent
_dup
.
Qed
.
CoreId
a'
→
IsOp
a
b1
b2
→
IsOp'
(
a
,
a'
)
(
b1
,
a'
)
(
b2
,
a'
).
Proof
.
constructor
=>
//=.
by
rewrite
-
core_id
_dup
.
Qed
.
Global
Instance
is_op_Some
{
A
:
cmraT
}
(
a
:
A
)
b1
b2
:
IsOp
a
b1
b2
→
IsOp'
(
Some
a
)
(
Some
b1
)
(
Some
b2
).
...
...
theories/tests/ipm_paper.v
View file @
c311eeca
...
...
@@ -162,7 +162,7 @@ Section M.
Qed
.
Canonical
Structure
M_UR
:
ucmraT
:
=
UcmraT
M
M_ucmra_mixin
.
Global
Instance
frag_
persistent
n
:
Persistent
(
Frag
n
).
Global
Instance
frag_
core_id
n
:
CoreId
(
Frag
n
).
Proof
.
by
constructor
.
Qed
.
Lemma
auth_frag_valid
j
n
:
✓
(
Auth
n
⋅
Frag
j
)
→
(
j
≤
n
)%
nat
.
Proof
.
simpl
.
case_decide
.
done
.
by
intros
[].
Qed
.
...
...
Write
Preview
Supports
Markdown
0%
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!
Cancel
Please
register
or
sign in
to comment