aboutsummaryrefslogtreecommitdiffstats
path: root/mppa_k1c/Asmblockgen.v
diff options
context:
space:
mode:
authorCyril SIX <cyril.six@kalray.eu>2018-10-12 18:23:12 +0200
committerCyril SIX <cyril.six@kalray.eu>2018-10-12 18:23:12 +0200
commit48d4481b6bd246e8912b3d01d98e0bc297f645de (patch)
tree23754c0b48bd4a771945e9de74704a17a3783138 /mppa_k1c/Asmblockgen.v
parent7e02f98ae8cb37c9d84df6d43b05104346fda691 (diff)
downloadcompcert-kvx-48d4481b6bd246e8912b3d01d98e0bc297f645de.tar.gz
compcert-kvx-48d4481b6bd246e8912b3d01d98e0bc297f645de.zip
Stuck on proving step_simu_body
Diffstat (limited to 'mppa_k1c/Asmblockgen.v')
-rw-r--r--mppa_k1c/Asmblockgen.v11
1 files changed, 8 insertions, 3 deletions
diff --git a/mppa_k1c/Asmblockgen.v b/mppa_k1c/Asmblockgen.v
index 28941e22..fce2cb47 100644
--- a/mppa_k1c/Asmblockgen.v
+++ b/mppa_k1c/Asmblockgen.v
@@ -850,13 +850,18 @@ Local Obligation Tactic := bblock_auto_correct.
(** Can generate two bblocks if the ctl is a PExpand (since the PExpand must be alone in its block) *)
Program Definition gen_bblocks (hd: list label) (c: list basic) (ctl: list instruction) :=
match (extract_ctl ctl) with
- | None => {| header := hd; body := Pnop::(c ++ (extract_basic ctl)); exit := None |} :: nil
+ | None => {| header := hd; body := (c ++ (Pnop :: extract_basic ctl)); exit := None |} :: nil
(* gen_bblock_noctl hd (c ++ (extract_basic ctl)) :: nil *)
- | Some (PExpand (Pbuiltin ef args res)) => ({| header := hd; body := Pnop::c; exit := None |}) ::
+ | Some (PExpand (Pbuiltin ef args res)) => ({| header := hd; body := c++(Pnop::nil); exit := None |}) ::
((PExpand (Pbuiltin ef args res)) ::b nil)
- | Some (PCtlFlow i) => {| header := hd; body := Pnop :: (c ++ (extract_basic ctl)); exit := Some (PCtlFlow i) |} :: nil
+ | Some (PCtlFlow i) => {| header := hd; body := (c ++ (Pnop :: extract_basic ctl)); exit := Some (PCtlFlow i) |} :: nil
end
.
+Next Obligation.
+ intros. constructor. intro. apply app_eq_nil in H. destruct H. discriminate.
+Qed. Next Obligation.
+ intros. constructor. intro. apply app_eq_nil in H. destruct H. discriminate.
+Qed.
Definition transl_block (f: Machblock.function) (fb: Machblock.bblock) : res (list bblock) :=
do c <- transl_basic_code' f fb.(Machblock.body) true;