Skip to content
GitLab
Menu
Projects
Groups
Snippets
Help
Help
Support
Community forum
Keyboard shortcuts
?
Submit feedback
Contribute to GitLab
Sign in / Register
Toggle navigation
Menu
Open sidebar
Rice Wine
Iris
Commits
01663351
Commit
01663351
authored
Jun 29, 2018
by
Ralf Jung
Browse files
wp_cas_fail/succ: Report error when proving vals_cas_compare_safe; fast_done is enough
parent
6ac4c4eb
Changes
3
Show whitespace changes
Inline
Side-by-side
tests/heap_lang.ref
View file @
01663351
...
@@ -68,3 +68,8 @@
...
@@ -68,3 +68,8 @@
let: "val2" := fun2 "val1" in
let: "val2" := fun2 "val1" in
let: "val3" := fun3 "val2" in if: "val1" = "val2" then "val" else "val3"
let: "val3" := fun3 "val2" in if: "val1" = "val2" then "val" else "val3"
{{{ (x y : val) (z : Z), RET (x, y, #z); True }}}
{{{ (x y : val) (z : Z), RET (x, y, #z); True }}}
"not_cas_compare_safe"
: string
The command has indeed failed with message:
Ltac call to "wp_cas_suc" failed.
Tactic failure: wp_cas_suc: Values are not safe to compare for CAS.
tests/heap_lang.v
View file @
01663351
...
@@ -161,5 +161,16 @@ Section printing_tests.
...
@@ -161,5 +161,16 @@ Section printing_tests.
End
printing_tests
.
End
printing_tests
.
Section
error_tests
.
Context
`
{
heapG
Σ
}.
Check
"not_cas_compare_safe"
.
Lemma
not_cas_compare_safe
(
l
:
loc
)
(
v
:
val
)
:
l
↦
v
-
∗
WP
CAS
#
l
v
v
{{
_
,
True
}}.
Proof
.
iIntros
"H↦"
.
Fail
wp_cas_suc
.
Abort
.
End
error_tests
.
Lemma
heap_e_adequate
σ
:
adequate
NotStuck
heap_e
σ
(=
#
2
).
Lemma
heap_e_adequate
σ
:
adequate
NotStuck
heap_e
σ
(=
#
2
).
Proof
.
eapply
(
heap_adequacy
heap
Σ
)=>
?.
by
apply
heap_e_spec
.
Qed
.
Proof
.
eapply
(
heap_adequacy
heap
Σ
)=>
?.
by
apply
heap_e_spec
.
Qed
.
theories/heap_lang/proofmode.v
View file @
01663351
...
@@ -406,7 +406,7 @@ Tactic Notation "wp_cas_fail" :=
...
@@ -406,7 +406,7 @@ Tactic Notation "wp_cas_fail" :=
[
iSolveTC
[
iSolveTC
|
solve_mapsto
()
|
solve_mapsto
()
|
try
congruence
|
try
congruence
|
done
|
fast_done
||
fail
"wp_cas_fail: Values are not safe to compare for CAS"
|
simpl
;
try
wp_value_head
]
|
simpl
;
try
wp_value_head
]
|
|-
envs_entails
_
(
twp
?s
?E
?e
?Q
)
=>
|
|-
envs_entails
_
(
twp
?s
?E
?e
?Q
)
=>
first
first
...
@@ -415,7 +415,7 @@ Tactic Notation "wp_cas_fail" :=
...
@@ -415,7 +415,7 @@ Tactic Notation "wp_cas_fail" :=
|
fail
1
"wp_cas_fail: cannot find 'CAS' in"
e
]
;
|
fail
1
"wp_cas_fail: cannot find 'CAS' in"
e
]
;
[
solve_mapsto
()
[
solve_mapsto
()
|
try
congruence
|
try
congruence
|
done
|
fast_done
||
fail
"wp_cas_fail: Values are not safe to compare for CAS"
|
wp_expr_simpl
;
try
wp_value_head
]
|
wp_expr_simpl
;
try
wp_value_head
]
|
_
=>
fail
"wp_cas_fail: not a 'wp'"
|
_
=>
fail
"wp_cas_fail: not a 'wp'"
end
.
end
.
...
@@ -434,7 +434,7 @@ Tactic Notation "wp_cas_suc" :=
...
@@ -434,7 +434,7 @@ Tactic Notation "wp_cas_suc" :=
[
iSolveTC
[
iSolveTC
|
solve_mapsto
()
|
solve_mapsto
()
|
try
congruence
|
try
congruence
|
done
|
fast_done
||
fail
"wp_cas_suc: Values are not safe to compare for CAS"
|
pm_reflexivity
|
pm_reflexivity
|
simpl
;
try
wp_value_head
]
|
simpl
;
try
wp_value_head
]
|
|-
envs_entails
_
(
twp
?E
?e
?Q
)
=>
|
|-
envs_entails
_
(
twp
?E
?e
?Q
)
=>
...
@@ -444,7 +444,7 @@ Tactic Notation "wp_cas_suc" :=
...
@@ -444,7 +444,7 @@ Tactic Notation "wp_cas_suc" :=
|
fail
1
"wp_cas_suc: cannot find 'CAS' in"
e
]
;
|
fail
1
"wp_cas_suc: cannot find 'CAS' in"
e
]
;
[
solve_mapsto
()
[
solve_mapsto
()
|
try
congruence
|
try
congruence
|
done
|
fast_done
||
fail
"wp_cas_suc: Values are not safe to compare for CAS"
|
pm_reflexivity
|
pm_reflexivity
|
wp_expr_simpl
;
try
wp_value_head
]
|
wp_expr_simpl
;
try
wp_value_head
]
|
_
=>
fail
"wp_cas_suc: not a 'wp'"
|
_
=>
fail
"wp_cas_suc: not a 'wp'"
...
...
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