aboutsummaryrefslogtreecommitdiffstats
path: root/mppa_k1c/Asmblock.v
diff options
context:
space:
mode:
Diffstat (limited to 'mppa_k1c/Asmblock.v')
-rw-r--r--mppa_k1c/Asmblock.v25
1 files changed, 25 insertions, 0 deletions
diff --git a/mppa_k1c/Asmblock.v b/mppa_k1c/Asmblock.v
index 91e5ac89..cce180ac 100644
--- a/mppa_k1c/Asmblock.v
+++ b/mppa_k1c/Asmblock.v
@@ -294,6 +294,31 @@ Fixpoint exec_body (body: list basic) (rs: regset) (m: mem): outcome :=
end
end.
+
+Theorem builtin_body_nil:
+ forall bb ef args res, exit bb = Some (PExpand (Pbuiltin ef args res)) -> body bb = nil.
+Proof.
+ intros. destruct bb as [hd bdy ex WF]. simpl in *.
+ apply wf_bblock_refl in WF. inv WF. unfold builtin_alone in H1.
+ eapply H1; eauto.
+Qed.
+
+Theorem exec_body_app:
+ forall l l' rs m rs'' m'',
+ exec_body (l ++ l') rs m = Next rs'' m'' ->
+ exists rs' m',
+ exec_body l rs m = Next rs' m'
+ /\ exec_body l' rs' m' = Next rs'' m''.
+Proof.
+ induction l.
+ - intros. simpl in H. repeat eexists. auto.
+ - intros. rewrite <- app_comm_cons in H. simpl in H.
+ destruct (exec_basic_instr a rs m) eqn:EXEBI.
+ + apply IHl in H. destruct H as (rs1 & m1 & EXEB1 & EXEB2).
+ repeat eexists. simpl. rewrite EXEBI. eauto. auto.
+ + discriminate.
+Qed.
+
(** Position corresponding to a label *)
Definition goto_label (f: function) (lbl: label) (rs: regset) (m: mem) : outcome := par_goto_label f lbl rs rs m.