aboutsummaryrefslogtreecommitdiffstats
diff options
context:
space:
mode:
-rw-r--r--backend/RTLtyping.v14
-rw-r--r--backend/ValueAnalysis.v28
-rw-r--r--mppa_k1c/ValueAOp.v20
3 files changed, 58 insertions, 4 deletions
diff --git a/backend/RTLtyping.v b/backend/RTLtyping.v
index 6d27df28..74bfa582 100644
--- a/backend/RTLtyping.v
+++ b/backend/RTLtyping.v
@@ -851,6 +851,16 @@ Proof.
eapply wt_regset_assign; eauto. rewrite H9; eapply Mem.load_type; eauto.
Qed.
+Lemma wt_exec_Iload_notrap:
+ forall env f chunk addr args dst s rs,
+ wt_instr f env (Iload NOTRAP chunk addr args dst s) ->
+ wt_regset env rs ->
+ wt_regset env (rs#dst <- (default_notrap_load_value chunk)).
+Proof.
+ intros.
+ eapply wt_regset_assign; eauto. simpl. trivial.
+Qed.
+
Lemma wt_exec_Ibuiltin:
forall env f ef (ge: genv) args res s vargs m t vres m' rs,
wt_instr f env (Ibuiltin ef args res s) ->
@@ -930,6 +940,10 @@ Proof.
econstructor; eauto. eapply wt_exec_Iop; eauto.
(* Iload *)
econstructor; eauto. eapply wt_exec_Iload; eauto.
+ (* Iload notrap1*)
+ econstructor; eauto. eapply wt_exec_Iload_notrap; eauto.
+ (* Iload notrap2*)
+ econstructor; eauto. eapply wt_exec_Iload_notrap; eauto.
(* Istore *)
econstructor; eauto.
(* Icall *)
diff --git a/backend/ValueAnalysis.v b/backend/ValueAnalysis.v
index 5072448a..21dd2c35 100644
--- a/backend/ValueAnalysis.v
+++ b/backend/ValueAnalysis.v
@@ -139,9 +139,11 @@ Definition transfer (f: function) (rm: romem) (pc: node) (ae: aenv) (am: amem) :
| Some(Iop op args res s) =>
let a := eval_static_operation op (aregs ae args) in
VA.State (AE.set res a ae) am
- | Some(Iload trap chunk addr args dst s) =>
+ | Some(Iload TRAP chunk addr args dst s) =>
let a := loadv chunk rm am (eval_static_addressing addr (aregs ae args)) in
VA.State (AE.set dst a ae) am
+ | Some(Iload NOTRAP chunk addr args dst s) =>
+ VA.State (AE.set dst Vtop ae) am
| Some(Istore chunk addr args src s) =>
let am' := storev chunk am (eval_static_addressing addr (aregs ae args)) (areg ae src) in
VA.State ae am'
@@ -1268,11 +1270,29 @@ Proof.
apply ematch_update; auto. eapply eval_static_operation_sound; eauto with va.
- (* load *)
+ destruct trap.
+ + eapply sound_succ_state; eauto. simpl; auto.
+ unfold transfer; rewrite H. eauto.
+ apply ematch_update; auto. eapply loadv_sound; eauto with va.
+ eapply eval_static_addressing_sound; eauto with va.
+ + eapply sound_succ_state; eauto. simpl; auto.
+ unfold transfer; rewrite H. eauto.
+ apply ematch_update; auto.
+ eapply vmatch_top.
+ eapply loadv_sound; try eassumption.
+ eapply eval_static_addressing_sound; eauto with va.
+- (* load notrap1 *)
eapply sound_succ_state; eauto. simpl; auto.
unfold transfer; rewrite H. eauto.
- apply ematch_update; auto. eapply loadv_sound; eauto with va.
- eapply eval_static_addressing_sound; eauto with va.
-
+ apply ematch_update; auto.
+ unfold default_notrap_load_value.
+ constructor.
+- (* load notrap2 *)
+ eapply sound_succ_state; eauto. simpl; auto.
+ unfold transfer; rewrite H. eauto.
+ apply ematch_update; auto.
+ unfold default_notrap_load_value.
+ constructor.
- (* store *)
exploit eval_static_addressing_sound; eauto with va. intros VMADDR.
eapply sound_succ_state; eauto. simpl; auto.
diff --git a/mppa_k1c/ValueAOp.v b/mppa_k1c/ValueAOp.v
index 2c9bdf3e..5e9eb455 100644
--- a/mppa_k1c/ValueAOp.v
+++ b/mppa_k1c/ValueAOp.v
@@ -472,6 +472,26 @@ Proof.
rewrite Ptrofs.add_zero_l; eauto with va.
Qed.
+(* not needed
+Theorem eval_static_addressing_sound_none:
+ forall addr vargs aargs,
+ eval_addressing ge (Vptr sp Ptrofs.zero) addr vargs = None ->
+ list_forall2 (vmatch bc) vargs aargs ->
+ (eval_static_addressing addr aargs) = Vbot.
+Proof.
+ unfold eval_addressing, eval_static_addressing.
+ intros until aargs. intros Heval_none Hlist.
+ inv Hlist.
+ destruct addr; trivial; discriminate.
+ inv H0.
+ destruct addr; trivial; discriminate.
+ inv H2.
+ destruct addr; trivial; discriminate.
+ inv H3;
+ destruct addr; trivial; discriminate.
+Qed.
+ *)
+
Theorem eval_static_operation_sound:
forall op vargs m vres aargs,
eval_operation ge (Vptr sp Ptrofs.zero) op vargs m = Some vres ->