aboutsummaryrefslogtreecommitdiffstats
path: root/cfrontend/Cexec.v
diff options
context:
space:
mode:
authorMaxime Dénès <mail@maximedenes.fr>2016-09-21 13:27:28 +0200
committerMaxime Dénès <mail@maximedenes.fr>2016-09-21 13:55:32 +0200
commit320f39edca5061b84a8a45b5c33e516049c7f7fe (patch)
tree8932bbf78dc6403e002a397ce4210f6a35b31eda /cfrontend/Cexec.v
parentf26267c6289e4fa306a0875ff149a00ee401e043 (diff)
downloadcompcert-kvx-320f39edca5061b84a8a45b5c33e516049c7f7fe.tar.gz
compcert-kvx-320f39edca5061b84a8a45b5c33e516049c7f7fe.zip
Fix minor issues in some proofs and tactics.
These minor problems were revealed by porting CompCert to Coq 8.6, where they trigger errors.
Diffstat (limited to 'cfrontend/Cexec.v')
-rw-r--r--cfrontend/Cexec.v13
1 files changed, 8 insertions, 5 deletions
diff --git a/cfrontend/Cexec.v b/cfrontend/Cexec.v
index bf88e033..b9deb204 100644
--- a/cfrontend/Cexec.v
+++ b/cfrontend/Cexec.v
@@ -217,8 +217,8 @@ Definition do_volatile_store (w: world) (chunk: memory_chunk) (m: mem) (b: block
Ltac mydestr :=
match goal with
- | [ |- None = Some _ -> _ ] => intro X; discriminate
- | [ |- Some _ = Some _ -> _ ] => intro X; inv X
+ | [ |- None = Some _ -> _ ] => let X := fresh "X" in intro X; discriminate
+ | [ |- Some _ = Some _ -> _ ] => let X := fresh "X" in intro X; inv X
| [ |- match ?x with Some _ => _ | None => _ end = Some _ -> _ ] => destruct x eqn:?; mydestr
| [ |- match ?x with true => _ | false => _ end = Some _ -> _ ] => destruct x eqn:?; mydestr
| [ |- match ?x with left _ => _ | right _ => _ end = Some _ -> _ ] => destruct x; mydestr
@@ -2061,12 +2061,15 @@ Definition do_step (w: world) (s: state) : list transition :=
Ltac myinv :=
match goal with
- | [ |- In _ nil -> _ ] => intro X; elim X
+ | [ |- In _ nil -> _ ] => let X := fresh "X" in intro X; elim X
| [ |- In _ (ret _ _) -> _ ] =>
+ let X := fresh "X" in
intro X; elim X; clear X;
- [intro EQ; unfold ret in EQ; inv EQ; myinv | myinv]
+ [let EQ := fresh "EQ" in intro EQ; unfold ret in EQ; inv EQ; myinv | myinv]
| [ |- In _ (_ :: nil) -> _ ] =>
- intro X; elim X; clear X; [intro EQ; inv EQ; myinv | myinv]
+ let X := fresh "X" in
+ intro X; elim X; clear X;
+ [let EQ := fresh "EQ" in intro EQ; inv EQ; myinv | myinv]
| [ |- In _ (match ?x with Some _ => _ | None => _ end) -> _ ] => destruct x eqn:?; myinv
| [ |- In _ (match ?x with false => _ | true => _ end) -> _ ] => destruct x eqn:?; myinv
| [ |- In _ (match ?x with left _ => _ | right _ => _ end) -> _ ] => destruct x; myinv