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
George Pirlea
Iris
Commits
da93f357
Commit
da93f357
authored
Mar 05, 2019
by
Ralf Jung
Browse files
use ! when possible to avoid overzealous generalization
parent
38abc449
Changes
52
Hide whitespace changes
Inline
Side-by-side
tests/algebra.v
View file @
da93f357
From
iris
.
base_logic
.
lib
Require
Import
invariants
.
Section
tests
.
Context
`
{
invG
Σ
}.
Context
`
{
!
invG
Σ
}.
Program
Definition
test
:
(
iProp
Σ
-
n
>
iProp
Σ
)
-
n
>
(
iProp
Σ
-
n
>
iProp
Σ
)
:
=
λ
ne
P
v
,
(
▷
(
P
v
))%
I
.
...
...
tests/heap_lang.ref
View file @
da93f357
1 subgoal
Σ : gFunctors
H
: heapG Σ
heapG0
: heapG Σ
E : coPset
============================
--------------------------------------∗
...
...
@@ -10,7 +10,7 @@
1 subgoal
Σ : gFunctors
H
: heapG Σ
heapG0
: heapG Σ
E : coPset
l : loc
============================
...
...
@@ -21,7 +21,7 @@
1 subgoal
Σ : gFunctors
H
: heapG Σ
heapG0
: heapG Σ
E : coPset
l : loc
============================
...
...
@@ -35,7 +35,7 @@
1 subgoal
Σ : gFunctors
H
: heapG Σ
heapG0
: heapG Σ
l : loc
============================
_ : ▷ l ↦ #0
...
...
@@ -45,7 +45,7 @@
1 subgoal
Σ : gFunctors
H
: heapG Σ
heapG0
: heapG Σ
l : loc
============================
_ : l ↦ #1
...
...
@@ -55,7 +55,7 @@
1 subgoal
Σ : gFunctors
H
: heapG Σ
heapG0
: heapG Σ
l : loc
============================
"Hl1" : l ↦{1 / 2} #0
...
...
@@ -66,7 +66,7 @@
1 subgoal
Σ : gFunctors
H
: heapG Σ
heapG0
: heapG Σ
l : loc
============================
--------------------------------------∗
...
...
@@ -81,7 +81,7 @@ Tactic failure: wp_pure: cannot find ?y in (Var "x") or
1 subgoal
Σ : gFunctors
H
: heapG Σ
heapG0
: heapG Σ
============================
--------------------------------------∗
WP "x" {{ _, True }}
...
...
@@ -89,7 +89,7 @@ Tactic failure: wp_pure: cannot find ?y in (Var "x") or
1 subgoal
Σ : gFunctors
H
: heapG Σ
heapG0
: heapG Σ
fun1, fun2, fun3 : expr
============================
--------------------------------------∗
...
...
@@ -101,7 +101,7 @@ Tactic failure: wp_pure: cannot find ?y in (Var "x") or
1 subgoal
Σ : gFunctors
H
: heapG Σ
heapG0
: heapG Σ
fun1, fun2, fun3 : expr
Φ : language.val heap_lang → iPropI Σ
============================
...
...
@@ -114,7 +114,7 @@ Tactic failure: wp_pure: cannot find ?y in (Var "x") or
1 subgoal
Σ : gFunctors
H
: heapG Σ
heapG0
: heapG Σ
fun1, fun2, fun3 : expr
Φ : language.val heap_lang → iPropI Σ
E : coPset
...
...
@@ -128,7 +128,7 @@ Tactic failure: wp_pure: cannot find ?y in (Var "x") or
1 subgoal
Σ : gFunctors
H
: heapG Σ
heapG0
: heapG Σ
fun1, fun2, fun3 : expr
============================
{{{ True }}}
...
...
tests/heap_lang.v
View file @
da93f357
...
...
@@ -6,7 +6,7 @@ Set Ltac Backtrace.
Set
Default
Proof
Using
"Type"
.
Section
tests
.
Context
`
{
heapG
Σ
}.
Context
`
{
!
heapG
Σ
}.
Implicit
Types
P
Q
:
iProp
Σ
.
Implicit
Types
Φ
:
val
→
iProp
Σ
.
...
...
@@ -147,7 +147,7 @@ Section tests.
End
tests
.
Section
printing_tests
.
Context
`
{
heapG
Σ
}.
Context
`
{
!
heapG
Σ
}.
(* These terms aren't even closed, but that's not what this is about. The
length of the variable names etc. has been carefully chosen to trigger
...
...
@@ -192,7 +192,7 @@ Section printing_tests.
End
printing_tests
.
Section
error_tests
.
Context
`
{
heapG
Σ
}.
Context
`
{
!
heapG
Σ
}.
Check
"not_cas"
.
Lemma
not_cas
:
...
...
tests/heap_lang2.ref
View file @
da93f357
1 subgoal
Σ : gFunctors
H
: heapG Σ
heapG0
: heapG Σ
fun1, fun2, fun3 : expr
============================
--------------------------------------∗
...
...
tests/heap_lang2.v
View file @
da93f357
...
...
@@ -6,7 +6,7 @@ From iris.heap_lang Require Import proofmode notation.
Set
Default
Proof
Using
"Type"
.
Section
printing_tests
.
Context
`
{
heapG
Σ
}.
Context
`
{
!
heapG
Σ
}.
Lemma
wp_print_long_expr
(
fun1
fun2
fun3
:
expr
)
:
True
-
∗
WP
let
:
"val1"
:
=
fun1
#()
in
...
...
tests/ipm_paper.v
View file @
da93f357
...
...
@@ -107,7 +107,7 @@ under max can be found in [theories/heap_lang/lib/counter.v]. *)
update modalities (which we did not cover in the paper). Normally we use these
mask changing update modalities directly in our proofs, but in this file we use
the first prove the rule as a lemma, and then use that. *)
Lemma
wp_inv_open
`
{
irisG
Λ
Σ
}
N
E
P
e
Φ
:
Lemma
wp_inv_open
`
{
!
irisG
Λ
Σ
}
N
E
P
e
Φ
:
nclose
N
⊆
E
→
Atomic
WeaklyAtomic
e
→
inv
N
P
∗
(
▷
P
-
∗
WP
e
@
E
∖
↑
N
{{
v
,
▷
P
∗
Φ
v
}})
⊢
WP
e
@
E
{{
Φ
}}.
Proof
.
...
...
tests/proofmode.ref
View file @
da93f357
...
...
@@ -131,7 +131,7 @@ Tactic failure: iFrame: cannot frame Q.
1 subgoal
PROP : sbi
H
: BiAffine PROP
BiAffine0
: BiAffine PROP
P, Q : PROP
============================
_ : □ P
...
...
tests/proofmode.v
View file @
da93f357
...
...
@@ -69,7 +69,7 @@ Lemma test_iDestruct_intuitionistic_2 P Q `{!Persistent P, !Affine P}:
Q
∗
(
Q
-
∗
P
)
-
∗
P
.
Proof
.
iIntros
"[HQ HQP]"
.
iDestruct
(
"HQP"
with
"HQ"
)
as
"#HP"
.
done
.
Qed
.
Lemma
test_iDestruct_intuitionistic_affine_bi
`
{
BiAffine
PROP
}
P
Q
`
{!
Persistent
P
}
:
Lemma
test_iDestruct_intuitionistic_affine_bi
`
{
!
BiAffine
PROP
}
P
Q
`
{!
Persistent
P
}
:
Q
∗
(
Q
-
∗
P
)
-
∗
P
∗
Q
.
Proof
.
iIntros
"[HQ HQP]"
.
iDestruct
(
"HQP"
with
"HQ"
)
as
"#HP"
.
by
iFrame
.
Qed
.
...
...
@@ -180,7 +180,7 @@ Lemma test_iFrame_conjunction_2 P Q :
P
-
∗
Q
-
∗
(
P
∧
P
)
∗
(
Q
∧
Q
).
Proof
.
iIntros
"HP HQ"
.
iFrame
"HP HQ"
.
Qed
.
Lemma
test_iFrame_later
`
{
BiAffine
PROP
}
P
Q
:
P
-
∗
Q
-
∗
▷
P
∗
Q
.
Lemma
test_iFrame_later
`
{
!
BiAffine
PROP
}
P
Q
:
P
-
∗
Q
-
∗
▷
P
∗
Q
.
Proof
.
iIntros
"H1 H2"
.
by
iFrame
"H1"
.
Qed
.
Lemma
test_iAssert_modality
P
:
◇
False
-
∗
▷
P
.
...
...
@@ -555,7 +555,7 @@ Proof.
Qed
.
Check
"test_and_sep_affine_bi"
.
Lemma
test_and_sep_affine_bi
`
{
BiAffine
PROP
}
P
Q
:
□
P
∧
Q
⊢
□
P
∗
Q
.
Lemma
test_and_sep_affine_bi
`
{
!
BiAffine
PROP
}
P
Q
:
□
P
∧
Q
⊢
□
P
∗
Q
.
Proof
.
iIntros
"[??]"
.
iSplit
;
last
done
.
Show
.
done
.
Qed
.
...
...
tests/proofmode_iris.ref
View file @
da93f357
1 subgoal
Σ : gFunctors
H
: invG Σ
H
0 : cinvG Σ
H1
: na_invG Σ
invG0
: invG Σ
cinvG
0 : cinvG Σ
na_invG0
: na_invG Σ
N : namespace
P : iProp Σ
============================
...
...
@@ -15,9 +15,9 @@
1 subgoal
Σ : gFunctors
H
: invG Σ
H
0 : cinvG Σ
H1
: na_invG Σ
invG0
: invG Σ
cinvG
0 : cinvG Σ
na_invG0
: na_invG Σ
N : namespace
P : iProp Σ
============================
...
...
@@ -31,9 +31,9 @@
1 subgoal
Σ : gFunctors
H
: invG Σ
H
0 : cinvG Σ
H1
: na_invG Σ
invG0
: invG Σ
cinvG
0 : cinvG Σ
na_invG0
: na_invG Σ
γ : gname
p : Qp
N : namespace
...
...
@@ -49,9 +49,9 @@
1 subgoal
Σ : gFunctors
H
: invG Σ
H
0 : cinvG Σ
H1
: na_invG Σ
invG0
: invG Σ
cinvG
0 : cinvG Σ
na_invG0
: na_invG Σ
γ : gname
p : Qp
N : namespace
...
...
@@ -68,14 +68,14 @@
1 subgoal
Σ : gFunctors
H
: invG Σ
H
0 : cinvG Σ
H1
: na_invG Σ
invG0
: invG Σ
cinvG
0 : cinvG Σ
na_invG0
: na_invG Σ
t : na_inv_pool_name
N : namespace
E1, E2 : coPset
P : iProp Σ
H
2
: ↑N ⊆ E2
H : ↑N ⊆ E2
============================
_ : na_inv t N (<pers> P)
"HP" : ▷ <pers> P
...
...
@@ -89,14 +89,14 @@
1 subgoal
Σ : gFunctors
H
: invG Σ
H
0 : cinvG Σ
H1
: na_invG Σ
invG0
: invG Σ
cinvG
0 : cinvG Σ
na_invG0
: na_invG Σ
t : na_inv_pool_name
N : namespace
E1, E2 : coPset
P : iProp Σ
H
2
: ↑N ⊆ E2
H : ↑N ⊆ E2
============================
_ : na_inv t N (<pers> P)
"HP" : ▷ <pers> P
...
...
@@ -132,12 +132,12 @@ Tactic failure: iInv: invariant "H2" not found.
1 subgoal
Σ : gFunctors
H
: invG Σ
invG0
: invG Σ
I : biIndex
N : namespace
E : coPset
𝓟 : iProp Σ
H
0
: ↑N ⊆ E
H : ↑N ⊆ E
============================
"HP" : ⎡ ▷ 𝓟 ⎤
--------------------------------------∗
...
...
@@ -148,12 +148,12 @@ Tactic failure: iInv: invariant "H2" not found.
1 subgoal
Σ : gFunctors
H
: invG Σ
invG0
: invG Σ
I : biIndex
N : namespace
E : coPset
𝓟 : iProp Σ
H
0
: ↑N ⊆ E
H : ↑N ⊆ E
============================
"HP" : ⎡ ▷ 𝓟 ⎤
"Hclose" : ⎡ ▷ 𝓟 ={E ∖ ↑N,E}=∗ emp ⎤
...
...
tests/proofmode_iris.v
View file @
da93f357
...
...
@@ -50,7 +50,7 @@ Section base_logic_tests.
End
base_logic_tests
.
Section
iris_tests
.
Context
`
{
invG
Σ
,
cinvG
Σ
,
na_invG
Σ
}.
Context
`
{
!
invG
Σ
,
!
cinvG
Σ
,
!
na_invG
Σ
}.
Implicit
Types
P
Q
R
:
iProp
Σ
.
Lemma
test_masks
N
E
P
Q
R
:
...
...
@@ -223,7 +223,7 @@ Section iris_tests.
End
iris_tests
.
Section
monpred_tests
.
Context
`
{
invG
Σ
}.
Context
`
{
!
invG
Σ
}.
Context
{
I
:
biIndex
}.
Local
Notation
monPred
:
=
(
monPred
I
(
iPropI
Σ
)).
Local
Notation
monPredI
:
=
(
monPredI
I
(
iPropI
Σ
)).
...
...
theories/algebra/agree.v
View file @
da93f357
...
...
@@ -266,7 +266,7 @@ Lemma agree_map_to_agree {A B} (f : A → B) (x : A) :
Proof
.
by
apply
agree_eq
.
Qed
.
Section
agree_map
.
Context
{
A
B
:
ofeT
}
(
f
:
A
→
B
)
`
{
Hf
:
NonExpansive
f
}.
Context
{
A
B
:
ofeT
}
(
f
:
A
→
B
)
{
Hf
:
NonExpansive
f
}.
Instance
agree_map_ne
:
NonExpansive
(
agree_map
f
).
Proof
.
...
...
theories/algebra/auth.v
View file @
da93f357
...
...
@@ -43,7 +43,7 @@ Definition auth_ofe_mixin : OfeMixin (auth A).
Proof
.
by
apply
(
iso_ofe_mixin
(
λ
x
,
(
authoritative
x
,
auth_own
x
))).
Qed
.
Canonical
Structure
authC
:
=
OfeT
(
auth
A
)
auth_ofe_mixin
.
Global
Instance
auth_cofe
`
{
Cofe
A
}
:
Cofe
authC
.
Global
Instance
auth_cofe
`
{
!
Cofe
A
}
:
Cofe
authC
.
Proof
.
apply
(
iso_cofe
(
λ
y
:
_
*
_
,
Auth
(
y
.
1
)
(
y
.
2
))
(
λ
x
,
(
authoritative
x
,
auth_own
x
)))
;
by
repeat
intro
.
...
...
@@ -113,7 +113,7 @@ Proof.
destruct
x
as
[[[]|]]
;
naive_solver
eauto
using
cmra_validN_includedN
.
Qed
.
Lemma
auth_valid_discrete
`
{
CmraDiscrete
A
}
x
:
Lemma
auth_valid_discrete
`
{
!
CmraDiscrete
A
}
x
:
✓
x
↔
match
authoritative
x
with
|
Excl'
a
=>
auth_own
x
≼
a
∧
✓
a
|
None
=>
✓
auth_own
x
...
...
@@ -125,12 +125,12 @@ Proof.
Qed
.
Lemma
auth_validN_2
n
a
b
:
✓
{
n
}
(
●
a
⋅
◯
b
)
↔
b
≼
{
n
}
a
∧
✓
{
n
}
a
.
Proof
.
by
rewrite
auth_validN_eq
/=
left_id
.
Qed
.
Lemma
auth_valid_discrete_2
`
{
CmraDiscrete
A
}
a
b
:
✓
(
●
a
⋅
◯
b
)
↔
b
≼
a
∧
✓
a
.
Lemma
auth_valid_discrete_2
`
{
!
CmraDiscrete
A
}
a
b
:
✓
(
●
a
⋅
◯
b
)
↔
b
≼
a
∧
✓
a
.
Proof
.
by
rewrite
auth_valid_discrete
/=
left_id
.
Qed
.
Lemma
authoritative_valid
x
:
✓
x
→
✓
authoritative
x
.
Proof
.
by
destruct
x
as
[[[]|]].
Qed
.
Lemma
auth_own_valid
`
{
CmraDiscrete
A
}
x
:
✓
x
→
✓
auth_own
x
.
Lemma
auth_own_valid
`
{
!
CmraDiscrete
A
}
x
:
✓
x
→
✓
auth_own
x
.
Proof
.
rewrite
auth_valid_discrete
.
destruct
x
as
[[[]|]]
;
naive_solver
eauto
using
cmra_valid_included
.
...
...
theories/algebra/excl.v
View file @
da93f357
...
...
@@ -52,7 +52,7 @@ Proof.
Qed
.
Canonical
Structure
exclC
:
ofeT
:
=
OfeT
(
excl
A
)
excl_ofe_mixin
.
Global
Instance
excl_cofe
`
{
Cofe
A
}
:
Cofe
exclC
.
Global
Instance
excl_cofe
`
{
!
Cofe
A
}
:
Cofe
exclC
.
Proof
.
apply
(
iso_cofe
(
from_option
Excl
ExclBot
)
(
maybe
Excl
)).
-
by
intros
n
[
a
|]
[
b
|]
;
split
;
inversion_clear
1
;
constructor
.
...
...
theories/algebra/gmap.v
View file @
da93f357
...
...
@@ -353,7 +353,7 @@ Qed.
Section
freshness
.
Local
Set
Default
Proof
Using
"Type*"
.
Context
`
{
Infinite
K
}.
Context
`
{
!
Infinite
K
}.
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
.
...
...
theories/algebra/gset.v
View file @
da93f357
...
...
@@ -163,7 +163,7 @@ Section gset_disj.
Section
fresh_updates
.
Local
Set
Default
Proof
Using
"Type*"
.
Context
`
{
Infinite
K
}.
Context
`
{
!
Infinite
K
}.
Lemma
gset_disj_alloc_updateP
(
Q
:
gset_disj
K
→
Prop
)
X
:
(
∀
i
,
i
∉
X
→
Q
(
GSet
({[
i
]}
∪
X
)))
→
GSet
X
~~>
:
Q
.
...
...
theories/algebra/list.v
View file @
da93f357
...
...
@@ -58,12 +58,12 @@ Program Definition list_chain
(
c
:
chain
listC
)
(
x
:
A
)
(
k
:
nat
)
:
chain
A
:
=
{|
chain_car
n
:
=
default
x
(
c
n
!!
k
)
|}.
Next
Obligation
.
intros
c
x
k
n
i
?.
by
rewrite
/=
(
chain_cauchy
c
n
i
).
Qed
.
Definition
list_compl
`
{
Cofe
A
}
:
Compl
listC
:
=
λ
c
,
Definition
list_compl
`
{
!
Cofe
A
}
:
Compl
listC
:
=
λ
c
,
match
c
0
with
|
[]
=>
[]
|
x
::
_
=>
compl
∘
list_chain
c
x
<$>
seq
0
(
length
(
c
0
))
end
.
Global
Program
Instance
list_cofe
`
{
Cofe
A
}
:
Cofe
listC
:
=
Global
Program
Instance
list_cofe
`
{
!
Cofe
A
}
:
Cofe
listC
:
=
{|
compl
:
=
list_compl
|}.
Next
Obligation
.
intros
?
n
c
;
rewrite
/
compl
/
list_compl
.
...
...
theories/base_logic/bi.v
View file @
da93f357
...
...
@@ -202,7 +202,7 @@ Lemma option_validI {A : cmraT} (mx : option A) :
Proof
.
exact
:
uPred_primitive
.
option_validI
.
Qed
.
Lemma
discrete_valid
{
A
:
cmraT
}
`
{!
CmraDiscrete
A
}
(
a
:
A
)
:
✓
a
⊣
⊢
⌜✓
a
⌝
.
Proof
.
exact
:
uPred_primitive
.
discrete_valid
.
Qed
.
Lemma
ofe_fun_validI
`
{
B
:
A
→
ucmraT
}
(
g
:
ofe_fun
B
)
:
✓
g
⊣
⊢
∀
i
,
✓
g
i
.
Lemma
ofe_fun_validI
{
A
}
{
B
:
A
→
ucmraT
}
(
g
:
ofe_fun
B
)
:
✓
g
⊣
⊢
∀
i
,
✓
g
i
.
Proof
.
exact
:
uPred_primitive
.
ofe_fun_validI
.
Qed
.
(** Consistency/soundness statement *)
...
...
theories/base_logic/derived.v
View file @
da93f357
...
...
@@ -60,7 +60,7 @@ Proof.
Qed
.
(** Timeless instances *)
Global
Instance
valid_timeless
{
A
:
cmraT
}
`
{
CmraDiscrete
A
}
(
a
:
A
)
:
Global
Instance
valid_timeless
{
A
:
cmraT
}
`
{
!
CmraDiscrete
A
}
(
a
:
A
)
:
Timeless
(
✓
a
:
uPred
M
)%
I
.
Proof
.
rewrite
/
Timeless
!
discrete_valid
.
apply
(
timeless
_
).
Qed
.
Global
Instance
ownM_timeless
(
a
:
M
)
:
Discrete
a
→
Timeless
(
uPred_ownM
a
).
...
...
theories/base_logic/lib/auth.v
View file @
da93f357
...
...
@@ -16,7 +16,7 @@ Instance subG_authΣ Σ A : subG (authΣ A) Σ → CmraDiscrete A → authG Σ A
Proof
.
solve_inG
.
Qed
.
Section
definitions
.
Context
`
{
invG
Σ
,
authG
Σ
A
}
{
T
:
Type
}
(
γ
:
gname
).
Context
`
{
!
invG
Σ
,
!
authG
Σ
A
}
{
T
:
Type
}
(
γ
:
gname
).
Definition
auth_own
(
a
:
A
)
:
iProp
Σ
:
=
own
γ
(
◯
a
).
...
...
@@ -60,7 +60,7 @@ Instance: Params (@auth_inv) 5 := {}.
Instance
:
Params
(@
auth_ctx
)
7
:
=
{}.
Section
auth
.
Context
`
{
invG
Σ
,
authG
Σ
A
}.
Context
`
{
!
invG
Σ
,
!
authG
Σ
A
}.
Context
{
T
:
Type
}
`
{!
Inhabited
T
}.
Context
(
f
:
T
→
A
)
(
φ
:
T
→
iProp
Σ
).
Implicit
Types
N
:
namespace
.
...
...
theories/base_logic/lib/boxes.v
View file @
da93f357
...
...
@@ -17,7 +17,7 @@ Instance subG_boxΣ Σ : subG boxΣ Σ → boxG Σ.
Proof
.
solve_inG
.
Qed
.
Section
box_defs
.
Context
`
{
invG
Σ
,
boxG
Σ
}
(
N
:
namespace
).
Context
`
{
!
invG
Σ
,
!
boxG
Σ
}
(
N
:
namespace
).
Definition
slice_name
:
=
gname
.
...
...
@@ -46,7 +46,7 @@ Instance: Params (@slice) 5 := {}.
Instance
:
Params
(@
box
)
5
:
=
{}.
Section
box
.
Context
`
{
invG
Σ
,
boxG
Σ
}
(
N
:
namespace
).
Context
`
{
!
invG
Σ
,
!
boxG
Σ
}
(
N
:
namespace
).
Implicit
Types
P
Q
:
iProp
Σ
.
Global
Instance
box_own_prop_ne
γ
:
NonExpansive
(
box_own_prop
γ
).
...
...
Prev
1
2
3
Next
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