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
Iris
examples
Commits
778406b2
Commit
778406b2
authored
Oct 06, 2018
by
Robbert Krebbers
Browse files
Bump Iris.
parent
b9e22933
Pipeline
#12146
passed with stage
in 10 minutes and 20 seconds
Changes
16
Pipelines
2
Hide whitespace changes
Inline
Side-by-side
opam
View file @
778406b2
...
...
@@ -9,6 +9,6 @@ build: [make "-j%{jobs}%"]
install: [make "install"]
remove: ["rm" "-rf" "%{lib}%/coq/user-contrib/iris_examples"]
depends: [
"coq-iris" { (= "dev.2018-07-13.2.
af5611c8
") | (= "dev") }
"coq-iris" { (= "dev.2018-07-13.2.
464c2449
") | (= "dev") }
"coq-autosubst" { = "dev.coq86" }
]
theories/barrier/example_client.v
View file @
778406b2
...
...
@@ -61,7 +61,7 @@ Section ClosedProofs.
Let
Σ
:
gFunctors
:
=
#[
heap
Σ
;
barrier
Σ
;
spawn
Σ
].
Lemma
client_adequate
σ
:
adequate
NotStuck
client
σ
(
λ
_
,
True
).
Lemma
client_adequate
σ
:
adequate
NotStuck
client
σ
(
λ
_
_
,
True
).
Proof
.
apply
(
heap_adequacy
Σ
)=>
?.
apply
client_safe
.
Qed
.
End
ClosedProofs
.
...
...
theories/hocap/cg_bag.v
View file @
778406b2
...
...
@@ -15,7 +15,9 @@ Set Default Proof Using "Type".
(** Coarse-grained bag implementation using a spin lock *)
Definition
newBag
:
val
:
=
λ
:
<>,
(
ref
NONE
,
newlock
#()).
let
:
"r"
:
=
ref
NONE
in
let
:
"l"
:
=
newlock
#()
in
(
"r"
,
"l"
).
Definition
pushBag
:
val
:
=
λ
:
"b"
"v"
,
let
:
"l"
:
=
Snd
"b"
in
let
:
"r"
:
=
Fst
"b"
in
...
...
@@ -102,12 +104,11 @@ Section proof.
Proof
.
iIntros
(
Φ
)
"_ HΦ"
.
unfold
newBag
.
wp_rec
.
wp_alloc
r
as
"Hr"
.
wp_alloc
r
as
"Hr"
.
wp_let
.
iMod
(
own_alloc
(
1
%
Qp
,
to_agree
∅
))
as
(
γ
b
)
"[Ha Hf]"
;
first
done
.
wp_apply
(
newlock_spec
N
(
bag_inv
γ
b
r
)
with
"[Hr Ha]"
).
{
iExists
[].
iFrame
.
}
iIntros
(
lk
γ
)
"#Hlk"
.
iApply
wp_value
.
iApply
"HΦ"
.
iIntros
(
lk
γ
)
"#Hlk"
.
wp_let
.
iApply
"HΦ"
.
rewrite
/
is_bag
/
bag_contents
.
iFrame
.
iExists
_
,
_
,
_
.
by
iFrame
"Hlk"
.
Qed
.
...
...
theories/hocap/concurrent_runners.v
View file @
778406b2
...
...
@@ -221,8 +221,8 @@ Section contents.
Proof
.
iIntros
(
Φ
)
"[#Hrunner HP] HΦ"
.
unfold
newTask
.
do
2
wp_rec
.
iApply
wp_fupd
.
wp_alloc
status
as
"Hstatus"
.
wp_alloc
res
as
"Hres"
.
wp_alloc
status
as
"Hstatus"
.
iMod
(
new_pending
)
as
(
γ
)
"[Htoken Htask]"
.
iMod
(
new_INIT
)
as
(
γ
'
)
"[Hinit Hinit']"
.
iMod
(
inv_alloc
(
N
.@
"task"
)
_
(
task_inv
γ
γ
'
status
res
(
Q
a
))%
I
with
"[-HP HΦ Htask Hinit]"
)
as
"#Hinv"
.
...
...
@@ -377,12 +377,12 @@ Section contents.
iL
ö
b
as
"IH"
forall
(
i
).
wp_rec
.
wp_op
.
case_bool_decide
;
wp_if
;
last
first
.
{
by
iApply
"HΦ"
.
}
wp_bind
(
Fork
_
).
iApply
wp_fork
.
iSplitL
.
wp_bind
(
Fork
_
).
iApply
(
wp_fork
with
"[]"
).
-
iNext
.
by
iApply
runner_runTasks_spec
.
-
iNext
.
wp_rec
.
wp_op
.
(* Set Printing Coercions. *)
assert
((
Z
.
of_nat
i
+
1
)
=
Z
.
of_nat
(
i
+
1
))
as
->
by
lia
.
iApply
(
"IH"
with
"HΦ"
).
-
iNext
.
by
iApply
runner_runTasks_spec
.
Qed
.
Lemma
newRunner_spec
P
Q
(
fe
ne
:
expr
)
(
f
:
val
)
(
n
:
nat
)
:
...
...
theories/hocap/parfib.v
View file @
778406b2
...
...
@@ -26,6 +26,7 @@ Section contents.
|
S
n'
=>
fib
n
+
fib
n'
end
end
.
Arguments
fib
!
_
/
:
simpl
nomatch
.
Definition
seqFib
:
val
:
=
rec
:
"fib"
"a"
:
=
if
:
"a"
=
#
0
...
...
@@ -38,30 +39,21 @@ Section contents.
{{{
True
}}}
seqFib
#
n
{{{
(
m
:
nat
),
RET
#
m
;
⌜
fib
n
=
m
⌝
}}}.
Proof
.
iIntros
(
Φ
)
"_ HΦ"
.
iL
ö
b
as
"IH"
forall
(
n
Φ
).
rewrite
/
seqFib
.
wp_rec
.
wp_op
.
case_bool_decide
;
simplify_eq
;
wp_if
.
iL
ö
b
as
"IH"
forall
(
n
Φ
).
wp_rec
.
simpl
.
wp_op
.
case_bool_decide
;
simplify_eq
;
wp_if
.
{
assert
(
n
=
O
)
as
->
by
lia
.
assert
(
1
=
S
O
)
as
->
by
lia
.
by
iApply
"HΦ"
.
}
by
iApply
(
"HΦ"
$!
1
%
nat
).
}
wp_op
.
case_bool_decide
;
simplify_eq
;
wp_if
.
{
assert
(
n
=
S
O
)
as
->
by
lia
.
assert
(
1
=
S
O
)
as
->
by
lia
.
by
iApply
"HΦ"
.
}
wp_op
.
wp_bind
((
rec
:
"fib"
"a"
:
=
_
)%
V
#(
n
-
1
)).
assert
(
∃
m
,
n
=
S
(
S
m
))
as
[
m
->].
{
assert
(
n
≠
O
)
by
naive_solver
.
assert
(
n
≠
S
O
)
by
naive_solver
.
exists
(
n
-
(
S
(
S
O
)))%
nat
.
lia
.
}
assert
((
S
(
S
m
)
-
1
)
=
S
m
)
as
->
by
lia
.
iApply
"IH"
.
iNext
.
iIntros
(?
<-).
assert
(
Z
.
of_nat
(
S
(
S
m
))
=
m
+
2
)
as
->
by
lia
.
Local
Opaque
fib
.
wp_op
.
assert
(
m
+
2
-
2
=
m
)
as
->
by
lia
.
wp_bind
((
rec
:
"fib"
"a"
:
=
_
)%
V
#
m
).
iApply
"IH"
.
iNext
.
iIntros
(?
<-).
wp_op
.
rewrite
-
Nat2Z
.
inj_add
.
iApply
"HΦ"
.
iPureIntro
.
Local
Transparent
fib
.
done
.
{
assert
(
n
=
1
%
nat
)
as
->
by
lia
.
by
iApply
(
"HΦ"
$!
1
%
nat
).
}
wp_op
.
destruct
n
as
[|[|
n
]]=>
//.
rewrite
!
Nat2Z
.
inj_succ
.
replace
(
Z
.
succ
(
Z
.
succ
n
)
-
2
)
with
(
Z
.
of_nat
n
)
by
lia
.
wp_apply
"IH"
.
iIntros
(?
<-).
wp_op
.
replace
(
Z
.
succ
(
Z
.
succ
n
)
-
1
)
with
(
Z
.
of_nat
(
S
n
))
by
lia
.
wp_apply
"IH"
.
iIntros
(?
<-).
wp_op
.
rewrite
-
Nat2Z
.
inj_add
.
by
iApply
"HΦ"
.
Qed
.
Definition
parFib
:
val
:
=
λ
:
"r"
"a"
,
...
...
@@ -85,38 +77,30 @@ Section contents.
Proof
.
iIntros
(
Φ
)
"[#Hrunner HP] HΦ"
.
iDestruct
"HP"
as
(
n
)
"%"
;
simplify_eq
.
do
2
wp_rec
.
wp_op
.
case_bool_decide
;
wp_if
.
do
2
wp_rec
.
simpl
.
wp_op
.
case_bool_decide
;
wp_if
.
-
iApply
seqFib_spec
;
eauto
.
iNext
.
iIntros
(?
<-).
iApply
"HΦ"
.
iExists
n
;
done
.
-
do
2
(
wp_op
;
wp_let
).
assert
(
∃
m
:
nat
,
n
=
S
(
S
m
))
as
[
m
->].
{
exists
(
n
-
2
)%
nat
.
lia
.
}
assert
((
S
(
S
m
)
-
1
)
=
S
m
)
as
->
by
lia
.
assert
((
S
(
S
m
)
-
2
)
=
m
)
as
->
by
lia
.
wp_bind
(
runner_Fork
b
r
_
).
iApply
(
runner_Fork_spec
).
{
iFrame
"Hrunner"
.
iExists
(
S
m
).
eauto
.
}
iNext
.
iIntros
(
γ
1
γ
1
'
t1
)
"Ht1"
.
wp_let
.
wp_bind
(
runner_Fork
b
r
_
).
iApply
(
runner_Fork_spec
).
destruct
n
as
[|[|
n
]]=>
//.
rewrite
!
Nat2Z
.
inj_succ
.
replace
(
Z
.
succ
(
Z
.
succ
n
)
-
2
)
with
(
Z
.
of_nat
n
)
by
lia
.
replace
(
Z
.
succ
(
Z
.
succ
n
)
-
1
)
with
(
Z
.
succ
n
)
by
lia
.
wp_apply
(
runner_Fork_spec
).
{
iFrame
"Hrunner"
.
iExists
(
S
n
).
by
rewrite
Nat2Z
.
inj_succ
.
}
iIntros
(
γ
1
γ
1
'
t1
)
"Ht1"
.
wp_let
.
wp_apply
(
runner_Fork_spec
).
{
iFrame
"Hrunner"
.
eauto
.
}
iNext
.
iIntros
(
γ
2
γ
2
'
t2
)
"Ht2"
.
wp_let
.
wp_bind
(
task_Join
_
).
iApply
(
task_Join_spec
with
"[$Ht1]"
)
;
try
done
.
iNext
.
iIntros
(
v1
)
"HQ"
;
simplify_eq
.
iDestruct
"HQ"
as
(
m1'
m1
)
"%"
.
simplify_eq
.
assert
(
m1'
=
S
m
)
as
->
by
lia
.
Local
Opaque
fib
.
wp_bind
(
task_Join
_
).
iApply
(
task_Join_spec
with
"[$Ht2]"
)
;
try
done
.
iNext
.
iIntros
(
v2
)
"HQ"
;
simplify_eq
.
iIntros
(
γ
2
γ
2
'
t2
)
"Ht2"
.
wp_let
.
wp_apply
(
task_Join_spec
with
"[$Ht2]"
)
;
try
done
.
iIntros
(
v2
)
"HQ"
;
simplify_eq
.
iDestruct
"HQ"
as
(
m2'
m2
)
"%"
.
simplify_eq
.
wp_op
.
iApply
"HΦ"
.
Local
Transparent
fib
.
iExists
(
S
(
S
m2'
)).
iSplit
;
eauto
.
assert
((
fib
(
S
m2'
)
+
fib
m2'
)
=
(
fib
(
S
m2'
)
+
fib
m2'
)%
nat
)
as
->
by
lia
.
done
.
wp_apply
(
task_Join_spec
with
"[$Ht1]"
)
;
try
done
.
iIntros
(
v1
)
"HQ"
;
simplify_eq
.
iDestruct
"HQ"
as
(
m1'
m1
)
"%"
.
simplify_eq
.
wp_op
.
iApply
"HΦ"
.
iExists
(
S
(
S
m2'
)).
rewrite
!
Nat2Z
.
inj_succ
-
Nat2Z
.
inj_add
.
by
replace
m1'
with
(
S
m2'
)
by
lia
.
Qed
.
Lemma
fibRunner_spec
(
n
a
:
nat
)
:
...
...
@@ -136,4 +120,3 @@ Section contents.
by
iApply
"HΦ"
.
Qed
.
End
contents
.
theories/logrel/F_mu/rules.v
View file @
778406b2
...
...
@@ -24,35 +24,36 @@ Section lang_rules.
Local
Ltac
solve_exec_safe
:
=
intros
;
subst
;
do
3
eexists
;
econstructor
;
eauto
.
Local
Ltac
solve_exec_puredet
:
=
simpl
;
intros
;
by
inv_head_step
.
Local
Ltac
solve_pure_exec
:
=
unfold
IntoVal
in
*
;
subst
;
repeat
match
goal
with
H
:
AsVal
_
|-
_
=>
destruct
H
as
[??]
end
;
apply
det_head_step_pure_exec
;
[
solve_exec_safe
|
solve_exec_puredet
].
unfold
IntoVal
in
*
;
repeat
match
goal
with
H
:
AsVal
_
|-
_
=>
destruct
H
as
[??]
end
;
subst
;
intros
?
;
apply
nsteps_once
,
pure_head_step_pure_step
;
constructor
;
[
solve_exec_safe
|
solve_exec_puredet
].
Global
Instance
pure_lam
e1
e2
`
{!
AsVal
e2
}
:
PureExec
True
(
App
(
Lam
e1
)
e2
)
e1
.[
e2
/].
PureExec
True
1
(
App
(
Lam
e1
)
e2
)
e1
.[
e2
/].
Proof
.
solve_pure_exec
.
Qed
.
Global
Instance
pure_tlam
e
:
PureExec
True
(
TApp
(
TLam
e
))
e
.
Global
Instance
pure_tlam
e
:
PureExec
True
1
(
TApp
(
TLam
e
))
e
.
Proof
.
solve_pure_exec
.
Qed
.
Global
Instance
pure_fold
e
`
{!
AsVal
e
}
:
PureExec
True
(
Unfold
(
Fold
e
))
e
.
PureExec
True
1
(
Unfold
(
Fold
e
))
e
.
Proof
.
solve_pure_exec
.
Qed
.
Global
Instance
pure_fst
e1
e2
`
{!
AsVal
e1
,
!
AsVal
e2
}
:
PureExec
True
(
Fst
(
Pair
e1
e2
))
e1
.
PureExec
True
1
(
Fst
(
Pair
e1
e2
))
e1
.
Proof
.
solve_pure_exec
.
Qed
.
Global
Instance
pure_snd
e1
e2
`
{!
AsVal
e1
,
!
AsVal
e2
}
:
PureExec
True
(
Snd
(
Pair
e1
e2
))
e2
.
PureExec
True
1
(
Snd
(
Pair
e1
e2
))
e2
.
Proof
.
solve_pure_exec
.
Qed
.
Global
Instance
pure_case_inl
e0
e1
e2
`
{!
AsVal
e0
}
:
PureExec
True
(
Case
(
InjL
e0
)
e1
e2
)
e1
.[
e0
/].
PureExec
True
1
(
Case
(
InjL
e0
)
e1
e2
)
e1
.[
e0
/].
Proof
.
solve_pure_exec
.
Qed
.
Global
Instance
pure_case_inr
e0
e1
e2
`
{!
AsVal
e0
}
:
PureExec
True
(
Case
(
InjR
e0
)
e1
e2
)
e2
.[
e0
/].
PureExec
True
1
(
Case
(
InjR
e0
)
e1
e2
)
e2
.[
e0
/].
Proof
.
solve_pure_exec
.
Qed
.
End
lang_rules
.
theories/logrel/F_mu/soundness.v
View file @
778406b2
...
...
@@ -7,7 +7,7 @@ Theorem soundness Σ `{invPreG Σ} e τ e' thp σ σ' :
rtc
step
([
e
],
σ
)
(
thp
,
σ
'
)
→
e'
∈
thp
→
is_Some
(
to_val
e'
)
∨
reducible
e'
σ
'
.
Proof
.
intros
Hlog
??.
cut
(
adequate
NotStuck
e
σ
(
λ
_
,
True
))
;
first
(
intros
[
_
?]
;
eauto
).
intros
Hlog
??.
cut
(
adequate
NotStuck
e
σ
(
λ
_
_
,
True
))
;
first
(
intros
[
_
?]
;
eauto
).
eapply
(
wp_adequacy
Σ
)
;
eauto
.
iIntros
(
Hinv
).
iModIntro
.
iExists
(
λ
_
,
True
%
I
).
iSplit
=>
//.
rewrite
-(
empty_env_subst
e
).
...
...
theories/logrel/F_mu_ref/rules.v
View file @
778406b2
...
...
@@ -94,34 +94,35 @@ Section lang_rules.
Local
Ltac
solve_exec_safe
:
=
intros
;
subst
;
do
3
eexists
;
econstructor
;
eauto
.
Local
Ltac
solve_exec_puredet
:
=
simpl
;
intros
;
by
inv_head_step
.
Local
Ltac
solve_pure_exec
:
=
unfold
IntoVal
in
*
;
subst
;
repeat
match
goal
with
H
:
AsVal
_
|-
_
=>
destruct
H
as
[??]
end
;
apply
det_head_step_pure_exec
;
[
solve_exec_safe
|
solve_exec_puredet
].
unfold
IntoVal
in
*
;
repeat
match
goal
with
H
:
AsVal
_
|-
_
=>
destruct
H
as
[??]
end
;
subst
;
intros
?
;
apply
nsteps_once
,
pure_head_step_pure_step
;
constructor
;
[
solve_exec_safe
|
solve_exec_puredet
].
Global
Instance
pure_lam
e1
e2
`
{!
AsVal
e2
}
:
PureExec
True
(
App
(
Lam
e1
)
e2
)
e1
.[
e2
/].
PureExec
True
1
(
App
(
Lam
e1
)
e2
)
e1
.[
e2
/].
Proof
.
solve_pure_exec
.
Qed
.
Global
Instance
pure_tlam
e
:
PureExec
True
(
TApp
(
TLam
e
))
e
.
Global
Instance
pure_tlam
e
:
PureExec
True
1
(
TApp
(
TLam
e
))
e
.
Proof
.
solve_pure_exec
.
Qed
.
Global
Instance
pure_fold
e
`
{!
AsVal
e
}
:
PureExec
True
(
Unfold
(
Fold
e
))
e
.
PureExec
True
1
(
Unfold
(
Fold
e
))
e
.
Proof
.
solve_pure_exec
.
Qed
.
Global
Instance
pure_fst
e1
e2
`
{!
AsVal
e1
,
!
AsVal
e2
}
:
PureExec
True
(
Fst
(
Pair
e1
e2
))
e1
.
PureExec
True
1
(
Fst
(
Pair
e1
e2
))
e1
.
Proof
.
solve_pure_exec
.
Qed
.
Global
Instance
pure_snd
e1
e2
`
{!
AsVal
e1
,
!
AsVal
e2
}
:
PureExec
True
(
Snd
(
Pair
e1
e2
))
e2
.
PureExec
True
1
(
Snd
(
Pair
e1
e2
))
e2
.
Proof
.
solve_pure_exec
.
Qed
.
Global
Instance
pure_case_inl
e0
e1
e2
`
{!
AsVal
e0
}
:
PureExec
True
(
Case
(
InjL
e0
)
e1
e2
)
e1
.[
e0
/].
PureExec
True
1
(
Case
(
InjL
e0
)
e1
e2
)
e1
.[
e0
/].
Proof
.
solve_pure_exec
.
Qed
.
Global
Instance
pure_case_inr
e0
e1
e2
`
{!
AsVal
e0
}
:
PureExec
True
(
Case
(
InjR
e0
)
e1
e2
)
e2
.[
e0
/].
PureExec
True
1
(
Case
(
InjR
e0
)
e1
e2
)
e2
.[
e0
/].
Proof
.
solve_pure_exec
.
Qed
.
End
lang_rules
.
theories/logrel/F_mu_ref/soundness.v
View file @
778406b2
...
...
@@ -13,7 +13,7 @@ Theorem soundness Σ `{heapPreG Σ} e τ e' thp σ σ' :
rtc
step
([
e
],
σ
)
(
thp
,
σ
'
)
→
e'
∈
thp
→
is_Some
(
to_val
e'
)
∨
reducible
e'
σ
'
.
Proof
.
intros
Hlog
??.
cut
(
adequate
NotStuck
e
σ
(
λ
_
,
True
))
;
first
(
intros
[
_
?]
;
eauto
).
intros
Hlog
??.
cut
(
adequate
NotStuck
e
σ
(
λ
_
_
,
True
))
;
first
(
intros
[
_
?]
;
eauto
).
eapply
(
wp_adequacy
Σ
_
)
;
eauto
.
iIntros
(
Hinv
).
iMod
(
own_alloc
(
●
to_gen_heap
σ
))
as
(
γ
)
"Hh"
.
...
...
theories/logrel/F_mu_ref/soundness_binary.v
View file @
778406b2
...
...
@@ -11,7 +11,7 @@ Lemma basic_soundness Σ `{heapPreG Σ, inG Σ (authR cfgUR)}
(
∃
thp'
hp'
v'
,
rtc
step
([
e'
],
∅
)
(
of_val
v'
::
thp'
,
hp'
)).
Proof
.
intros
Hlog
Hsteps
.
cut
(
adequate
NotStuck
e
∅
(
λ
_
,
∃
thp'
h
v
,
rtc
step
([
e'
],
∅
)
(
of_val
v
::
thp'
,
h
))).
cut
(
adequate
NotStuck
e
∅
(
λ
_
_
,
∃
thp'
h
v
,
rtc
step
([
e'
],
∅
)
(
of_val
v
::
thp'
,
h
))).
{
destruct
1
;
naive_solver
.
}
eapply
(
wp_adequacy
Σ
)
;
first
by
apply
_
.
iIntros
(
Hinv
).
...
...
theories/logrel/F_mu_ref_conc/rules.v
View file @
778406b2
...
...
@@ -124,47 +124,48 @@ Section lang_rules.
Local
Ltac
solve_exec_safe
:
=
intros
;
subst
;
do
3
eexists
;
econstructor
;
eauto
.
Local
Ltac
solve_exec_puredet
:
=
simpl
;
intros
;
by
inv_head_step
.
Local
Ltac
solve_pure_exec
:
=
unfold
IntoVal
in
*
;
subst
;
repeat
match
goal
with
H
:
AsVal
_
|-
_
=>
destruct
H
as
[??]
end
;
apply
det_head_step_pure_exec
;
[
solve_exec_safe
|
solve_exec_puredet
].
unfold
IntoVal
in
*
;
repeat
match
goal
with
H
:
AsVal
_
|-
_
=>
destruct
H
as
[??]
end
;
subst
;
intros
?
;
apply
nsteps_once
,
pure_head_step_pure_step
;
constructor
;
[
solve_exec_safe
|
solve_exec_puredet
].
Global
Instance
pure_rec
e1
e2
`
{!
AsVal
e2
}
:
PureExec
True
(
App
(
Rec
e1
)
e2
)
e1
.[(
Rec
e1
),
e2
/].
PureExec
True
1
(
App
(
Rec
e1
)
e2
)
e1
.[(
Rec
e1
),
e2
/].
Proof
.
solve_pure_exec
.
Qed
.
Global
Instance
pure_tlam
e
:
PureExec
True
(
TApp
(
TLam
e
))
e
.
Global
Instance
pure_tlam
e
:
PureExec
True
1
(
TApp
(
TLam
e
))
e
.
Proof
.
solve_pure_exec
.
Qed
.
Global
Instance
pure_fold
e
`
{!
AsVal
e
}
:
PureExec
True
(
Unfold
(
Fold
e
))
e
.
PureExec
True
1
(
Unfold
(
Fold
e
))
e
.
Proof
.
solve_pure_exec
.
Qed
.
Global
Instance
pure_fst
e1
e2
`
{!
AsVal
e1
,
!
AsVal
e2
}
:
PureExec
True
(
Fst
(
Pair
e1
e2
))
e1
.
PureExec
True
1
(
Fst
(
Pair
e1
e2
))
e1
.
Proof
.
solve_pure_exec
.
Qed
.
Global
Instance
pure_snd
e1
e2
`
{!
AsVal
e1
,
!
AsVal
e2
}
:
PureExec
True
(
Snd
(
Pair
e1
e2
))
e2
.
PureExec
True
1
(
Snd
(
Pair
e1
e2
))
e2
.
Proof
.
solve_pure_exec
.
Qed
.
Global
Instance
pure_case_inl
e0
e1
e2
`
{!
AsVal
e0
}
:
PureExec
True
(
Case
(
InjL
e0
)
e1
e2
)
e1
.[
e0
/].
PureExec
True
1
(
Case
(
InjL
e0
)
e1
e2
)
e1
.[
e0
/].
Proof
.
solve_pure_exec
.
Qed
.
Global
Instance
pure_case_inr
e0
e1
e2
`
{!
AsVal
e0
}
:
PureExec
True
(
Case
(
InjR
e0
)
e1
e2
)
e2
.[
e0
/].
PureExec
True
1
(
Case
(
InjR
e0
)
e1
e2
)
e2
.[
e0
/].
Proof
.
solve_pure_exec
.
Qed
.
Global
Instance
wp_if_true
e1
e2
:
PureExec
True
(
If
(#
♭
true
)
e1
e2
)
e1
.
PureExec
True
1
(
If
(#
♭
true
)
e1
e2
)
e1
.
Proof
.
solve_pure_exec
.
Qed
.
Global
Instance
wp_if_false
e1
e2
:
PureExec
True
(
If
(#
♭
false
)
e1
e2
)
e2
.
PureExec
True
1
(
If
(#
♭
false
)
e1
e2
)
e2
.
Proof
.
solve_pure_exec
.
Qed
.
Global
Instance
wp_nat_binop
op
a
b
:
PureExec
True
(
BinOp
op
(#
n
a
)
(#
n
b
))
(
of_val
(
binop_eval
op
a
b
)).
PureExec
True
1
(
BinOp
op
(#
n
a
)
(#
n
b
))
(
of_val
(
binop_eval
op
a
b
)).
Proof
.
solve_pure_exec
.
Qed
.
End
lang_rules
.
theories/logrel/F_mu_ref_conc/soundness_binary.v
View file @
778406b2
...
...
@@ -11,7 +11,7 @@ Lemma basic_soundness Σ `{heapPreIG Σ, inG Σ (authR cfgUR)}
(
∃
thp'
hp'
v'
,
rtc
step
([
e'
],
∅
)
(
of_val
v'
::
thp'
,
hp'
)).
Proof
.
intros
Hlog
Hsteps
.
cut
(
adequate
NotStuck
e
∅
(
λ
_
,
∃
thp'
h
v
,
rtc
step
([
e'
],
∅
)
(
of_val
v
::
thp'
,
h
))).
cut
(
adequate
NotStuck
e
∅
(
λ
_
_
,
∃
thp'
h
v
,
rtc
step
([
e'
],
∅
)
(
of_val
v
::
thp'
,
h
))).
{
destruct
1
;
naive_solver
.
}
eapply
(
wp_adequacy
Σ
_
)
;
iIntros
(
Hinv
).
iMod
(
own_alloc
(
●
to_gen_heap
∅
))
as
(
γ
)
"Hh"
.
...
...
theories/logrel/F_mu_ref_conc/soundness_unary.v
View file @
778406b2
...
...
@@ -13,7 +13,7 @@ Theorem soundness Σ `{heapPreIG Σ} e τ e' thp σ σ' :
rtc
step
([
e
],
σ
)
(
thp
,
σ
'
)
→
e'
∈
thp
→
is_Some
(
to_val
e'
)
∨
reducible
e'
σ
'
.
Proof
.
intros
Hlog
??.
cut
(
adequate
NotStuck
e
σ
(
λ
_
,
True
))
;
first
(
intros
[
_
?]
;
eauto
).
intros
Hlog
??.
cut
(
adequate
NotStuck
e
σ
(
λ
_
_
,
True
))
;
first
(
intros
[
_
?]
;
eauto
).
eapply
(
wp_adequacy
Σ
_
).
iIntros
(
Hinv
).
iMod
(
own_alloc
(
●
to_gen_heap
σ
))
as
(
γ
)
"Hh"
.
{
apply
(
auth_auth_valid
_
(
to_gen_heap_valid
_
_
σ
)).
}
...
...
theories/logrel/stlc/rules.v
View file @
778406b2
...
...
@@ -25,30 +25,31 @@ Section stlc_rules.
Local
Ltac
solve_exec_safe
:
=
intros
;
subst
;
do
3
eexists
;
econstructor
;
eauto
.
Local
Ltac
solve_exec_puredet
:
=
simpl
;
intros
;
by
inv_head_step
.
Local
Ltac
solve_pure_exec
:
=
unfold
IntoVal
in
*
;
subst
;
repeat
match
goal
with
H
:
AsVal
_
|-
_
=>
destruct
H
as
[??]
end
;
apply
det_head_step_pure_exec
;
[
solve_exec_safe
|
solve_exec_puredet
].
unfold
IntoVal
in
*
;
repeat
match
goal
with
H
:
AsVal
_
|-
_
=>
destruct
H
as
[??]
end
;
subst
;
intros
?
;
apply
nsteps_once
,
pure_head_step_pure_step
;
constructor
;
[
solve_exec_safe
|
solve_exec_puredet
].
(** Helper Lemmas for weakestpre. *)
Global
Instance
pure_lam
e1
e2
`
{!
AsVal
e2
}
:
PureExec
True
(
App
(
Lam
e1
)
e2
)
(
e1
.[
e2
/]).
PureExec
True
1
(
App
(
Lam
e1
)
e2
)
(
e1
.[
e2
/]).
Proof
.
solve_pure_exec
.
Qed
.
Global
Instance
pure_fst
e1
e2
`
{!
AsVal
e1
,
!
AsVal
e2
}
:
PureExec
True
(
Fst
(
Pair
e1
e2
))
e1
.
PureExec
True
1
(
Fst
(
Pair
e1
e2
))
e1
.
Proof
.
solve_pure_exec
.
Qed
.
Global
Instance
pure_snd
e1
e2
`
{!
AsVal
e1
,
!
AsVal
e2
}
:
PureExec
True
(
Snd
(
Pair
e1
e2
))
e2
.
PureExec
True
1
(
Snd
(
Pair
e1
e2
))
e2
.
Proof
.
solve_pure_exec
.
Qed
.
Global
Instance
pure_case_inl
e0
e1
e2
`
{!
AsVal
e0
}
:
PureExec
True
(
Case
(
InjL
e0
)
e1
e2
)
e1
.[
e0
/].
PureExec
True
1
(
Case
(
InjL
e0
)
e1
e2
)
e1
.[
e0
/].
Proof
.
solve_pure_exec
.
Qed
.
Global
Instance
pure_case_inr
e0
e1
e2
`
{!
AsVal
e0
}
:
PureExec
True
(
Case
(
InjR
e0
)
e1
e2
)
e2
.[
e0
/].
PureExec
True
1
(
Case
(
InjR
e0
)
e1
e2
)
e2
.[
e0
/].
Proof
.
solve_pure_exec
.
Qed
.
End
stlc_rules
.
theories/logrel/stlc/soundness.v
View file @
778406b2
...
...
@@ -13,7 +13,7 @@ Theorem soundness e τ e' thp :
is_Some
(
to_val
e'
)
∨
reducible
e'
().
Proof
.
set
(
Σ
:
=
inv
Σ
).
intros
.
cut
(
adequate
NotStuck
e
()
(
λ
_
,
True
))
;
first
(
intros
[
_
Hsafe
]
;
eauto
).
cut
(
adequate
NotStuck
e
()
(
λ
_
_
,
True
))
;
first
(
intros
[
_
Hsafe
]
;
eauto
).
eapply
(
wp_adequacy
Σ
_
).
iIntros
(
Hinv
).
iModIntro
.
iExists
(
λ
_
,
True
%
I
).
iSplit
=>//.
set
(
H
Σ
:
=
IrisG
_
_
Hinv
(
λ
_
,
True
)%
I
).
...
...
theories/spanning_tree/spanning.v
View file @
778406b2
...
...
@@ -210,7 +210,7 @@ Section Helpers.
iIntros
(
u
)
"(H1 & Hagree & Hx & key)"
.
iDestruct
"H1"
as
(
m
)
"[Hmrk Hu]"
.
iDestruct
"Hagree"
as
%
Hagree
.
iDestruct
"Hmrk"
as
%
Hmrk
;
iDestruct
"Hu"
as
%
Hu
;
subst
.
wp_let
.
wp_bind
(
Fst
_
).
do
3
wp_proj
.
wp_let
.
do
3
wp_proj
.
iApply
(
wp_store_graph
_
(
None
,
w2
)
with
"[Hx key]"
)
;
eauto
;
[|
iFrame
"Hgr"
;
by
iFrame
].
{
by
destruct
w1
;
destruct
w2
;
destruct
v
;
inversion
Hagree
;
subst
.
}
...
...
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