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
Iris
Commits
bf42c065
Commit
bf42c065
authored
Jul 19, 2016
by
Robbert Krebbers
Browse files
Repair tree_sum test case and add to _CoqProject.
parent
f96e229a
Pipeline
#2291
passed with stage
Changes
2
Pipelines
1
Hide whitespace changes
Inline
Side-by-side
_CoqProject
View file @
bf42c065
...
...
@@ -106,6 +106,7 @@ tests/joining_existentials.v
tests/proofmode.v
tests/barrier_client.v
tests/list_reverse.v
tests/tree_sum.v
proofmode/coq_tactics.v
proofmode/pviewshifts.v
proofmode/environments.v
...
...
tests/tree_sum.v
View file @
bf42c065
...
...
@@ -21,16 +21,15 @@ Fixpoint sum (t : tree) : Z :=
Definition
sum_loop
:
val
:
=
rec
:
"sum_loop"
"t"
"l"
:
=
match
:
'
"t"
with
InjL
"n"
=>
'
"l"
<-
'
"n"
+
!
'
"l"
|
InjR
"tt"
=>
'
"sum_loop"
!(
Fst
'
"tt"
)
'
"l"
;;
'
"sum_loop"
!(
Snd
'
"tt"
)
'
"l"
match
:
"t"
with
InjL
"n"
=>
"l"
<-
"n"
+
!
"l"
|
InjR
"tt"
=>
"sum_loop"
!(
Fst
"tt"
)
"l"
;;
"sum_loop"
!(
Snd
"tt"
)
"l"
end
.
Definition
sum'
:
val
:
=
λ
:
"t"
,
let
:
"l"
:
=
ref
#
0
in
^
sum_loop
'
"t"
'
"l"
;;
!
'
"l"
.
sum_loop
"t"
"l"
;;
!
"l"
.
Global
Opaque
sum_loop
sum'
.
...
...
@@ -40,12 +39,12 @@ Lemma sum_loop_wp `{!heapG Σ} heapN v t l (n : Z) (Φ : val → iPropG heap_lan
⊢
WP
sum_loop
v
#
l
{{
Φ
}}.
Proof
.
iIntros
"(#Hh & Hl & Ht & HΦ)"
.
iL
ö
b
{
v
t
l
n
Φ
}
as
"IH"
.
wp_rec
;
wp_let
.
iL
ö
b
(
v
t
l
n
Φ
)
as
"IH"
.
wp_rec
.
wp_let
.
destruct
t
as
[
n'
|
tl
tr
]
;
simpl
in
*.
-
iDestruct
"Ht"
as
"%"
;
subst
.
wp_match
.
wp_load
.
wp_op
.
wp_store
.
by
iApply
(
"HΦ"
with
"Hl"
).
-
iDestruct
"Ht"
as
{
ll
lr
vl
vr
}
"(% & Hll & Htl & Hlr & Htr)"
;
subst
.
-
iDestruct
"Ht"
as
(
ll
lr
vl
vr
)
"(% & Hll & Htl & Hlr & Htr)"
;
subst
.
wp_match
.
wp_proj
.
wp_load
.
wp_apply
(
"IH"
with
"Hl Htl"
).
iIntros
"Hl Htl"
.
wp_seq
.
wp_proj
.
wp_load
.
...
...
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