aboutsummaryrefslogtreecommitdiffstats
path: root/backend
diff options
context:
space:
mode:
authorDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2019-12-09 14:51:52 +0100
committerDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2019-12-09 14:51:52 +0100
commit37de1399449067121a8bb9a51a7cc7a043ad17e2 (patch)
tree11429171ae1f49b57f8d7f98a09d5ea0511680c0 /backend
parentdd1c78c840bf5eb1ad7e101f2da05578245c5998 (diff)
parente2c64b54bf5df0927c684a70167378c91cba0ff4 (diff)
downloadcompcert-kvx-37de1399449067121a8bb9a51a7cc7a043ad17e2.tar.gz
compcert-kvx-37de1399449067121a8bb9a51a7cc7a043ad17e2.zip
Merge remote-tracking branch 'origin/mppa-work' into mppa-work-upstream-merge
Diffstat (limited to 'backend')
-rw-r--r--backend/Allnontrap.v26
-rw-r--r--backend/Allnontrapproof.v215
-rw-r--r--backend/Allocation.v18
-rw-r--r--backend/Allocproof.v115
-rw-r--r--backend/Bounds.v6
-rw-r--r--backend/CSE.v10
-rw-r--r--backend/CSEdomain.v13
-rw-r--r--backend/CSEproof.v170
-rw-r--r--backend/CleanupLabelsproof.v12
-rw-r--r--backend/Constprop.v4
-rw-r--r--backend/Constpropproof.v55
-rw-r--r--backend/Deadcode.v4
-rw-r--r--backend/Deadcodeproof.v77
-rw-r--r--backend/Debugvar.v2
-rw-r--r--backend/Debugvarproof.v16
-rw-r--r--backend/Duplicate.v24
-rw-r--r--backend/Duplicateproof.v31
-rw-r--r--backend/Inlining.v4
-rw-r--r--backend/Inliningaux.ml2
-rw-r--r--backend/Inliningproof.v54
-rw-r--r--backend/Inliningspec.v6
-rw-r--r--backend/LTL.v19
-rw-r--r--backend/Linear.v23
-rw-r--r--backend/Linearize.v4
-rw-r--r--backend/Linearizeproof.v48
-rw-r--r--backend/Lineartyping.v23
-rw-r--r--backend/Liveness.v2
-rw-r--r--backend/Mach.v19
-rw-r--r--backend/PrintLTL.ml7
-rw-r--r--backend/PrintMach.ml5
-rw-r--r--backend/PrintRTL.ml7
-rw-r--r--backend/PrintXTL.ml7
-rw-r--r--backend/RTL.v36
-rw-r--r--backend/RTLgen.v2
-rw-r--r--backend/RTLgenspec.v2
-rw-r--r--backend/RTLtyping.v26
-rw-r--r--backend/Regalloc.ml26
-rw-r--r--backend/Renumber.v2
-rw-r--r--backend/Renumberproof.v12
-rw-r--r--backend/Splitting.ml4
-rw-r--r--backend/Stacking.v4
-rw-r--r--backend/Stackingproof.v40
-rw-r--r--backend/Tailcallproof.v37
-rw-r--r--backend/Tunnelingproof.v25
-rw-r--r--backend/Unusedglob.v2
-rw-r--r--backend/Unusedglobproof.v32
-rw-r--r--backend/ValueAnalysis.v31
-rw-r--r--backend/XTL.ml4
-rw-r--r--backend/XTL.mli2
49 files changed, 1158 insertions, 157 deletions
diff --git a/backend/Allnontrap.v b/backend/Allnontrap.v
new file mode 100644
index 00000000..acf03eca
--- /dev/null
+++ b/backend/Allnontrap.v
@@ -0,0 +1,26 @@
+Require Import Coqlib Maps Errors Integers Floats Lattice Kildall.
+Require Import AST Linking.
+Require Import Memory Registers Op RTL.
+
+
+Definition transf_ros (ros: reg + ident) : reg + ident := ros.
+
+Definition transf_instr (pc: node) (instr: instruction) :=
+ match instr with
+ | Iload trap chunk addr args dst s => Iload NOTRAP chunk addr args dst s
+ | _ => instr
+ end.
+
+Definition transf_function (f: function) : function :=
+ {| fn_sig := f.(fn_sig);
+ fn_params := f.(fn_params);
+ fn_stacksize := f.(fn_stacksize);
+ fn_code := PTree.map transf_instr f.(fn_code);
+ fn_entrypoint := f.(fn_entrypoint) |}.
+
+Definition transf_fundef (fd: fundef) : fundef :=
+ AST.transf_fundef transf_function fd.
+
+Definition transf_program (p: program) : program :=
+ transform_program transf_fundef p.
+
diff --git a/backend/Allnontrapproof.v b/backend/Allnontrapproof.v
new file mode 100644
index 00000000..92e5a88c
--- /dev/null
+++ b/backend/Allnontrapproof.v
@@ -0,0 +1,215 @@
+Require Import FunInd.
+Require Import Coqlib Maps Errors Integers Floats Lattice Kildall.
+Require Import AST Linking.
+Require Import Values Memory Globalenvs Events Smallstep.
+Require Import Registers Op RTL.
+Require Import Allnontrap.
+
+
+Definition match_prog (p tp: RTL.program) :=
+ match_program (fun ctx f tf => tf = transf_fundef f) eq p tp.
+
+Lemma transf_program_match:
+ forall p, match_prog p (transf_program p).
+Proof.
+ intros. eapply match_transform_program; eauto.
+Qed.
+
+Section PRESERVATION.
+
+Variables prog tprog: program.
+Hypothesis TRANSL: match_prog prog tprog.
+Let ge := Genv.globalenv prog.
+Let tge := Genv.globalenv tprog.
+
+Lemma functions_translated:
+ forall v f,
+ Genv.find_funct ge v = Some f ->
+ Genv.find_funct tge v = Some (transf_fundef f).
+Proof (Genv.find_funct_transf TRANSL).
+
+Lemma function_ptr_translated:
+ forall v f,
+ Genv.find_funct_ptr ge v = Some f ->
+ Genv.find_funct_ptr tge v = Some (transf_fundef f).
+Proof (Genv.find_funct_ptr_transf TRANSL).
+
+Lemma symbols_preserved:
+ forall id,
+ Genv.find_symbol tge id = Genv.find_symbol ge id.
+Proof (Genv.find_symbol_transf TRANSL).
+
+Lemma senv_preserved:
+ Senv.equiv ge tge.
+Proof (Genv.senv_transf TRANSL).
+
+Lemma sig_preserved:
+ forall f, funsig (transf_fundef f) = funsig f.
+Proof.
+ destruct f; reflexivity.
+Qed.
+
+Lemma find_function_translated:
+ forall ros rs fd,
+ find_function ge ros rs = Some fd ->
+ find_function tge ros rs = Some (transf_fundef fd).
+Proof.
+ unfold find_function; intros. destruct ros as [r|id].
+ eapply functions_translated; eauto.
+ rewrite symbols_preserved. destruct (Genv.find_symbol ge id); try congruence.
+ eapply function_ptr_translated; eauto.
+Qed.
+
+Lemma transf_function_at:
+ forall f pc i,
+ f.(fn_code)!pc = Some i ->
+ (transf_function f).(fn_code)!pc = Some(transf_instr pc i).
+Proof.
+ intros until i. intro Hcode.
+ unfold transf_function; simpl.
+ rewrite PTree.gmap.
+ unfold option_map.
+ rewrite Hcode.
+ reflexivity.
+Qed.
+
+Ltac TR_AT :=
+ match goal with
+ | [ A: (fn_code _)!_ = Some _ |- _ ] =>
+ generalize (transf_function_at _ _ _ A); intros
+ end.
+
+
+Inductive match_frames: RTL.stackframe -> RTL.stackframe -> Prop :=
+ | match_frames_intro: forall res f sp pc rs,
+ match_frames (Stackframe res f sp pc rs)
+ (Stackframe res (transf_function f) sp pc rs).
+
+Inductive match_states: RTL.state -> RTL.state -> Prop :=
+ | match_regular_states: forall stk f sp pc rs m stk'
+ (STACKS: list_forall2 match_frames stk stk'),
+ match_states (State stk f sp pc rs m)
+ (State stk' (transf_function f) sp pc rs m)
+ | match_callstates: forall stk f args m stk'
+ (STACKS: list_forall2 match_frames stk stk'),
+ match_states (Callstate stk f args m)
+ (Callstate stk' (transf_fundef f) args m)
+ | match_returnstates: forall stk v m stk'
+ (STACKS: list_forall2 match_frames stk stk'),
+ match_states (Returnstate stk v m)
+ (Returnstate stk' v m).
+
+Lemma step_simulation:
+ forall S1 t S2, RTL.step ge S1 t S2 ->
+ forall S1', match_states S1 S1' ->
+ exists S2', RTL.step tge S1' t S2' /\ match_states S2 S2'.
+Proof.
+ induction 1; intros S1' MS; inv MS; try TR_AT.
+- (* nop *)
+ econstructor; split. eapply exec_Inop; eauto.
+ constructor; auto.
+- (* op *)
+ econstructor; split.
+ eapply exec_Iop with (v := v); eauto.
+ rewrite <- H0. apply eval_operation_preserved. exact symbols_preserved.
+ constructor; auto.
+(* load *)
+- econstructor; split.
+ assert (eval_addressing tge sp addr rs ## args = Some a).
+ rewrite <- H0. apply eval_addressing_preserved. exact symbols_preserved.
+ eapply exec_Iload; eauto.
+ constructor; auto.
+- (* load notrap1 *)
+ econstructor; split.
+ assert (eval_addressing tge sp addr rs ## args = None).
+ rewrite <- H0. apply eval_addressing_preserved. exact symbols_preserved.
+ eapply exec_Iload_notrap1; eauto.
+ constructor; auto.
+- (* load notrap2 *)
+ econstructor; split.
+ assert (eval_addressing tge sp addr rs ## args = Some a).
+ rewrite <- H0. apply eval_addressing_preserved. exact symbols_preserved.
+ eapply exec_Iload_notrap2; eauto.
+ constructor; auto.
+- (* store *)
+ econstructor; split.
+ assert (eval_addressing tge sp addr rs ## args = Some a).
+ rewrite <- H0. apply eval_addressing_preserved. exact symbols_preserved.
+ eapply exec_Istore; eauto.
+ constructor; auto.
+(* call *)
+- econstructor; split.
+ eapply exec_Icall with (fd := transf_fundef fd); eauto.
+ eapply find_function_translated; eauto.
+ apply sig_preserved.
+ constructor. constructor; auto. constructor.
+(* tailcall *)
+- econstructor; split.
+ eapply exec_Itailcall with (fd := transf_fundef fd); eauto.
+ eapply find_function_translated; eauto.
+ apply sig_preserved.
+ constructor. auto.
+(* builtin *)
+- econstructor; split.
+ eapply exec_Ibuiltin; eauto.
+ eapply eval_builtin_args_preserved with (ge1 := ge); eauto. exact symbols_preserved.
+ eapply external_call_symbols_preserved; eauto. apply senv_preserved.
+ constructor; auto.
+(* cond *)
+- econstructor; split.
+ eapply exec_Icond; eauto.
+ constructor; auto.
+(* jumptbl *)
+- econstructor; split.
+ eapply exec_Ijumptable; eauto.
+ constructor; auto.
+(* return *)
+- econstructor; split.
+ eapply exec_Ireturn; eauto.
+ constructor; auto.
+(* internal function *)
+- simpl. econstructor; split.
+ eapply exec_function_internal; eauto.
+ constructor; auto.
+(* external function *)
+- econstructor; split.
+ eapply exec_function_external; eauto.
+ eapply external_call_symbols_preserved; eauto. apply senv_preserved.
+ constructor; auto.
+(* return *)
+- inv STACKS. inv H1.
+ econstructor; split.
+ eapply exec_return; eauto.
+ constructor; auto.
+Qed.
+
+Lemma transf_initial_states:
+ forall S1, RTL.initial_state prog S1 ->
+ exists S2, RTL.initial_state tprog S2 /\ match_states S1 S2.
+Proof.
+ intros. inv H. econstructor; split.
+ econstructor.
+ eapply (Genv.init_mem_transf TRANSL); eauto.
+ rewrite symbols_preserved. rewrite (match_program_main TRANSL). eauto.
+ eapply function_ptr_translated; eauto.
+ rewrite <- H3; apply sig_preserved.
+ constructor. constructor.
+Qed.
+
+Lemma transf_final_states:
+ forall S1 S2 r, match_states S1 S2 -> RTL.final_state S1 r -> RTL.final_state S2 r.
+Proof.
+ intros. inv H0. inv H. inv STACKS. constructor.
+Qed.
+
+Theorem transf_program_correct:
+ forward_simulation (RTL.semantics prog) (RTL.semantics tprog).
+Proof.
+ eapply forward_simulation_step.
+ apply senv_preserved.
+ eexact transf_initial_states.
+ eexact transf_final_states.
+ exact step_simulation.
+Qed.
+
+End PRESERVATION.
diff --git a/backend/Allocation.v b/backend/Allocation.v
index 13e14530..6e4fcc82 100644
--- a/backend/Allocation.v
+++ b/backend/Allocation.v
@@ -58,7 +58,7 @@ Inductive block_shape: Type :=
(mv2: moves) (s: node)
| BSopdead (op: operation) (args: list reg) (res: reg)
(mv: moves) (s: node)
- | BSload (chunk: memory_chunk) (addr: addressing) (args: list reg) (dst: reg)
+ | BSload (trap : trapping_mode) (chunk: memory_chunk) (addr: addressing) (args: list reg) (dst: reg)
(mv1: moves) (args': list mreg) (dst': mreg)
(mv2: moves) (s: node)
| BSloaddead (chunk: memory_chunk) (addr: addressing) (args: list reg) (dst: reg)
@@ -226,15 +226,19 @@ Definition pair_instr_block
| operation_other _ _ =>
pair_Iop_block op args res s b
end
- | Iload chunk addr args dst s =>
+ | Iload trap chunk addr args dst s =>
let (mv1, b1) := extract_moves nil b in
match b1 with
- | Lload chunk' addr' args' dst' :: b2 =>
+ | Lload trap' chunk' addr' args' dst' :: b2 =>
+ assertion (trapping_mode_eq trap' trap);
if chunk_eq chunk Mint64 && Archi.splitlong then
+ (* TODO: do not support non trapping split loads *)
+ assertion (trapping_mode_eq trap TRAP);
assertion (chunk_eq chunk' Mint32);
let (mv2, b3) := extract_moves nil b2 in
match b3 with
- | Lload chunk'' addr'' args'' dst'' :: b4 =>
+ | Lload trap'' chunk'' addr'' args'' dst'' :: b4 =>
+ assertion (trapping_mode_eq trap'' TRAP);
let (mv3, b5) := extract_moves nil b4 in
assertion (chunk_eq chunk'' Mint32);
assertion (eq_addressing addr addr');
@@ -254,7 +258,7 @@ Definition pair_instr_block
assertion (chunk_eq chunk chunk');
assertion (eq_addressing addr addr');
assertion (check_succ s b3);
- Some(BSload chunk addr args dst mv1 args' dst' mv2 s))
+ Some(BSload trap chunk addr args dst mv1 args' dst' mv2 s))
| _ =>
assertion (check_succ s b1);
Some(BSloaddead chunk addr args dst mv1 s)
@@ -1023,7 +1027,7 @@ Definition transfer_aux (f: RTL.function) (env: regenv)
| BSopdead op args res mv s =>
assertion (reg_unconstrained res e);
track_moves env mv e
- | BSload chunk addr args dst mv1 args' dst' mv2 s =>
+ | BSload trap chunk addr args dst mv1 args' dst' mv2 s =>
do e1 <- track_moves env mv2 e;
do e2 <- transfer_use_def args dst args' dst' (destroyed_by_load chunk addr) e1;
track_moves env mv1 e2
@@ -1263,7 +1267,7 @@ Definition successors_block_shape (bsh: block_shape) : list node :=
| BShighlong src dst mv s => s :: nil
| BSop op args res mv1 args' res' mv2 s => s :: nil
| BSopdead op args res mv s => s :: nil
- | BSload chunk addr args dst mv1 args' dst' mv2 s => s :: nil
+ | BSload trap chunk addr args dst mv1 args' dst' mv2 s => s :: nil
| BSload2 addr addr' args dst mv1 args1' dst1' mv2 args2' dst2' mv3 s => s :: nil
| BSload2_1 addr args dst mv1 args' dst' mv2 s => s :: nil
| BSload2_2 addr addr' args dst mv1 args' dst' mv2 s => s :: nil
diff --git a/backend/Allocproof.v b/backend/Allocproof.v
index 1804f46b..3d8fb451 100644
--- a/backend/Allocproof.v
+++ b/backend/Allocproof.v
@@ -96,44 +96,44 @@ Inductive expand_block_shape: block_shape -> RTL.instruction -> LTL.bblock -> Pr
expand_block_shape (BSopdead op args res mv s)
(Iop op args res s)
(expand_moves mv (Lbranch s :: k))
- | ebs_load: forall chunk addr args dst mv1 args' dst' mv2 s k,
+ | ebs_load: forall trap chunk addr args dst mv1 args' dst' mv2 s k,
wf_moves mv1 -> wf_moves mv2 ->
- expand_block_shape (BSload chunk addr args dst mv1 args' dst' mv2 s)
- (Iload chunk addr args dst s)
+ expand_block_shape (BSload trap chunk addr args dst mv1 args' dst' mv2 s)
+ (Iload trap chunk addr args dst s)
(expand_moves mv1
- (Lload chunk addr args' dst' :: expand_moves mv2 (Lbranch s :: k)))
+ (Lload trap chunk addr args' dst' :: expand_moves mv2 (Lbranch s :: k)))
| ebs_load2: forall addr addr2 args dst mv1 args1' dst1' mv2 args2' dst2' mv3 s k,
wf_moves mv1 -> wf_moves mv2 -> wf_moves mv3 ->
Archi.splitlong = true ->
offset_addressing addr 4 = Some addr2 ->
expand_block_shape (BSload2 addr addr2 args dst mv1 args1' dst1' mv2 args2' dst2' mv3 s)
- (Iload Mint64 addr args dst s)
+ (Iload TRAP Mint64 addr args dst s)
(expand_moves mv1
- (Lload Mint32 addr args1' dst1' ::
+ (Lload TRAP Mint32 addr args1' dst1' ::
expand_moves mv2
- (Lload Mint32 addr2 args2' dst2' ::
+ (Lload TRAP Mint32 addr2 args2' dst2' ::
expand_moves mv3 (Lbranch s :: k))))
| ebs_load2_1: forall addr args dst mv1 args' dst' mv2 s k,
wf_moves mv1 -> wf_moves mv2 ->
Archi.splitlong = true ->
expand_block_shape (BSload2_1 addr args dst mv1 args' dst' mv2 s)
- (Iload Mint64 addr args dst s)
+ (Iload TRAP Mint64 addr args dst s)
(expand_moves mv1
- (Lload Mint32 addr args' dst' ::
+ (Lload TRAP Mint32 addr args' dst' ::
expand_moves mv2 (Lbranch s :: k)))
| ebs_load2_2: forall addr addr2 args dst mv1 args' dst' mv2 s k,
wf_moves mv1 -> wf_moves mv2 ->
Archi.splitlong = true ->
offset_addressing addr 4 = Some addr2 ->
expand_block_shape (BSload2_2 addr addr2 args dst mv1 args' dst' mv2 s)
- (Iload Mint64 addr args dst s)
+ (Iload TRAP Mint64 addr args dst s)
(expand_moves mv1
- (Lload Mint32 addr2 args' dst' ::
+ (Lload TRAP Mint32 addr2 args' dst' ::
expand_moves mv2 (Lbranch s :: k)))
- | ebs_load_dead: forall chunk addr args dst mv s k,
+ | ebs_load_dead: forall trap chunk addr args dst mv s k,
wf_moves mv ->
expand_block_shape (BSloaddead chunk addr args dst mv s)
- (Iload chunk addr args dst s)
+ (Iload trap chunk addr args dst s)
(expand_moves mv (Lbranch s :: k))
| ebs_store: forall chunk addr args src mv1 args' src' s k,
wf_moves mv1 ->
@@ -1970,8 +1970,8 @@ Ltac UseShape :=
end.
Remark addressing_not_long:
- forall env f addr args dst s r,
- wt_instr f env (Iload Mint64 addr args dst s) -> Archi.splitlong = true ->
+ forall trap env f addr args dst s r,
+ wt_instr f env (Iload trap Mint64 addr args dst s) -> Archi.splitlong = true ->
In r args -> r <> dst.
Proof.
intros. inv H.
@@ -1981,7 +1981,7 @@ Proof.
{ rewrite <- H5. apply in_map; auto. }
assert (C: env r = Tint).
{ apply A in B. rewrite B. unfold Tptr. rewrite Archi.splitlong_ptr32 by auto. auto. }
- red; intros; subst r. rewrite C in H8; discriminate.
+ red; intros; subst r. rewrite C in H9; discriminate.
Qed.
(** The proof of semantic preservation is a simulation argument of the
@@ -2082,8 +2082,8 @@ Proof.
econstructor; eauto.
eapply wt_exec_Iop; eauto.
-(* load regular *)
-- generalize (wt_exec_Iload _ _ _ _ _ _ _ _ _ _ _ WTI H1 WTRS). intros WTRS'.
+(* load regular TRAP *)
+- generalize (wt_exec_Iload _ _ _ _ _ _ _ _ _ _ _ _ WTI H1 WTRS). intros WTRS'.
exploit (exec_moves mv1); eauto. intros [ls1 [A1 B1]].
exploit transfer_use_def_satisf; eauto. intros [X Y].
exploit eval_addressing_lessdef; eauto. intros [a' [F G]].
@@ -2100,7 +2100,7 @@ Proof.
econstructor; eauto.
(* load pair *)
-- generalize (wt_exec_Iload _ _ _ _ _ _ _ _ _ _ _ WTI H1 WTRS). intros WTRS'.
+- generalize (wt_exec_Iload _ _ _ _ _ _ _ _ _ _ _ _ WTI H1 WTRS). intros WTRS'.
exploit loadv_int64_split; eauto. intros (v1 & v2 & LOAD1 & LOAD2 & V1 & V2).
set (v2' := if Archi.big_endian then v2 else v1) in *.
set (v1' := if Archi.big_endian then v1 else v2) in *.
@@ -2155,7 +2155,7 @@ Proof.
econstructor; eauto.
(* load first word of a pair *)
-- generalize (wt_exec_Iload _ _ _ _ _ _ _ _ _ _ _ WTI H1 WTRS). intros WTRS'.
+- generalize (wt_exec_Iload _ _ _ _ _ _ _ _ _ _ _ _ WTI H1 WTRS). intros WTRS'.
exploit loadv_int64_split; eauto. intros (v1 & v2 & LOAD1 & LOAD2 & V1 & V2).
set (v2' := if Archi.big_endian then v2 else v1) in *.
set (v1' := if Archi.big_endian then v1 else v2) in *.
@@ -2185,7 +2185,7 @@ Proof.
econstructor; eauto.
(* load second word of a pair *)
-- generalize (wt_exec_Iload _ _ _ _ _ _ _ _ _ _ _ WTI H1 WTRS). intros WTRS'.
+- generalize (wt_exec_Iload _ _ _ _ _ _ _ _ _ _ _ _ WTI H1 WTRS). intros WTRS'.
exploit loadv_int64_split; eauto. intros (v1 & v2 & LOAD1 & LOAD2 & V1 & V2).
set (v2' := if Archi.big_endian then v2 else v1) in *.
set (v1' := if Archi.big_endian then v1 else v2) in *.
@@ -2229,6 +2229,79 @@ Proof.
econstructor; eauto.
eapply wt_exec_Iload; eauto.
+- (* load notrap1 *)
+ generalize (wt_exec_Iload_notrap _ _ _ _ _ _ _ _ WTI WTRS).
+ intro WTRS'.
+ exploit (exec_moves mv1); eauto. intros [ls1 [A1 B1]].
+ exploit transfer_use_def_satisf; eauto. intros [X Y].
+ exploit eval_addressing_lessdef_none; eauto. intro Haddr.
+ exploit (exec_moves mv2); eauto. intros [ls2 [A2 B2]].
+ econstructor; split.
+ eapply plus_left. econstructor; eauto.
+ eapply star_trans. eexact A1.
+ eapply star_left. eapply exec_Lload_notrap1. rewrite <- Haddr.
+ apply eval_addressing_preserved. exact symbols_preserved. eauto.
+
+ eapply star_right. eexact A2. constructor.
+ eauto. eauto. eauto. traceEq.
+ exploit satisf_successors; eauto. simpl; eauto. intros [enext [U V]].
+ econstructor; eauto.
+
+(* load notrap1 dead *)
+- exploit exec_moves; eauto. intros [ls1 [X Y]].
+ econstructor; split.
+ eapply plus_left. econstructor; eauto.
+ eapply star_right. eexact X. econstructor; eauto.
+ eauto. traceEq.
+ exploit satisf_successors. eauto. eauto. simpl; eauto. eauto.
+ eapply reg_unconstrained_satisf; eauto.
+ intros [enext [U V]].
+ econstructor; eauto.
+ eapply wt_exec_Iload_notrap; eauto.
+
+(* load regular notrap2 *)
+- generalize (wt_exec_Iload_notrap _ _ _ _ _ _ _ _ WTI WTRS).
+ intro WTRS'.
+ exploit (exec_moves mv1); eauto. intros [ls1 [A1 B1]].
+ exploit transfer_use_def_satisf; eauto. intros [X Y].
+ exploit eval_addressing_lessdef; eauto. intros [a' [F G]].
+ destruct (Mem.loadv chunk m' a') as [v' |] eqn:Hload.
+ { exploit (exec_moves mv2 env (rs # dst <- Vundef)); eauto. intros [ls2 [A2 B2]].
+ econstructor; split.
+ eapply plus_left. econstructor; eauto.
+ eapply star_trans. eexact A1.
+ eapply star_left. econstructor. instantiate (1 := a'). rewrite <- F.
+ apply eval_addressing_preserved. exact symbols_preserved. eauto. eauto.
+ eapply star_right. eexact A2. constructor.
+ eauto. eauto. eauto. traceEq.
+ exploit satisf_successors; eauto. simpl; eauto. intros [enext [U V]].
+ econstructor; eauto.
+ }
+ { exploit (exec_moves mv2 env (rs # dst <- Vundef)); eauto. intros [ls2 [A2 B2]].
+ econstructor; split.
+ eapply plus_left. econstructor; eauto.
+ eapply star_trans. eexact A1.
+ eapply star_left. eapply exec_Lload_notrap2. rewrite <- F.
+ apply eval_addressing_preserved. exact symbols_preserved. assumption.
+ eauto.
+ eapply star_right. eexact A2. constructor.
+ eauto. eauto. eauto. traceEq.
+ exploit satisf_successors; eauto. simpl; eauto. intros [enext [U V]].
+ econstructor; eauto.
+ }
+
+- (* load notrap2 dead *)
+ exploit exec_moves; eauto. intros [ls1 [X Y]].
+ econstructor; split.
+ eapply plus_left. econstructor; eauto.
+ eapply star_right. eexact X. econstructor; eauto.
+ eauto. traceEq.
+ exploit satisf_successors. eauto. eauto. simpl; eauto. eauto.
+ eapply reg_unconstrained_satisf; eauto.
+ intros [enext [U V]].
+ econstructor; eauto.
+ eapply wt_exec_Iload_notrap; eauto.
+
(* store *)
- exploit exec_moves; eauto. intros [ls1 [X Y]].
exploit add_equations_lessdef; eauto. intros LD. simpl in LD. inv LD.
diff --git a/backend/Bounds.v b/backend/Bounds.v
index fa695234..b8c12166 100644
--- a/backend/Bounds.v
+++ b/backend/Bounds.v
@@ -67,7 +67,7 @@ Definition instr_within_bounds (i: instruction) :=
| Lgetstack sl ofs ty r => slot_within_bounds sl ofs ty /\ mreg_within_bounds r
| Lsetstack r sl ofs ty => slot_within_bounds sl ofs ty
| Lop op args res => mreg_within_bounds res
- | Lload chunk addr args dst => mreg_within_bounds dst
+ | Lload trap chunk addr args dst => mreg_within_bounds dst
| Lcall sig ros => size_arguments sig <= bound_outgoing b
| Lbuiltin ef args res =>
(forall r, In r (params_of_builtin_res res) \/ In r (destroyed_by_builtin ef) -> mreg_within_bounds r)
@@ -104,7 +104,7 @@ Definition record_regs_of_instr (u: RegSet.t) (i: instruction) : RegSet.t :=
| Lgetstack sl ofs ty r => record_reg u r
| Lsetstack r sl ofs ty => record_reg u r
| Lop op args res => record_reg u res
- | Lload chunk addr args dst => record_reg u dst
+ | Lload trap chunk addr args dst => record_reg u dst
| Lstore chunk addr args src => u
| Lcall sig ros => u
| Ltailcall sig ros => u
@@ -280,7 +280,7 @@ Definition defined_by_instr (r': mreg) (i: instruction) :=
match i with
| Lgetstack sl ofs ty r => r' = r
| Lop op args res => r' = res
- | Lload chunk addr args dst => r' = dst
+ | Lload trap chunk addr args dst => r' = dst
| Lbuiltin ef args res => In r' (params_of_builtin_res res) \/ In r' (destroyed_by_builtin ef)
| _ => False
end.
diff --git a/backend/CSE.v b/backend/CSE.v
index ecfa1f9e..2827161d 100644
--- a/backend/CSE.v
+++ b/backend/CSE.v
@@ -459,8 +459,10 @@ Definition transfer (f: function) (approx: PMap.t VA.t) (pc: node) (before: numb
before
| Iop op args res s =>
add_op before res op args
- | Iload chunk addr args dst s =>
- add_load before dst chunk addr args
+ | Iload TRAP chunk addr args dst s =>
+ add_load before dst chunk addr args
+ | Iload NOTRAP _ _ _ dst _ =>
+ set_unknown before dst
| Istore chunk addr args src s =>
let app := approx!!pc in
let n := kill_loads_after_store app before chunk addr args in
@@ -534,14 +536,14 @@ Definition transf_instr (n: numbering) (instr: instruction) :=
let (op', args') := reduce _ combine_op n1 op args vl in
Iop op' args' res s
end
- | Iload chunk addr args dst s =>
+ | Iload TRAP chunk addr args dst s =>
let (n1, vl) := valnum_regs n args in
match find_rhs n1 (Load chunk addr vl) with
| Some r =>
Iop Omove (r :: nil) dst s
| None =>
let (addr', args') := reduce _ combine_addr n1 addr args vl in
- Iload chunk addr' args' dst s
+ Iload TRAP chunk addr' args' dst s
end
| Istore chunk addr args src s =>
let (n1, vl) := valnum_regs n args in
diff --git a/backend/CSEdomain.v b/backend/CSEdomain.v
index 9b1243c8..34ec0118 100644
--- a/backend/CSEdomain.v
+++ b/backend/CSEdomain.v
@@ -43,7 +43,7 @@ Definition eq_list_valnum: forall (x y: list valnum), {x=y}+{x<>y} := list_eq_de
Definition eq_rhs (x y: rhs) : {x=y}+{x<>y}.
Proof.
- generalize chunk_eq eq_operation eq_addressing eq_valnum eq_list_valnum.
+ generalize trapping_mode_eq chunk_eq eq_operation eq_addressing eq_valnum eq_list_valnum.
decide equality.
Defined.
@@ -109,7 +109,16 @@ Inductive rhs_eval_to (valu: valuation) (ge: genv) (sp: val) (m: mem):
| load_eval_to: forall chunk addr vl a v,
eval_addressing ge sp addr (map valu vl) = Some a ->
Mem.loadv chunk m a = Some v ->
- rhs_eval_to valu ge sp m (Load chunk addr vl) v.
+ rhs_eval_to valu ge sp m (Load chunk addr vl) v
+(* | load_notrap1_eval_to: forall chunk addr vl,
+ eval_addressing ge sp addr (map valu vl) = None ->
+ rhs_eval_to valu ge sp m (Load NOTRAP chunk addr vl)
+ (default_notrap_load_value chunk)
+ | load_notrap2_eval_to: forall chunk addr vl a,
+ eval_addressing ge sp addr (map valu vl) = Some a ->
+ Mem.loadv chunk m a = None ->
+ rhs_eval_to valu ge sp m (Load NOTRAP chunk addr vl)
+ (default_notrap_load_value chunk) *).
Inductive equation_holds (valu: valuation) (ge: genv) (sp: val) (m: mem):
equation -> Prop :=
diff --git a/backend/CSEproof.v b/backend/CSEproof.v
index 03c7ecfc..5bbb7508 100644
--- a/backend/CSEproof.v
+++ b/backend/CSEproof.v
@@ -71,7 +71,11 @@ Lemma rhs_eval_to_exten:
Proof.
intros. inv H; simpl in *.
- constructor. rewrite valnums_val_exten by assumption. auto.
-- econstructor; eauto. rewrite valnums_val_exten by assumption. auto.
+- eapply load_eval_to; eauto. rewrite valnums_val_exten by assumption. auto.
+(*
+- apply load_notrap1_eval_to; auto. rewrite valnums_val_exten by assumption. assumption.
+- eapply load_notrap2_eval_to; eauto. rewrite valnums_val_exten by assumption. assumption.
+*)
Qed.
Lemma equation_holds_exten:
@@ -393,6 +397,39 @@ Proof.
+ intros. apply Regmap.gso; auto.
Qed.
+(*
+Lemma add_load_holds_none1:
+ forall valu1 ge sp rs m n addr (args: list reg) chunk dst,
+ numbering_holds valu1 ge sp rs m n ->
+ eval_addressing ge sp addr rs##args = None ->
+ exists valu2, numbering_holds valu2 ge sp (rs#dst <- (default_notrap_load_value chunk)) m (add_load n dst chunk addr args).
+Proof.
+ unfold add_load; intros.
+ destruct (valnum_regs n args) as [n1 vl] eqn:VN.
+ exploit valnum_regs_holds; eauto.
+ intros (valu2 & A & B & C & D & E).
+ eapply add_rhs_holds; eauto.
++ rewrite Regmap.gss; auto. eapply load_notrap1_eval_to. rewrite <- B; eauto.
++ intros. apply Regmap.gso; auto.
+Qed.
+
+Lemma add_load_holds_none2:
+ forall valu1 ge sp rs m n addr (args: list reg) a chunk dst,
+ numbering_holds valu1 ge sp rs m n ->
+ eval_addressing ge sp addr rs##args = Some a ->
+ Mem.loadv chunk m a = None ->
+ exists valu2, numbering_holds valu2 ge sp (rs#dst <- (default_notrap_load_value chunk)) m (add_load n dst NOTRAP chunk addr args).
+Proof.
+ unfold add_load; intros.
+ destruct (valnum_regs n args) as [n1 vl] eqn:VN.
+ exploit valnum_regs_holds; eauto.
+ intros (valu2 & A & B & C & D & E).
+ eapply add_rhs_holds; eauto.
++ rewrite Regmap.gss; auto. eapply load_notrap2_eval_to. rewrite <- B; eauto. assumption.
++ intros. apply Regmap.gso; auto.
+Qed.
+ *)
+
Lemma set_unknown_holds:
forall valu ge sp rs m n r v,
numbering_holds valu ge sp rs m n ->
@@ -456,8 +493,8 @@ Lemma kill_all_loads_hold:
Proof.
intros. eapply kill_equations_hold; eauto.
unfold filter_loads; intros. inv H1.
- constructor. rewrite <- H2. apply op_depends_on_memory_correct; auto.
- discriminate.
+ 1: constructor; rewrite <- H2; apply op_depends_on_memory_correct; auto.
+ all: discriminate.
Qed.
Lemma kill_loads_after_store_holds:
@@ -486,6 +523,20 @@ Proof.
apply match_aptr_of_aval. eapply eval_static_addressing_sound; eauto.
erewrite <- regs_valnums_sound by eauto. eauto with va.
apply match_aptr_of_aval. eapply eval_static_addressing_sound; eauto with va.
+(*
+- eapply load_notrap1_eval_to; assumption.
+- destruct (regs_valnums n vl) as [rl|] eqn:RV; try discriminate.
+ eapply load_notrap2_eval_to; eauto.
+ rewrite <- H9.
+ destruct a; simpl in H1; try discriminate.
+ destruct a0; simpl in H9; try discriminate; simpl; trivial.
+ rewrite negb_false_iff in H6. unfold aaddressing in H6.
+ eapply Mem.load_store_other. eauto.
+ eapply pdisjoint_sound; eauto.
+ apply match_aptr_of_aval. eapply eval_static_addressing_sound; eauto.
+ erewrite <- regs_valnums_sound by eauto. eauto with va.
+ apply match_aptr_of_aval. eapply eval_static_addressing_sound; eauto with va.
+*)
Qed.
Lemma store_normalized_range_sound:
@@ -562,6 +613,19 @@ Proof.
unfold aaddressing. apply match_aptr_of_aval. eapply eval_static_addressing_sound; eauto.
erewrite <- regs_valnums_sound by eauto. eauto with va.
auto.
+(*
+- eapply load_notrap1_eval_to; assumption.
+- destruct (regs_valnums n vl) as [rl|] eqn:RV; try discriminate.
+ eapply load_notrap2_eval_to; eauto. rewrite <- H11.
+ destruct a; simpl in H10; try discriminate; simpl; trivial.
+ rewrite negb_false_iff in H8.
+ eapply Mem.load_storebytes_other. eauto.
+ rewrite H6. rewrite Z2Nat.id by omega.
+ eapply pdisjoint_sound. eauto.
+ unfold aaddressing. apply match_aptr_of_aval. eapply eval_static_addressing_sound; eauto.
+ erewrite <- regs_valnums_sound by eauto. eauto with va.
+ auto.
+*)
Qed.
Lemma load_memcpy:
@@ -1034,6 +1098,10 @@ Proof.
destruct (valnum_regs approx!!pc args) as [n1 vl] eqn:?.
destruct SAT as [valu1 NH1].
exploit valnum_regs_holds; eauto. intros (valu2 & NH2 & EQ & AG & P & Q).
+ destruct trap.
+
+ (* TRAP *)
+ {
destruct (find_rhs n1 (Load chunk addr vl)) as [r|] eqn:?.
+ (* replaced by move *)
exploit find_rhs_sound; eauto. intros (v' & EV & LD).
@@ -1063,7 +1131,103 @@ Proof.
unfold transfer; rewrite H.
eapply add_load_holds; eauto.
apply set_reg_lessdef; auto.
+ }
+
+ (* NOTRAP *)
+ {
+ assert (exists a' : val,
+ eval_addressing ge sp addr rs' ## args = Some a' /\ Val.lessdef a a')
+ as Haa'.
+ apply eval_addressing_lessdef with (vl1 := rs ## args).
+ apply regs_lessdef_regs; assumption.
+ assumption.
+ destruct Haa' as [a' [Ha'1 Ha'2]].
+
+ assert (
+ exists v' : val,
+ Mem.loadv chunk m' a' = Some v' /\ Val.lessdef v v') as Hload' by
+ (apply Mem.loadv_extends with (m1 := m) (addr1 := a); assumption).
+ destruct Hload' as [v' [Hv'1 Hv'2]].
+
+ econstructor. split.
+ eapply exec_Iload; eauto.
+ try (rewrite eval_addressing_preserved with (ge1 := ge); auto; exact symbols_preserved).
+
+ econstructor; eauto.
+ eapply analysis_correct_1; eauto. simpl; eauto.
+ unfold transfer. rewrite H.
+ exists valu1.
+ apply set_unknown_holds.
+ assumption.
+ apply set_reg_lessdef; assumption.
+ }
+
+- (* Iload notrap 1*)
+ destruct (valnum_regs approx!!pc args) as [n1 vl] eqn:?.
+ destruct SAT as [valu1 NH1].
+ exploit valnum_regs_holds; eauto. intros (valu2 & NH2 & EQ & AG & P & Q).
+
+ econstructor. split.
+ eapply exec_Iload_notrap1; eauto.
+ rewrite eval_addressing_preserved with (ge1 := ge).
+ apply eval_addressing_lessdef_none with (vl1 := rs ## args).
+ apply regs_lessdef_regs; assumption.
+ assumption.
+ exact symbols_preserved.
+
+ econstructor; eauto.
+ eapply analysis_correct_1; eauto. simpl; eauto.
+ unfold transfer. rewrite H.
+ exists valu1.
+ apply set_unknown_holds.
+ assumption.
+ apply set_reg_lessdef.
+ constructor. assumption.
+
+- (* Iload notrap 2*)
+ destruct (valnum_regs approx!!pc args) as [n1 vl] eqn:?.
+ destruct SAT as [valu1 NH1].
+ exploit valnum_regs_holds; eauto. intros (valu2 & NH2 & EQ & AG & P & Q).
+ assert (exists a' : val,
+ eval_addressing ge sp addr rs' ## args = Some a' /\ Val.lessdef a a')
+ as Haa'.
+ apply eval_addressing_lessdef with (vl1 := rs ## args).
+ apply regs_lessdef_regs; assumption.
+ assumption.
+ destruct Haa' as [a' [Ha'1 Ha'2]].
+
+ destruct (Mem.loadv chunk m' a') eqn:Hload'.
+
+ {
+ econstructor. split.
+ eapply exec_Iload; eauto.
+ try (rewrite eval_addressing_preserved with (ge1 := ge); auto; exact symbols_preserved).
+
+ econstructor; eauto.
+ eapply analysis_correct_1; eauto. simpl; eauto.
+ unfold transfer. rewrite H.
+ exists valu1.
+ apply set_unknown_holds.
+ assumption.
+ unfold default_notrap_load_value.
+ apply set_reg_lessdef; eauto.
+ }
+ {
+ econstructor. split.
+ eapply exec_Iload_notrap2; eauto.
+ try (rewrite eval_addressing_preserved with (ge1 := ge); auto; exact symbols_preserved).
+
+ econstructor; eauto.
+ eapply analysis_correct_1; eauto. simpl; eauto.
+ unfold transfer. rewrite H.
+ exists valu1.
+ apply set_unknown_holds.
+ assumption.
+ apply set_reg_lessdef.
+ constructor. assumption.
+ }
+
- (* Istore *)
destruct (valnum_regs approx!!pc args) as [n1 vl] eqn:?.
destruct SAT as [valu1 NH1].
diff --git a/backend/CleanupLabelsproof.v b/backend/CleanupLabelsproof.v
index e92be2b4..84ca403e 100644
--- a/backend/CleanupLabelsproof.v
+++ b/backend/CleanupLabelsproof.v
@@ -255,6 +255,18 @@ Proof.
left; econstructor; split.
econstructor; eauto.
econstructor; eauto with coqlib.
+(* Lload notrap1 *)
+ assert (eval_addressing tge sp addr (LTL.reglist rs args) = None).
+ rewrite <- H. apply eval_addressing_preserved. exact symbols_preserved.
+ left; econstructor; split.
+ eapply exec_Lload_notrap1; eauto.
+ econstructor; eauto with coqlib.
+(* Lload notrap2 *)
+ assert (eval_addressing tge sp addr (LTL.reglist rs args) = Some a).
+ rewrite <- H. apply eval_addressing_preserved. exact symbols_preserved.
+ left; econstructor; split.
+ eapply exec_Lload_notrap2; eauto.
+ econstructor; eauto with coqlib.
(* Lstore *)
assert (eval_addressing tge sp addr (LTL.reglist rs args) = Some a).
rewrite <- H. apply eval_addressing_preserved. exact symbols_preserved.
diff --git a/backend/Constprop.v b/backend/Constprop.v
index 4aab7677..eda41b39 100644
--- a/backend/Constprop.v
+++ b/backend/Constprop.v
@@ -181,7 +181,7 @@ Definition transf_instr (f: function) (an: PMap.t VA.t) (rm: romem)
let (op', args') := op_strength_reduction op args aargs in
Iop op' args' res s'
end
- | Iload chunk addr args dst s =>
+ | Iload TRAP chunk addr args dst s =>
let aargs := aregs ae args in
let a := ValueDomain.loadv chunk rm am (eval_static_addressing addr aargs) in
match const_for_result a with
@@ -189,7 +189,7 @@ Definition transf_instr (f: function) (an: PMap.t VA.t) (rm: romem)
Iop cop nil dst s
| None =>
let (addr', args') := addr_strength_reduction addr args aargs in
- Iload chunk addr' args' dst s
+ Iload TRAP chunk addr' args' dst s
end
| Istore chunk addr args src s =>
let aargs := aregs ae args in
diff --git a/backend/Constpropproof.v b/backend/Constpropproof.v
index a5d08a0f..63cfee24 100644
--- a/backend/Constpropproof.v
+++ b/backend/Constpropproof.v
@@ -406,6 +406,8 @@ Proof.
assert (VM1: vmatch bc a aa) by (eapply eval_static_addressing_sound; eauto with va).
set (av := loadv chunk (romem_for cu) am aa).
assert (VM2: vmatch bc v av) by (eapply loadv_sound; eauto).
+ destruct trap.
+ {
destruct (const_for_result av) as [cop|] eqn:?; intros.
+ (* constant-propagated *)
exploit const_for_result_correct; eauto. intros (v' & A & B).
@@ -431,6 +433,59 @@ Proof.
left; econstructor; econstructor; split.
eapply exec_Iload; eauto.
eapply match_states_succ; eauto. apply set_reg_lessdef; auto.
+ }
+ {
+ assert (exists v2 : val,
+ eval_addressing ge (Vptr sp0 Ptrofs.zero) addr (rs' ## args) = Some v2 /\ Val.lessdef a v2) as Hexist2.
+ apply eval_addressing_lessdef with (vl1 := rs ## args).
+ apply regs_lessdef_regs; assumption.
+ assumption.
+ destruct Hexist2 as [v2 [Heval2 Hlessdef2]].
+ destruct (Mem.loadv_extends chunk m m' a v2 v MEM H1 Hlessdef2) as [vX [Hvx1 Hvx2]].
+ left; econstructor; econstructor; split.
+ eapply exec_Iload with (a := v2); eauto.
+ try (erewrite eval_addressing_preserved with (ge1:=ge); auto;
+ exact symbols_preserved).
+ eapply match_states_succ; eauto. apply set_reg_lessdef; auto.
+
+ }
+
+- (* Iload notrap1 *)
+ rename pc'0 into pc. TransfInstr.
+ assert (eval_addressing tge (Vptr sp0 Ptrofs.zero) addr (rs' ## args) = None).
+ rewrite eval_addressing_preserved with (ge1 := ge); eauto.
+ apply eval_addressing_lessdef_none with (vl1 := rs ## args).
+ apply regs_lessdef_regs; assumption.
+ assumption.
+ exact symbols_preserved.
+
+ left; econstructor; econstructor; split.
+ eapply exec_Iload_notrap1; eauto.
+ eapply match_states_succ; eauto. apply set_reg_lessdef; auto.
+
+- (* Iload notrap2 *)
+ rename pc'0 into pc. TransfInstr.
+ assert (exists v2 : val,
+ eval_addressing ge (Vptr sp0 Ptrofs.zero) addr (rs' ## args) = Some v2 /\ Val.lessdef a v2) as Hexist2.
+ apply eval_addressing_lessdef with (vl1 := rs ## args).
+ apply regs_lessdef_regs; assumption.
+ assumption.
+ destruct Hexist2 as [a' [Heval' Hlessdef']].
+ destruct (Mem.loadv chunk m' a') eqn:Hload'.
+ {
+ left; econstructor; econstructor; split.
+ eapply exec_Iload; eauto.
+
+ try (rewrite eval_addressing_preserved with (ge1 := ge); auto; exact symbols_preserved).
+ eapply match_states_succ; eauto. apply set_reg_lessdef; auto.
+ }
+ {
+ left; econstructor; econstructor; split.
+ eapply exec_Iload_notrap2; eauto.
+
+ try (rewrite eval_addressing_preserved with (ge1 := ge); auto; exact symbols_preserved).
+ eapply match_states_succ; eauto. apply set_reg_lessdef; auto.
+ }
- (* Istore *)
rename pc'0 into pc. TransfInstr.
diff --git a/backend/Deadcode.v b/backend/Deadcode.v
index 2286876e..1f208a91 100644
--- a/backend/Deadcode.v
+++ b/backend/Deadcode.v
@@ -123,7 +123,7 @@ Definition transfer (f: function) (approx: PMap.t VA.t)
if is_dead nres then after
else if is_int_zero nres then (kill res ne, nm)
else (add_needs args (needs_of_operation op nres) (kill res ne), nm)
- | Some (Iload chunk addr args dst s) =>
+ | Some (Iload trap chunk addr args dst s) =>
let ndst := nreg ne dst in
if is_dead ndst then after
else if is_int_zero ndst then (kill dst ne, nm)
@@ -175,7 +175,7 @@ Definition transf_instr (approx: PMap.t VA.t) (an: PMap.t NA.t)
end
else
instr
- | Iload chunk addr args dst s =>
+ | Iload trap chunk addr args dst s =>
let ndst := nreg (fst an!!pc) dst in
if is_dead ndst then
Inop s
diff --git a/backend/Deadcodeproof.v b/backend/Deadcodeproof.v
index 2edc0395..6919fe78 100644
--- a/backend/Deadcodeproof.v
+++ b/backend/Deadcodeproof.v
@@ -829,6 +829,83 @@ Ltac UseTransfer :=
apply eagree_update; eauto 2 with na.
eapply magree_monotone; eauto. intros. apply incl_nmem_add; auto.
+- (* load notrap1 *)
+ TransfInstr; UseTransfer.
+ destruct (is_dead (nreg ne dst)) eqn:DEAD;
+ [idtac|destruct (is_int_zero (nreg ne dst)) eqn:INTZERO];
+ simpl in *.
++ (* dead instruction, turned into a nop *)
+ econstructor; split.
+ eapply exec_Inop; eauto.
+ eapply match_succ_states; eauto. simpl; auto.
+ apply eagree_update_dead; auto with na.
++ (* instruction with needs = [I Int.zero], turned into a load immediate of zero. *)
+ econstructor; split.
+ eapply exec_Iop with (v := Vint Int.zero); eauto.
+ eapply match_succ_states; eauto. simpl; auto.
+ apply eagree_update; auto.
+ rewrite is_int_zero_sound by auto.
+ unfold default_notrap_load_value.
+ constructor.
++ (* preserved *)
+ exploit eval_addressing_lessdef_none. eapply add_needs_all_lessdef; eauto. eassumption.
+ intro Hnone'.
+ assert (eval_addressing tge (Vptr sp0 Ptrofs.zero) addr te ## args = None) as Hnone2'.
+ erewrite eval_addressing_preserved with (ge1 := ge).
+ assumption.
+ exact symbols_preserved.
+
+ econstructor; split.
+ eapply exec_Iload_notrap1; eauto.
+ eapply match_succ_states; eauto. simpl; auto.
+ apply eagree_update; eauto 2 with na.
+ eapply magree_monotone; eauto. intros. apply incl_nmem_add; auto.
+
+- (* load notrap2 *)
+ TransfInstr; UseTransfer.
+
+ destruct (is_dead (nreg ne dst)) eqn:DEAD;
+ [idtac|destruct (is_int_zero (nreg ne dst)) eqn:INTZERO];
+ simpl in *.
++ (* dead instruction, turned into a nop *)
+ econstructor; split.
+ eapply exec_Inop; eauto.
+ eapply match_succ_states; eauto. simpl; auto.
+ apply eagree_update_dead; auto with na.
++ (* instruction with needs = [I Int.zero], turned into a load immediate of zero. *)
+ econstructor; split.
+ eapply exec_Iop with (v := Vint Int.zero); eauto.
+ eapply match_succ_states; eauto. simpl; auto.
+ apply eagree_update; auto.
+ rewrite is_int_zero_sound by auto.
+ unfold default_notrap_load_value.
+ constructor.
++ (* preserved *)
+ exploit eval_addressing_lessdef. eapply add_needs_all_lessdef; eauto. eauto.
+ intros (ta & U & V).
+ destruct (Mem.loadv chunk tm ta) eqn:Hchunk2.
+ {
+ econstructor; split.
+ eapply exec_Iload. eauto.
+ erewrite eval_addressing_preserved with (ge1 := ge).
+ eassumption.
+ exact symbols_preserved.
+ eassumption.
+ eapply match_succ_states; eauto. simpl; auto.
+ apply eagree_update; eauto 2 with na.
+ eapply magree_monotone; eauto. intros. apply incl_nmem_add; auto.
+ }
+ {
+ econstructor; split.
+ eapply exec_Iload_notrap2. eauto.
+ erewrite eval_addressing_preserved with (ge1 := ge).
+ eassumption.
+ exact symbols_preserved.
+ eassumption.
+ eapply match_succ_states; eauto. simpl; auto.
+ apply eagree_update; eauto 2 with na.
+ eapply magree_monotone; eauto. intros. apply incl_nmem_add; auto.
+ }
- (* store *)
TransfInstr; UseTransfer.
destruct (nmem_contains nm (aaddressing (vanalyze cu f) # pc addr args)
diff --git a/backend/Debugvar.v b/backend/Debugvar.v
index 1f361030..56908855 100644
--- a/backend/Debugvar.v
+++ b/backend/Debugvar.v
@@ -233,7 +233,7 @@ Definition transfer (lm: labelmap) (before: option avail) (i: instruction):
(lm, Some (kill (S sl ofs ty) s))
| Lop op args dst =>
(lm, Some (kill (R dst) s))
- | Lload chunk addr args dst =>
+ | Lload trap chunk addr args dst =>
(lm, Some (kill (R dst) s))
| Lstore chunk addr args src =>
(lm, before)
diff --git a/backend/Debugvarproof.v b/backend/Debugvarproof.v
index d31c63ec..95020637 100644
--- a/backend/Debugvarproof.v
+++ b/backend/Debugvarproof.v
@@ -449,6 +449,22 @@ Proof.
eauto. eauto.
apply eval_add_delta_ranges. traceEq.
constructor; auto.
+- (* load notrap1 *)
+ econstructor; split.
+ eapply plus_left.
+ eapply exec_Lload_notrap1.
+ rewrite <- H; apply eval_addressing_preserved; exact symbols_preserved.
+ eauto. eauto.
+ apply eval_add_delta_ranges. traceEq.
+ constructor; auto.
+- (* load notrap2 *)
+ econstructor; split.
+ eapply plus_left.
+ eapply exec_Lload_notrap2.
+ rewrite <- H; apply eval_addressing_preserved; exact symbols_preserved.
+ eauto. eauto.
+ apply eval_add_delta_ranges. traceEq.
+ constructor; auto.
- (* store *)
econstructor; split.
eapply plus_left.
diff --git a/backend/Duplicate.v b/backend/Duplicate.v
index 3ad37c83..46f0855d 100644
--- a/backend/Duplicate.v
+++ b/backend/Duplicate.v
@@ -72,17 +72,19 @@ Definition verify_match_inst dupmap inst tinst :=
else Error(msg "Different operations in Iop")
| _ => Error(msg "verify_match_inst Inop") end
- | Iload m a lr r n => match tinst with
- | Iload m' a' lr' r' n' =>
+ | Iload tm m a lr r n => match tinst with
+ | Iload tm' m' a' lr' r' n' =>
do u <- verify_is_copy dupmap n n';
- if (chunk_eq m m') then
- if (eq_addressing a a') then
- if (list_eq_dec Pos.eq_dec lr lr') then
- if (Pos.eq_dec r r') then OK tt
- else Error (msg "Different r in Iload")
- else Error (msg "Different lr in Iload")
- else Error (msg "Different addressing in Iload")
- else Error (msg "Different mchunk in Iload")
+ if (trapping_mode_eq tm tm') then
+ if (chunk_eq m m') then
+ if (eq_addressing a a') then
+ if (list_eq_dec Pos.eq_dec lr lr') then
+ if (Pos.eq_dec r r') then OK tt
+ else Error (msg "Different r in Iload")
+ else Error (msg "Different lr in Iload")
+ else Error (msg "Different addressing in Iload")
+ else Error (msg "Different mchunk in Iload")
+ else Error (msg "Different trapping_mode in Iload")
| _ => Error (msg "verify_match_inst Iload") end
| Istore m a lr r n => match tinst with
@@ -195,4 +197,4 @@ Definition transf_fundef (f: fundef) : res fundef :=
transf_partial_fundef transf_function f.
Definition transf_program (p: program) : res program :=
- transform_partial_program transf_fundef p. \ No newline at end of file
+ transform_partial_program transf_fundef p.
diff --git a/backend/Duplicateproof.v b/backend/Duplicateproof.v
index 9d56e86f..ebb17774 100644
--- a/backend/Duplicateproof.v
+++ b/backend/Duplicateproof.v
@@ -13,8 +13,8 @@ Inductive match_inst (dupmap: PTree.t node): instruction -> instruction -> Prop
dupmap!n' = (Some n) -> match_inst dupmap (Inop n) (Inop n')
| match_inst_op: forall n n' op lr r,
dupmap!n' = (Some n) -> match_inst dupmap (Iop op lr r n) (Iop op lr r n')
- | match_inst_load: forall n n' m a lr r,
- dupmap!n' = (Some n) -> match_inst dupmap (Iload m a lr r n) (Iload m a lr r n')
+ | match_inst_load: forall n n' tm m a lr r,
+ dupmap!n' = (Some n) -> match_inst dupmap (Iload tm m a lr r n) (Iload tm m a lr r n')
| match_inst_store: forall n n' m a lr r,
dupmap!n' = (Some n) -> match_inst dupmap (Istore m a lr r n) (Istore m a lr r n')
| match_inst_call: forall n n' s ri lr r,
@@ -137,6 +137,7 @@ Proof.
(* Iload *)
- destruct i'; try (inversion H; fail). monadInv H.
destruct x. eapply verify_is_copy_correct in EQ.
+ destruct (trapping_mode_eq _ _); try discriminate.
destruct (chunk_eq _ _); try discriminate.
destruct (eq_addressing _ _); try discriminate.
destruct (list_eq_dec _ _ _); try discriminate.
@@ -344,15 +345,16 @@ Proof.
intros. inv H.
exploit function_ptr_translated; eauto. intros (tf & FIND & TRANSF).
eexists. split.
- - econstructor.
+ - econstructor; eauto.
+ eapply (Genv.init_mem_transf_partial TRANSL); eauto.
+ replace (prog_main tprog) with (prog_main prog). rewrite symbols_preserved. eauto.
symmetry. eapply match_program_main. eauto.
- + exploit function_ptr_translated; eauto.
+ destruct f.
* monadInv TRANSF. rewrite <- H3. symmetry; eapply transf_function_preserves. assumption.
* monadInv TRANSF. assumption.
- - constructor; eauto. constructor. apply transf_fundef_correct; auto.
+ - constructor; eauto.
+ + constructor.
+ + apply transf_fundef_correct; auto.
Qed.
Theorem transf_final_states:
@@ -390,14 +392,29 @@ Proof.
destruct DUPLIC as (i' & H2 & H3). inv H3.
pose symbols_preserved as SYMPRES.
eexists. split.
- + eapply exec_Iload; eauto. erewrite eval_addressing_preserved; eauto.
+ + eapply exec_Iload; eauto; (* is the follow still needed?*) erewrite eval_addressing_preserved; eauto.
+ econstructor; eauto.
+(* Iload notrap1 *)
+ - eapply dupmap_correct in DUPLIC; eauto.
+ destruct DUPLIC as (i' & H2 & H3). inv H3.
+ pose symbols_preserved as SYMPRES.
+ eexists. split.
+ + eapply exec_Iload_notrap1; eauto; erewrite eval_addressing_preserved; eauto.
+ + econstructor; eauto.
+(* Iload notrap2 *)
+ - eapply dupmap_correct in DUPLIC; eauto.
+ destruct DUPLIC as (i' & H2 & H3). inv H3.
+ pose symbols_preserved as SYMPRES.
+ eexists. split.
+ + eapply exec_Iload_notrap2; eauto; erewrite eval_addressing_preserved; eauto.
+ + econstructor; eauto.
+
(* Istore *)
- eapply dupmap_correct in DUPLIC; eauto.
destruct DUPLIC as (i' & H2 & H3). inv H3.
pose symbols_preserved as SYMPRES.
eexists. split.
- + eapply exec_Istore; eauto. erewrite eval_addressing_preserved; eauto.
+ + eapply exec_Istore; eauto; erewrite eval_addressing_preserved; eauto.
+ econstructor; eauto.
(* Icall *)
- eapply dupmap_correct in DUPLIC; eauto.
diff --git a/backend/Inlining.v b/backend/Inlining.v
index f7ee4166..9cf535b9 100644
--- a/backend/Inlining.v
+++ b/backend/Inlining.v
@@ -364,9 +364,9 @@ Definition expand_instr (ctx: context) (pc: node) (i: instruction): mon unit :=
| Iop op args res s =>
set_instr (spc ctx pc)
(Iop (sop ctx op) (sregs ctx args) (sreg ctx res) (spc ctx s))
- | Iload chunk addr args dst s =>
+ | Iload trap chunk addr args dst s =>
set_instr (spc ctx pc)
- (Iload chunk (saddr ctx addr) (sregs ctx args) (sreg ctx dst) (spc ctx s))
+ (Iload trap chunk (saddr ctx addr) (sregs ctx args) (sreg ctx dst) (spc ctx s))
| Istore chunk addr args src s =>
set_instr (spc ctx pc)
(Istore chunk (saddr ctx addr) (sregs ctx args) (sreg ctx src) (spc ctx s))
diff --git a/backend/Inliningaux.ml b/backend/Inliningaux.ml
index 842e0c93..609b2637 100644
--- a/backend/Inliningaux.ml
+++ b/backend/Inliningaux.ml
@@ -57,7 +57,7 @@ let used_in_globvar io gv =
let fun_inline_analysis id io fn =
let inst io nid = function
| Iop (op, args, dest, succ) -> used_id io (globals_operation op)
- | Iload (chunk, addr, args, dest, succ)
+ | Iload (_, chunk, addr, args, dest, succ)
| Istore (chunk, addr, args, dest, succ) -> used_id io (globals_addressing addr)
| Ibuiltin (ef, args, dest, succ) -> used_id io (globals_of_builtin_args args)
| Icall (_, Coq_inr cid, _, _, _)
diff --git a/backend/Inliningproof.v b/backend/Inliningproof.v
index cc84b1cc..c4efaf18 100644
--- a/backend/Inliningproof.v
+++ b/backend/Inliningproof.v
@@ -929,6 +929,15 @@ Proof.
intros. inv H. eauto.
Qed.
+Lemma eval_addressing_none:
+ forall sp' ctx addr rs,
+ eval_addressing ge (Vptr sp' (Ptrofs.repr (dstk ctx))) addr rs = None ->
+ eval_addressing ge (Vptr sp' Ptrofs.zero) (saddr ctx addr) rs = None.
+Proof.
+ intros until rs; intro Heval.
+ destruct addr; destruct rs as [| r0 rs1]; simpl in *; trivial; discriminate.
+Qed.
+
Theorem step_simulation:
forall S1 t S2,
step ge S1 t S2 ->
@@ -976,6 +985,51 @@ Proof.
apply match_stacks_inside_set_reg; auto.
apply agree_set_reg; auto.
+- (* load notrap1 *)
+ exploit tr_funbody_inv; eauto. intros TR; inv TR.
+ left; econstructor; split.
+ eapply plus_one. eapply exec_Iload_notrap1. eassumption.
+ rewrite eval_addressing_preserved with (ge1:=ge) (ge2:=tge).
+ exploit eval_addressing_inj_none.
+ 4: eassumption.
+ intros. eapply symbol_address_inject.
+ eapply match_stacks_inside_globals; eauto.
+ eauto.
+ instantiate (1 := rs'##(sregs ctx args)). eapply agree_val_regs; eauto.
+ rewrite Ptrofs.add_zero_l.
+ apply eval_addressing_none.
+ exact symbols_preserved.
+ econstructor; eauto.
+ apply match_stacks_inside_set_reg; auto.
+ apply agree_set_reg; auto.
+
+- (* load notrap2 *)
+ exploit tr_funbody_inv; eauto. intros TR; inv TR.
+
+ exploit eval_addressing_inject.
+ eapply match_stacks_inside_globals; eauto.
+ eexact SP.
+ instantiate (2 := rs##args). instantiate (1 := rs'##(sregs ctx args)). eapply agree_val_regs; eauto.
+ eauto.
+ fold (saddr ctx addr). intros [a' [P Q]].
+
+ destruct (Mem.loadv chunk m' a') eqn:Hload'.
+ + left; econstructor; split.
+ eapply plus_one.
+ eapply exec_Iload; eauto.
+ try (rewrite <- P; apply eval_addressing_preserved; exact symbols_preserved).
+ econstructor; eauto.
+ apply match_stacks_inside_set_reg; auto.
+ apply agree_set_reg; auto.
+
+ + left; econstructor; split.
+ eapply plus_one.
+ eapply exec_Iload_notrap2; eauto.
+ try (rewrite <- P; apply eval_addressing_preserved; exact symbols_preserved).
+ econstructor; eauto.
+ apply match_stacks_inside_set_reg; auto.
+ apply agree_set_reg; auto.
+
- (* store *)
exploit tr_funbody_inv; eauto. intros TR; inv TR.
exploit eval_addressing_inject.
diff --git a/backend/Inliningspec.v b/backend/Inliningspec.v
index c345c942..e20fb373 100644
--- a/backend/Inliningspec.v
+++ b/backend/Inliningspec.v
@@ -270,10 +270,10 @@ Inductive tr_instr: context -> node -> instruction -> code -> Prop :=
Ple res ctx.(mreg) ->
c!(spc ctx pc) = Some (Iop (sop ctx op) (sregs ctx args) (sreg ctx res) (spc ctx s)) ->
tr_instr ctx pc (Iop op args res s) c
- | tr_load: forall ctx pc c chunk addr args res s,
+ | tr_load: forall ctx pc c trap chunk addr args res s,
Ple res ctx.(mreg) ->
- c!(spc ctx pc) = Some (Iload chunk (saddr ctx addr) (sregs ctx args) (sreg ctx res) (spc ctx s)) ->
- tr_instr ctx pc (Iload chunk addr args res s) c
+ c!(spc ctx pc) = Some (Iload trap chunk (saddr ctx addr) (sregs ctx args) (sreg ctx res) (spc ctx s)) ->
+ tr_instr ctx pc (Iload trap chunk addr args res s) c
| tr_store: forall ctx pc c chunk addr args src s,
c!(spc ctx pc) = Some (Istore chunk (saddr ctx addr) (sregs ctx args) (sreg ctx src) (spc ctx s)) ->
tr_instr ctx pc (Istore chunk addr args src s) c
diff --git a/backend/LTL.v b/backend/LTL.v
index 5e7eec8c..ee8b4826 100644
--- a/backend/LTL.v
+++ b/backend/LTL.v
@@ -29,7 +29,7 @@ Definition node := positive.
Inductive instruction: Type :=
| Lop (op: operation) (args: list mreg) (res: mreg)
- | Lload (chunk: memory_chunk) (addr: addressing) (args: list mreg) (dst: mreg)
+ | Lload (trap : trapping_mode) (chunk: memory_chunk) (addr: addressing) (args: list mreg) (dst: mreg)
| Lgetstack (sl: slot) (ofs: Z) (ty: typ) (dst: mreg)
| Lsetstack (src: mreg) (sl: slot) (ofs: Z) (ty: typ)
| Lstore (chunk: memory_chunk) (addr: addressing) (args: list mreg) (src: mreg)
@@ -209,11 +209,24 @@ Inductive step: state -> trace -> state -> Prop :=
rs' = Locmap.set (R res) v (undef_regs (destroyed_by_op op) rs) ->
step (Block s f sp (Lop op args res :: bb) rs m)
E0 (Block s f sp bb rs' m)
- | exec_Lload: forall s f sp chunk addr args dst bb rs m a v rs',
+ | exec_Lload: forall s f sp trap chunk addr args dst bb rs m a v rs',
eval_addressing ge sp addr (reglist rs args) = Some a ->
Mem.loadv chunk m a = Some v ->
rs' = Locmap.set (R dst) v (undef_regs (destroyed_by_load chunk addr) rs) ->
- step (Block s f sp (Lload chunk addr args dst :: bb) rs m)
+ step (Block s f sp (Lload trap chunk addr args dst :: bb) rs m)
+ E0 (Block s f sp bb rs' m)
+ | exec_Lload_notrap1: forall s f sp chunk addr args dst bb rs m rs',
+ eval_addressing ge sp addr (reglist rs args) = None ->
+ rs' = Locmap.set (R dst) (default_notrap_load_value chunk)
+ (undef_regs (destroyed_by_load chunk addr) rs) ->
+ step (Block s f sp (Lload NOTRAP chunk addr args dst :: bb) rs m)
+ E0 (Block s f sp bb rs' m)
+ | exec_Lload_notrap2: forall s f sp chunk addr args dst bb rs m a rs',
+ eval_addressing ge sp addr (reglist rs args) = Some a ->
+ Mem.loadv chunk m a = None ->
+ rs' = Locmap.set (R dst) (default_notrap_load_value chunk)
+ (undef_regs (destroyed_by_load chunk addr) rs) ->
+ step (Block s f sp (Lload NOTRAP chunk addr args dst :: bb) rs m)
E0 (Block s f sp bb rs' m)
| exec_Lgetstack: forall s f sp sl ofs ty dst bb rs m rs',
rs' = Locmap.set (R dst) (rs (S sl ofs ty)) (undef_regs (destroyed_by_getstack sl) rs) ->
diff --git a/backend/Linear.v b/backend/Linear.v
index 447c6ba6..1443f795 100644
--- a/backend/Linear.v
+++ b/backend/Linear.v
@@ -28,7 +28,7 @@ Inductive instruction: Type :=
| Lgetstack: slot -> Z -> typ -> mreg -> instruction
| Lsetstack: mreg -> slot -> Z -> typ -> instruction
| Lop: operation -> list mreg -> mreg -> instruction
- | Lload: memory_chunk -> addressing -> list mreg -> mreg -> instruction
+ | Lload: trapping_mode -> memory_chunk -> addressing -> list mreg -> mreg -> instruction
| Lstore: memory_chunk -> addressing -> list mreg -> mreg -> instruction
| Lcall: signature -> mreg + ident -> instruction
| Ltailcall: signature -> mreg + ident -> instruction
@@ -160,11 +160,28 @@ Inductive step: state -> trace -> state -> Prop :=
step (State s f sp (Lop op args res :: b) rs m)
E0 (State s f sp b rs' m)
| exec_Lload:
- forall s f sp chunk addr args dst b rs m a v rs',
+ forall s f sp trap chunk addr args dst b rs m a v rs',
eval_addressing ge sp addr (reglist rs args) = Some a ->
Mem.loadv chunk m a = Some v ->
rs' = Locmap.set (R dst) v (undef_regs (destroyed_by_load chunk addr) rs) ->
- step (State s f sp (Lload chunk addr args dst :: b) rs m)
+ step (State s f sp (Lload trap chunk addr args dst :: b) rs m)
+ E0 (State s f sp b rs' m)
+ | exec_Lload_notrap1:
+ forall s f sp chunk addr args dst b rs m rs',
+ eval_addressing ge sp addr (reglist rs args) = None ->
+ rs' = Locmap.set (R dst)
+ (default_notrap_load_value chunk)
+ (undef_regs (destroyed_by_load chunk addr) rs) ->
+ step (State s f sp (Lload NOTRAP chunk addr args dst :: b) rs m)
+ E0 (State s f sp b rs' m)
+ | exec_Lload_notrap2:
+ forall s f sp chunk addr args dst b rs m a rs',
+ eval_addressing ge sp addr (reglist rs args) = Some a ->
+ Mem.loadv chunk m a = None ->
+ rs' = Locmap.set (R dst)
+ (default_notrap_load_value chunk)
+ (undef_regs (destroyed_by_load chunk addr) rs) ->
+ step (State s f sp (Lload NOTRAP chunk addr args dst :: b) rs m)
E0 (State s f sp b rs' m)
| exec_Lstore:
forall s f sp chunk addr args src b rs m m' a rs',
diff --git a/backend/Linearize.v b/backend/Linearize.v
index 2cfa4d3c..4216958c 100644
--- a/backend/Linearize.v
+++ b/backend/Linearize.v
@@ -163,8 +163,8 @@ Fixpoint linearize_block (b: LTL.bblock) (k: code) : code :=
| nil => k
| LTL.Lop op args res :: b' =>
Lop op args res :: linearize_block b' k
- | LTL.Lload chunk addr args dst :: b' =>
- Lload chunk addr args dst :: linearize_block b' k
+ | LTL.Lload trap chunk addr args dst :: b' =>
+ Lload trap chunk addr args dst :: linearize_block b' k
| LTL.Lgetstack sl ofs ty dst :: b' =>
Lgetstack sl ofs ty dst :: linearize_block b' k
| LTL.Lsetstack src sl ofs ty :: b' =>
diff --git a/backend/Linearizeproof.v b/backend/Linearizeproof.v
index 10a3d8b2..18dc52a5 100644
--- a/backend/Linearizeproof.v
+++ b/backend/Linearizeproof.v
@@ -585,45 +585,61 @@ Proof.
intros; eapply reachable_successors; eauto.
eapply is_tail_lin_block; eauto. eapply is_tail_find_label; eauto.
- (* Lop *)
+- (* Lop *)
left; econstructor; split. simpl.
apply plus_one. econstructor; eauto.
instantiate (1 := v); rewrite <- H; apply eval_operation_preserved.
exact symbols_preserved.
econstructor; eauto.
- (* Lload *)
+- (* Lload *)
left; econstructor; split. simpl.
- apply plus_one. econstructor.
+ apply plus_one. eapply exec_Lload.
instantiate (1 := a). rewrite <- H; apply eval_addressing_preserved.
exact symbols_preserved. eauto. eauto.
econstructor; eauto.
- (* Lgetstack *)
+- (* Lload notrap1 *)
+ left; econstructor; split. simpl.
+ apply plus_one. eapply exec_Lload_notrap1.
+ rewrite <- H.
+ apply eval_addressing_preserved.
+ exact symbols_preserved. eauto.
+ econstructor; eauto.
+
+- (* Lload notrap2 *)
+ left; econstructor; split. simpl.
+ apply plus_one. eapply exec_Lload_notrap2.
+ rewrite <- H.
+ apply eval_addressing_preserved.
+ exact symbols_preserved. eauto. eauto.
+ econstructor; eauto.
+
+- (* Lgetstack *)
left; econstructor; split. simpl.
apply plus_one. econstructor; eauto.
econstructor; eauto.
- (* Lsetstack *)
+- (* Lsetstack *)
left; econstructor; split. simpl.
apply plus_one. econstructor; eauto.
econstructor; eauto.
- (* Lstore *)
+- (* Lstore *)
left; econstructor; split. simpl.
apply plus_one. econstructor.
instantiate (1 := a). rewrite <- H; apply eval_addressing_preserved.
exact symbols_preserved. eauto. eauto.
econstructor; eauto.
- (* Lcall *)
+- (* Lcall *)
exploit find_function_translated; eauto. intros [tfd [A B]].
left; econstructor; split. simpl.
apply plus_one. econstructor; eauto.
symmetry; eapply sig_preserved; eauto.
econstructor; eauto. constructor; auto. econstructor; eauto.
- (* Ltailcall *)
+- (* Ltailcall *)
exploit find_function_translated; eauto. intros [tfd [A B]].
left; econstructor; split. simpl.
apply plus_one. econstructor; eauto.
@@ -633,18 +649,18 @@ Proof.
rewrite (match_parent_locset _ _ STACKS).
econstructor; eauto.
- (* Lbuiltin *)
+- (* Lbuiltin *)
left; econstructor; split. simpl.
apply plus_one. eapply exec_Lbuiltin; eauto.
eapply eval_builtin_args_preserved with (ge1 := ge); eauto. exact symbols_preserved.
eapply external_call_symbols_preserved; eauto. apply senv_preserved.
econstructor; eauto.
- (* Lbranch *)
+- (* Lbranch *)
assert ((reachable f)!!pc = true). apply REACH; simpl; auto.
right; split. simpl; omega. split. auto. simpl. econstructor; eauto.
- (* Lcond *)
+- (* Lcond *)
assert (REACH1: (reachable f)!!pc1 = true) by (apply REACH; simpl; auto).
assert (REACH2: (reachable f)!!pc2 = true) by (apply REACH; simpl; auto).
simpl linearize_block.
@@ -670,18 +686,18 @@ Proof.
apply plus_one. eapply exec_Lcond_false. eauto. eauto.
econstructor; eauto.
- (* Ljumptable *)
+- (* Ljumptable *)
assert (REACH': (reachable f)!!pc = true).
apply REACH. simpl. eapply list_nth_z_in; eauto.
right; split. simpl; omega. split. auto. econstructor; eauto.
- (* Lreturn *)
+- (* Lreturn *)
left; econstructor; split.
simpl. apply plus_one. econstructor; eauto.
rewrite (stacksize_preserved _ _ TRF). eauto.
rewrite (match_parent_locset _ _ STACKS). econstructor; eauto.
- (* internal functions *)
+- (* internal functions *)
assert (REACH: (reachable f)!!(LTL.fn_entrypoint f) = true).
apply reachable_entrypoint.
monadInv H7.
@@ -691,13 +707,13 @@ Proof.
generalize EQ; intro EQ'; monadInv EQ'. simpl.
econstructor; eauto. simpl. eapply is_tail_add_branch. constructor.
- (* external function *)
+- (* external function *)
monadInv H8. left; econstructor; split.
apply plus_one. eapply exec_function_external; eauto.
eapply external_call_symbols_preserved; eauto. apply senv_preserved.
econstructor; eauto.
- (* return *)
+- (* return *)
inv H3. inv H1.
left; econstructor; split.
apply plus_one. econstructor.
diff --git a/backend/Lineartyping.v b/backend/Lineartyping.v
index 0e3b7c8e..3fe61470 100644
--- a/backend/Lineartyping.v
+++ b/backend/Lineartyping.v
@@ -76,7 +76,7 @@ Definition wt_instr (i: instruction) : bool :=
let (targs, tres) := type_of_operation op in
subtype tres (mreg_type res)
end
- | Lload chunk addr args dst =>
+ | Lload trap chunk addr args dst =>
subtype (type_of_chunk chunk) (mreg_type dst)
| Ltailcall sg ros =>
zeq (size_arguments sg) 0
@@ -321,17 +321,34 @@ Local Opaque mreg_type.
+ (* other ops *)
destruct (type_of_operation op) as [ty_args ty_res] eqn:TYOP. InvBooleans.
econstructor; eauto.
- apply wt_setreg. eapply Val.has_subtype; eauto.
+
+ apply wt_setreg; auto; try (apply wt_undef_regs; auto).
+ eapply Val.has_subtype; eauto.
+
change ty_res with (snd (ty_args, ty_res)). rewrite <- TYOP. eapply type_of_operation_sound; eauto.
red; intros; subst op. simpl in ISMOVE.
destruct args; try discriminate. destruct args; discriminate.
- apply wt_undef_regs; auto.
+ (* no longer needed apply wt_undef_regs; auto. *)
- (* load *)
simpl in *; InvBooleans.
econstructor; eauto.
apply wt_setreg. eapply Val.has_subtype; eauto.
destruct a; simpl in H0; try discriminate. eapply Mem.load_type; eauto.
apply wt_undef_regs; auto.
+- (* load notrap1 *)
+ simpl in *; InvBooleans.
+ econstructor; eauto.
+ apply wt_setreg. eapply Val.has_subtype; eauto.
+ unfold default_notrap_load_value.
+ constructor.
+ apply wt_undef_regs; auto.
+- (* load notrap2 *)
+ simpl in *; InvBooleans.
+ econstructor; eauto.
+ apply wt_setreg. eapply Val.has_subtype; eauto.
+ unfold default_notrap_load_value.
+ constructor.
+ apply wt_undef_regs; auto.
- (* store *)
simpl in *; InvBooleans.
econstructor. eauto. eauto. eauto.
diff --git a/backend/Liveness.v b/backend/Liveness.v
index 16533158..afe11ae6 100644
--- a/backend/Liveness.v
+++ b/backend/Liveness.v
@@ -79,7 +79,7 @@ Definition transfer
reg_list_live args (reg_dead res after)
else
after
- | Iload chunk addr args dst s =>
+ | Iload trap chunk addr args dst s =>
if Regset.mem dst after then
reg_list_live args (reg_dead dst after)
else
diff --git a/backend/Mach.v b/backend/Mach.v
index 9fdee9eb..1c6fdb18 100644
--- a/backend/Mach.v
+++ b/backend/Mach.v
@@ -56,7 +56,7 @@ Inductive instruction: Type :=
| Msetstack: mreg -> ptrofs -> typ -> instruction
| Mgetparam: ptrofs -> typ -> mreg -> instruction
| Mop: operation -> list mreg -> mreg -> instruction
- | Mload: memory_chunk -> addressing -> list mreg -> mreg -> instruction
+ | Mload: trapping_mode -> memory_chunk -> addressing -> list mreg -> mreg -> instruction
| Mstore: memory_chunk -> addressing -> list mreg -> mreg -> instruction
| Mcall: signature -> mreg + ident -> instruction
| Mtailcall: signature -> mreg + ident -> instruction
@@ -321,11 +321,24 @@ Inductive step: state -> trace -> state -> Prop :=
step (State s f sp (Mop op args res :: c) rs m)
E0 (State s f sp c rs' m)
| exec_Mload:
- forall s f sp chunk addr args dst c rs m a v rs',
+ forall s f sp trap chunk addr args dst c rs m a v rs',
eval_addressing ge sp addr rs##args = Some a ->
Mem.loadv chunk m a = Some v ->
rs' = ((undef_regs (destroyed_by_load chunk addr) rs)#dst <- v) ->
- step (State s f sp (Mload chunk addr args dst :: c) rs m)
+ step (State s f sp (Mload trap chunk addr args dst :: c) rs m)
+ E0 (State s f sp c rs' m)
+ | exec_Mload_notrap1:
+ forall s f sp chunk addr args dst c rs m rs',
+ eval_addressing ge sp addr rs##args = None ->
+ rs' = ((undef_regs (destroyed_by_load chunk addr) rs)#dst <- (default_notrap_load_value chunk)) ->
+ step (State s f sp (Mload NOTRAP chunk addr args dst :: c) rs m)
+ E0 (State s f sp c rs' m)
+ | exec_Mload_notrap2:
+ forall s f sp chunk addr args dst c rs m a rs',
+ eval_addressing ge sp addr rs##args = Some a ->
+ Mem.loadv chunk m a = None ->
+ rs' = ((undef_regs (destroyed_by_load chunk addr) rs)#dst <- (default_notrap_load_value chunk)) ->
+ step (State s f sp (Mload NOTRAP chunk addr args dst :: c) rs m)
E0 (State s f sp c rs' m)
| exec_Mstore:
forall s f sp chunk addr args src c rs m m' a rs',
diff --git a/backend/PrintLTL.ml b/backend/PrintLTL.ml
index 1c449e74..b309a9f2 100644
--- a/backend/PrintLTL.ml
+++ b/backend/PrintLTL.ml
@@ -61,9 +61,10 @@ let print_succ pp s dfl =
let print_instruction pp succ = function
| Lop(op, args, res) ->
fprintf pp "%a = %a" mreg res (print_operation mreg) (op, args)
- | Lload(chunk, addr, args, dst) ->
- fprintf pp "%a = %s[%a]"
- mreg dst (name_of_chunk chunk) (print_addressing mreg) (addr, args)
+ | Lload(trap,chunk, addr, args, dst) ->
+ fprintf pp "%a = %s[%a]%a"
+ mreg dst (name_of_chunk chunk) (print_addressing mreg) (addr, args)
+ print_trapping_mode trap
| Lgetstack(sl, ofs, ty, dst) ->
fprintf pp "%a = %a" mreg dst slot (sl, ofs, ty)
| Lsetstack(src, sl, ofs, ty) ->
diff --git a/backend/PrintMach.ml b/backend/PrintMach.ml
index 517f3037..70e65832 100644
--- a/backend/PrintMach.ml
+++ b/backend/PrintMach.ml
@@ -48,10 +48,11 @@ let print_instruction pp i =
| Mop(op, args, res) ->
fprintf pp "\t%a = %a\n"
reg res (PrintOp.print_operation reg) (op, args)
- | Mload(chunk, addr, args, dst) ->
- fprintf pp "\t%a = %s[%a]\n"
+ | Mload(trap, chunk, addr, args, dst) ->
+ fprintf pp "\t%a = %s[%a]%a\n"
reg dst (name_of_chunk chunk)
(PrintOp.print_addressing reg) (addr, args)
+ print_trapping_mode trap
| Mstore(chunk, addr, args, src) ->
fprintf pp "\t%s[%a] = %a\n"
(name_of_chunk chunk)
diff --git a/backend/PrintRTL.ml b/backend/PrintRTL.ml
index 841540b6..c25773e5 100644
--- a/backend/PrintRTL.ml
+++ b/backend/PrintRTL.ml
@@ -50,10 +50,11 @@ let print_instruction pp (pc, i) =
fprintf pp "%a = %a\n"
reg res (PrintOp.print_operation reg) (op, args);
print_succ pp s (pc - 1)
- | Iload(chunk, addr, args, dst, s) ->
- fprintf pp "%a = %s[%a]\n"
+ | Iload(trap, chunk, addr, args, dst, s) ->
+ fprintf pp "%a = %s[%a]%a\n"
reg dst (name_of_chunk chunk)
- (PrintOp.print_addressing reg) (addr, args);
+ (PrintOp.print_addressing reg) (addr, args)
+ print_trapping_mode trap;
print_succ pp s (pc - 1)
| Istore(chunk, addr, args, src, s) ->
fprintf pp "%s[%a] = %a\n"
diff --git a/backend/PrintXTL.ml b/backend/PrintXTL.ml
index 6432682a..1c7655fb 100644
--- a/backend/PrintXTL.ml
+++ b/backend/PrintXTL.ml
@@ -86,9 +86,10 @@ let print_instruction pp succ = function
fprintf pp "(%a) = (%a) using %a, %a" vars dsts vars srcs var t1 var t2
| Xop(op, args, res) ->
fprintf pp "%a = %a" var res (print_operation var) (op, args)
- | Xload(chunk, addr, args, dst) ->
- fprintf pp "%a = %s[%a]"
- var dst (name_of_chunk chunk) (print_addressing var) (addr, args)
+ | Xload(trap, chunk, addr, args, dst) ->
+ fprintf pp "%a = %s[%a]%a"
+ var dst (name_of_chunk chunk) (print_addressing var) (addr, args)
+ print_trapping_mode trap
| Xstore(chunk, addr, args, src) ->
fprintf pp "%s[%a] = %a"
(name_of_chunk chunk) (print_addressing var) (addr, args) var src
diff --git a/backend/RTL.v b/backend/RTL.v
index 9599a24a..29a49311 100644
--- a/backend/RTL.v
+++ b/backend/RTL.v
@@ -43,11 +43,12 @@ Inductive instruction: Type :=
(** [Iop op args dest succ] performs the arithmetic operation [op]
over the values of registers [args], stores the result in [dest],
and branches to [succ]. *)
- | Iload: memory_chunk -> addressing -> list reg -> reg -> node -> instruction
- (** [Iload chunk addr args dest succ] loads a [chunk] quantity from
+ | Iload: trapping_mode -> memory_chunk -> addressing -> list reg -> reg -> node -> instruction
+ (** [Iload trap chunk addr args dest succ] loads a [chunk] quantity from
the address determined by the addressing mode [addr] and the
values of the [args] registers, stores the quantity just read
- into [dest], and branches to [succ]. *)
+ into [dest], and branches to [succ].
+ If trap=NOTRAP, then failures lead to a default value written to [dest]. *)
| Istore: memory_chunk -> addressing -> list reg -> reg -> node -> instruction
(** [Istore chunk addr args src succ] stores the value of register
[src] in the [chunk] quantity at the
@@ -212,12 +213,25 @@ Inductive step: state -> trace -> state -> Prop :=
step (State s f sp pc rs m)
E0 (State s f sp pc' (rs#res <- v) m)
| exec_Iload:
- forall s f sp pc rs m chunk addr args dst pc' a v,
- (fn_code f)!pc = Some(Iload chunk addr args dst pc') ->
+ forall s f sp pc rs m trap chunk addr args dst pc' a v,
+ (fn_code f)!pc = Some(Iload trap chunk addr args dst pc') ->
eval_addressing ge sp addr rs##args = Some a ->
Mem.loadv chunk m a = Some v ->
step (State s f sp pc rs m)
E0 (State s f sp pc' (rs#dst <- v) m)
+ | exec_Iload_notrap1:
+ forall s f sp pc rs m chunk addr args dst pc',
+ (fn_code f)!pc = Some(Iload NOTRAP chunk addr args dst pc') ->
+ eval_addressing ge sp addr rs##args = None ->
+ step (State s f sp pc rs m)
+ E0 (State s f sp pc' (rs#dst <- (default_notrap_load_value chunk)) m)
+ | exec_Iload_notrap2:
+ forall s f sp pc rs m chunk addr args dst pc' a,
+ (fn_code f)!pc = Some(Iload NOTRAP chunk addr args dst pc') ->
+ eval_addressing ge sp addr rs##args = Some a ->
+ Mem.loadv chunk m a = None->
+ step (State s f sp pc rs m)
+ E0 (State s f sp pc' (rs#dst <- (default_notrap_load_value chunk)) m)
| exec_Istore:
forall s f sp pc rs m chunk addr args src pc' a m',
(fn_code f)!pc = Some(Istore chunk addr args src pc') ->
@@ -299,8 +313,8 @@ Proof.
Qed.
Lemma exec_Iload':
- forall s f sp pc rs m chunk addr args dst pc' rs' a v,
- (fn_code f)!pc = Some(Iload chunk addr args dst pc') ->
+ forall s f sp pc rs m trap chunk addr args dst pc' rs' a v,
+ (fn_code f)!pc = Some(Iload trap chunk addr args dst pc') ->
eval_addressing ge sp addr rs##args = Some a ->
Mem.loadv chunk m a = Some v ->
rs' = (rs#dst <- v) ->
@@ -384,7 +398,7 @@ Definition successors_instr (i: instruction) : list node :=
match i with
| Inop s => s :: nil
| Iop op args res s => s :: nil
- | Iload chunk addr args dst s => s :: nil
+ | Iload trap chunk addr args dst s => s :: nil
| Istore chunk addr args src s => s :: nil
| Icall sig ros args res s => s :: nil
| Itailcall sig ros args => nil
@@ -403,7 +417,7 @@ Definition instr_uses (i: instruction) : list reg :=
match i with
| Inop s => nil
| Iop op args res s => args
- | Iload chunk addr args dst s => args
+ | Iload trap chunk addr args dst s => args
| Istore chunk addr args src s => src :: args
| Icall sig (inl r) args res s => r :: args
| Icall sig (inr id) args res s => args
@@ -422,7 +436,7 @@ Definition instr_defs (i: instruction) : option reg :=
match i with
| Inop s => None
| Iop op args res s => Some res
- | Iload chunk addr args dst s => Some dst
+ | Iload trap chunk addr args dst s => Some dst
| Istore chunk addr args src s => None
| Icall sig ros args res s => Some res
| Itailcall sig ros args => None
@@ -462,7 +476,7 @@ Definition max_reg_instr (m: positive) (pc: node) (i: instruction) :=
match i with
| Inop s => m
| Iop op args res s => fold_left Pos.max args (Pos.max res m)
- | Iload chunk addr args dst s => fold_left Pos.max args (Pos.max dst m)
+ | Iload trap chunk addr args dst s => fold_left Pos.max args (Pos.max dst m)
| Istore chunk addr args src s => fold_left Pos.max args (Pos.max src m)
| Icall sig (inl r) args res s => fold_left Pos.max args (Pos.max r (Pos.max res m))
| Icall sig (inr id) args res s => fold_left Pos.max args (Pos.max res m)
diff --git a/backend/RTLgen.v b/backend/RTLgen.v
index 9d7a8506..a0ca0f17 100644
--- a/backend/RTLgen.v
+++ b/backend/RTLgen.v
@@ -436,7 +436,7 @@ Fixpoint transl_expr (map: mapping) (a: expr) (rd: reg) (nd: node)
transl_exprlist map al rl no
| Eload chunk addr al =>
do rl <- alloc_regs map al;
- do no <- add_instr (Iload chunk addr rl rd nd);
+ do no <- add_instr (Iload TRAP chunk addr rl rd nd);
transl_exprlist map al rl no
| Econdition a b c =>
do nfalse <- transl_expr map c rd nd;
diff --git a/backend/RTLgenspec.v b/backend/RTLgenspec.v
index 17022a7d..aa83da6d 100644
--- a/backend/RTLgenspec.v
+++ b/backend/RTLgenspec.v
@@ -707,7 +707,7 @@ Inductive tr_expr (c: code):
tr_expr c map pr (Eop op al) ns nd rd dst
| tr_Eload: forall map pr chunk addr al ns nd rd n1 rl dst,
tr_exprlist c map pr al ns n1 rl ->
- c!n1 = Some (Iload chunk addr rl rd nd) ->
+ c!n1 = Some (Iload TRAP chunk addr rl rd nd) ->
reg_map_ok map rd dst -> ~In rd pr ->
tr_expr c map pr (Eload chunk addr al) ns nd rd dst
| tr_Econdition: forall map pr a ifso ifnot ns nd rd ntrue nfalse dst,
diff --git a/backend/RTLtyping.v b/backend/RTLtyping.v
index 8336d1bf..74bfa582 100644
--- a/backend/RTLtyping.v
+++ b/backend/RTLtyping.v
@@ -104,11 +104,11 @@ Inductive wt_instr : instruction -> Prop :=
valid_successor s ->
wt_instr (Iop op args res s)
| wt_Iload:
- forall chunk addr args dst s,
+ forall trap chunk addr args dst s,
map env args = type_of_addressing addr ->
env dst = type_of_chunk chunk ->
valid_successor s ->
- wt_instr (Iload chunk addr args dst s)
+ wt_instr (Iload trap chunk addr args dst s)
| wt_Istore:
forall chunk addr args src s,
map env args = type_of_addressing addr ->
@@ -282,7 +282,7 @@ Definition type_instr (e: S.typenv) (i: instruction) : res S.typenv :=
else
(let (targs, tres) := type_of_operation op in
do e1 <- S.set_list e args targs; S.set e1 res tres)
- | Iload chunk addr args dst s =>
+ | Iload trap chunk addr args dst s =>
do x <- check_successor s;
do e1 <- S.set_list e args (type_of_addressing addr);
S.set e1 dst (type_of_chunk chunk)
@@ -841,14 +841,24 @@ Proof.
Qed.
Lemma wt_exec_Iload:
- forall env f chunk addr args dst s m a v rs,
- wt_instr f env (Iload chunk addr args dst s) ->
+ forall env f trap chunk addr args dst s m a v rs,
+ wt_instr f env (Iload trap chunk addr args dst s) ->
Mem.loadv chunk m a = Some v ->
wt_regset env rs ->
wt_regset env (rs#dst <- v).
Proof.
intros. destruct a; simpl in H0; try discriminate. inv H.
- eapply wt_regset_assign; eauto. rewrite H8; eapply Mem.load_type; eauto.
+ 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:
@@ -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/Regalloc.ml b/backend/Regalloc.ml
index 7db8a866..f2658b04 100644
--- a/backend/Regalloc.ml
+++ b/backend/Regalloc.ml
@@ -249,18 +249,18 @@ let block_of_RTL_instr funsig tyenv = function
else
let t = new_temp (tyenv res) in (t :: args2', t) in
movelist args1 args3 (Xop(op, args3, res3) :: move res3 res1 [Xbranch s])
- | RTL.Iload(chunk, addr, args, dst, s) ->
+ | RTL.Iload(trap, chunk, addr, args, dst, s) ->
if Archi.splitlong && chunk = Mint64 then begin
match offset_addressing addr (coqint_of_camlint 4l) with
| None -> assert false
| Some addr' ->
- [Xload(Mint32, addr, vregs tyenv args,
+ [Xload(trap, Mint32, addr, vregs tyenv args,
V((if Archi.big_endian then dst else twin_reg dst), Tint));
- Xload(Mint32, addr', vregs tyenv args,
+ Xload(trap, Mint32, addr', vregs tyenv args,
V((if Archi.big_endian then twin_reg dst else dst), Tint));
Xbranch s]
end else
- [Xload(chunk, addr, vregs tyenv args, vreg tyenv dst); Xbranch s]
+ [Xload(trap, chunk, addr, vregs tyenv args, vreg tyenv dst); Xbranch s]
| RTL.Istore(chunk, addr, args, src, s) ->
if Archi.splitlong && chunk = Mint64 then begin
match offset_addressing addr (coqint_of_camlint 4l) with
@@ -364,7 +364,7 @@ let live_before instr after =
if VSet.mem res after
then vset_addlist args (VSet.remove res after)
else after
- | Xload(chunk, addr, args, dst) ->
+ | Xload(trap, chunk, addr, args, dst) ->
if VSet.mem dst after
then vset_addlist args (VSet.remove dst after)
else after
@@ -459,7 +459,7 @@ let dce_instr instr after k =
if VSet.mem res after
then instr :: k
else k
- | Xload(chunk, addr, args, dst) ->
+ | Xload(trap, chunk, addr, args, dst) ->
if VSet.mem dst after
then instr :: k
else k
@@ -550,7 +550,7 @@ let spill_costs f =
(* temps must not be spilled *)
| Xop(op, args, res) ->
charge_list 10 1 args; charge 10 1 res
- | Xload(chunk, addr, args, dst) ->
+ | Xload(trap, chunk, addr, args, dst) ->
charge_list 10 1 args; charge 10 1 dst
| Xstore(chunk, addr, args, src) ->
charge_list 10 1 args; charge 10 1 src
@@ -677,7 +677,7 @@ let add_interfs_instr g instr live =
(vset_addlist (res :: argl) (VSet.remove res live))
end;
add_interfs_destroyed g (VSet.remove res live) (destroyed_by_op op)
- | Xload(chunk, addr, args, dst) ->
+ | Xload(trap, chunk, addr, args, dst) ->
add_interfs_def g dst live;
add_interfs_destroyed g (VSet.remove dst live)
(destroyed_by_load chunk addr)
@@ -782,7 +782,7 @@ let tospill_instr alloc instr ts =
ts
| Xop(op, args, res) ->
addlist_tospill alloc args (add_tospill alloc res ts)
- | Xload(chunk, addr, args, dst) ->
+ | Xload(trap, chunk, addr, args, dst) ->
addlist_tospill alloc args (add_tospill alloc dst ts)
| Xstore(chunk, addr, args, src) ->
addlist_tospill alloc args (add_tospill alloc src ts)
@@ -964,10 +964,10 @@ let spill_instr tospill eqs instr =
add res tmp (kill tmp (kill res eqs2)))
end
end
- | Xload(chunk, addr, args, dst) ->
+ | Xload(trap, chunk, addr, args, dst) ->
let (args', c1, eqs1) = reload_vars tospill eqs args in
let (dst', c2, eqs2) = save_var tospill eqs1 dst in
- (c1 @ Xload(chunk, addr, args', dst') :: c2, eqs2)
+ (c1 @ Xload(trap, chunk, addr, args', dst') :: c2, eqs2)
| Xstore(chunk, addr, args, src) ->
let (args', c1, eqs1) = reload_vars tospill eqs args in
let (src', c2, eqs2) = reload_var tospill eqs1 src in
@@ -1115,8 +1115,8 @@ let transl_instr alloc instr k =
LTL.Lop(Omove, [rarg1], rres) ::
LTL.Lop(op, rres :: rargl, rres) :: k
end
- | Xload(chunk, addr, args, dst) ->
- LTL.Lload(chunk, addr, mregs_of alloc args, mreg_of alloc dst) :: k
+ | Xload(trap, chunk, addr, args, dst) ->
+ LTL.Lload(trap, chunk, addr, mregs_of alloc args, mreg_of alloc dst) :: k
| Xstore(chunk, addr, args, src) ->
LTL.Lstore(chunk, addr, mregs_of alloc args, mreg_of alloc src) :: k
| Xcall(sg, vos, args, res) ->
diff --git a/backend/Renumber.v b/backend/Renumber.v
index 10f58251..7ba16658 100644
--- a/backend/Renumber.v
+++ b/backend/Renumber.v
@@ -43,7 +43,7 @@ Definition renum_instr (i: instruction) : instruction :=
match i with
| Inop s => Inop (renum_pc s)
| Iop op args res s => Iop op args res (renum_pc s)
- | Iload chunk addr args res s => Iload chunk addr args res (renum_pc s)
+ | Iload trap chunk addr args res s => Iload trap chunk addr args res (renum_pc s)
| Istore chunk addr args src s => Istore chunk addr args src (renum_pc s)
| Icall sg ros args res s => Icall sg ros args res (renum_pc s)
| Itailcall sg ros args => i
diff --git a/backend/Renumberproof.v b/backend/Renumberproof.v
index 7cda9425..2e161965 100644
--- a/backend/Renumberproof.v
+++ b/backend/Renumberproof.v
@@ -175,6 +175,18 @@ Proof.
rewrite <- H0. apply eval_addressing_preserved. exact symbols_preserved.
eapply exec_Iload; eauto.
constructor; auto. eapply reach_succ; eauto. simpl; auto.
+ (* load notrap1 *)
+ econstructor; split.
+ assert (eval_addressing tge sp addr rs ## args = None).
+ rewrite <- H0. apply eval_addressing_preserved. exact symbols_preserved.
+ eapply exec_Iload_notrap1; eauto.
+ constructor; auto. eapply reach_succ; eauto. simpl; auto.
+ (* load notrap2 *)
+ econstructor; split.
+ assert (eval_addressing tge sp addr rs ## args = Some a).
+ rewrite <- H0. apply eval_addressing_preserved. exact symbols_preserved.
+ eapply exec_Iload_notrap2; eauto.
+ constructor; auto. eapply reach_succ; eauto. simpl; auto.
(* store *)
econstructor; split.
assert (eval_addressing tge sp addr rs ## args = Some a).
diff --git a/backend/Splitting.ml b/backend/Splitting.ml
index 40f09c3d..78eb66a5 100644
--- a/backend/Splitting.ml
+++ b/backend/Splitting.ml
@@ -151,8 +151,8 @@ let ren_instr f maps pc i =
| Inop s -> Inop s
| Iop(op, args, res, s) ->
Iop(op, ren_regs before args, ren_reg after res, s)
- | Iload(chunk, addr, args, dst, s) ->
- Iload(chunk, addr, ren_regs before args, ren_reg after dst, s)
+ | Iload(trap, chunk, addr, args, dst, s) ->
+ Iload(trap, chunk, addr, ren_regs before args, ren_reg after dst, s)
| Istore(chunk, addr, args, src, s) ->
Istore(chunk, addr, ren_regs before args, ren_reg before src, s)
| Icall(sg, ros, args, res, s) ->
diff --git a/backend/Stacking.v b/backend/Stacking.v
index 7b382d05..0e3f2832 100644
--- a/backend/Stacking.v
+++ b/backend/Stacking.v
@@ -133,8 +133,8 @@ Definition transl_instr
end
| Lop op args res =>
Mop (transl_op fe op) args res :: k
- | Lload chunk addr args dst =>
- Mload chunk (transl_addr fe addr) args dst :: k
+ | Lload trap chunk addr args dst =>
+ Mload trap chunk (transl_addr fe addr) args dst :: k
| Lstore chunk addr args src =>
Mstore chunk (transl_addr fe addr) args src :: k
| Lcall sig ros =>
diff --git a/backend/Stackingproof.v b/backend/Stackingproof.v
index 326fab61..d3fcdb91 100644
--- a/backend/Stackingproof.v
+++ b/backend/Stackingproof.v
@@ -1918,6 +1918,46 @@ Proof.
apply agree_regs_set_reg. rewrite transl_destroyed_by_load. apply agree_regs_undef_regs; auto. auto.
apply agree_locs_set_reg. apply agree_locs_undef_locs. auto. apply destroyed_by_load_caller_save. auto.
+- (* Lload notrap1*)
+ assert (eval_addressing ge (Vptr sp' Ptrofs.zero) (transl_addr (make_env (function_bounds f)) addr) rs0##args = None) as Haddress.
+ eapply eval_addressing_inject_none; eauto.
+ eapply globalenv_inject_preserves_globals. eapply sep_proj2. eapply sep_proj2. eapply sep_proj2. eexact SEP.
+ eapply agree_reglist; eauto.
+ econstructor; split.
+ apply plus_one. apply exec_Mload_notrap1.
+ rewrite <- Haddress. apply eval_addressing_preserved. exact symbols_preserved.
+ eauto. econstructor; eauto with coqlib.
+ apply agree_regs_set_reg. rewrite transl_destroyed_by_load. apply agree_regs_undef_regs; auto. auto.
+ apply agree_locs_set_reg. apply agree_locs_undef_locs. auto. apply destroyed_by_load_caller_save. auto.
+
+- (* Lload notrap2 *)
+ assert (exists a',
+ eval_addressing ge (Vptr sp' Ptrofs.zero) (transl_addr (make_env (function_bounds f)) addr) rs0##args = Some a'
+ /\ Val.inject j a a').
+ eapply eval_addressing_inject; eauto.
+ eapply globalenv_inject_preserves_globals. eapply sep_proj2. eapply sep_proj2. eapply sep_proj2. eexact SEP.
+ eapply agree_reglist; eauto.
+ destruct H1 as [a' [A B]].
+
+ destruct ( Mem.loadv chunk m' a') as [v'|] eqn:Hloadv.
+ {
+ econstructor; split.
+ apply plus_one. apply exec_Mload with (a:=a') (v:=v'); eauto.
+ try (rewrite <- A; apply eval_addressing_preserved; auto; exact symbols_preserved).
+ econstructor; eauto with coqlib.
+ apply agree_regs_set_reg. rewrite transl_destroyed_by_load. apply agree_regs_undef_regs; auto. auto.
+ apply agree_locs_set_reg. apply agree_locs_undef_locs. auto. apply destroyed_by_load_caller_save. auto.
+ }
+ {
+ econstructor; split.
+ apply plus_one. apply exec_Mload_notrap2 with (a:=a'); eauto.
+ try (rewrite <- A; apply eval_addressing_preserved; auto; exact symbols_preserved).
+
+ econstructor; eauto with coqlib.
+ apply agree_regs_set_reg. rewrite transl_destroyed_by_load. apply agree_regs_undef_regs; auto. auto.
+ apply agree_locs_set_reg. apply agree_locs_undef_locs. auto. apply destroyed_by_load_caller_save. auto.
+ }
+
- (* Lstore *)
assert (exists a',
eval_addressing ge (Vptr sp' Ptrofs.zero) (transl_addr (make_env (function_bounds f)) addr) rs0##args = Some a'
diff --git a/backend/Tailcallproof.v b/backend/Tailcallproof.v
index 06e314f3..5cb1b980 100644
--- a/backend/Tailcallproof.v
+++ b/backend/Tailcallproof.v
@@ -438,6 +438,43 @@ Proof.
apply eval_addressing_preserved. exact symbols_preserved. eauto.
econstructor; eauto. apply set_reg_lessdef; auto.
+- (* load notrap1 *)
+ TransfInstr.
+ assert (Val.lessdef_list (rs##args) (rs'##args)). apply regs_lessdef_regs; auto.
+ left.
+ exists (State s' (transf_function f) (Vptr sp0 Ptrofs.zero) pc' (rs'#dst <- (default_notrap_load_value chunk)) m'); split.
+ eapply exec_Iload_notrap1.
+ eassumption.
+ eapply eval_addressing_lessdef_none. eassumption.
+ erewrite eval_addressing_preserved.
+ eassumption. exact symbols_preserved.
+
+ econstructor; eauto. apply set_reg_lessdef; auto.
+
+- (* load notrap2 *)
+ TransfInstr.
+ assert (Val.lessdef_list (rs##args) (rs'##args)). apply regs_lessdef_regs; auto.
+ left.
+
+ exploit eval_addressing_lessdef; eauto.
+ intros [a' [ADDR' ALD]].
+
+ destruct (Mem.loadv chunk m' a') eqn:Echunk2.
+ + exists (State s' (transf_function f) (Vptr sp0 Ptrofs.zero) pc' (rs'#dst <- v) m'); split.
+ eapply exec_Iload with (a:=a'). eassumption.
+ erewrite eval_addressing_preserved.
+ eassumption.
+ exact symbols_preserved.
+ assumption.
+ econstructor; eauto. apply set_reg_lessdef; auto.
+ + exists (State s' (transf_function f) (Vptr sp0 Ptrofs.zero) pc' (rs'#dst <- (default_notrap_load_value chunk)) m'); split.
+ eapply exec_Iload_notrap2. eassumption.
+ erewrite eval_addressing_preserved.
+ eassumption.
+ exact symbols_preserved.
+ assumption.
+ econstructor; eauto. apply set_reg_lessdef; auto.
+
- (* store *)
TransfInstr.
assert (Val.lessdef_list (rs##args) (rs'##args)). apply regs_lessdef_regs; auto.
diff --git a/backend/Tunnelingproof.v b/backend/Tunnelingproof.v
index 4f95ac9b..d3b8a9f0 100644
--- a/backend/Tunnelingproof.v
+++ b/backend/Tunnelingproof.v
@@ -441,6 +441,31 @@ Proof.
rewrite <- EV. apply eval_addressing_preserved. exact symbols_preserved.
eauto. eauto.
econstructor; eauto using locmap_set_lessdef, locmap_undef_regs_lessdef.
+- (* Lload notrap1 *)
+ exploit eval_addressing_lessdef_none. apply reglist_lessdef; eauto. eassumption.
+ left; simpl; econstructor; split.
+ eapply exec_Lload_notrap1.
+ rewrite <- H0.
+ apply eval_addressing_preserved. exact symbols_preserved. eauto.
+ econstructor; eauto using locmap_set_lessdef, locmap_undef_regs_lessdef.
+- (* Lload notrap2 *)
+ exploit eval_addressing_lessdef. apply reglist_lessdef; eauto. eauto.
+ intros (ta & EV & LD).
+ destruct (Mem.loadv chunk tm ta) eqn:Htload.
+ {
+ left; simpl; econstructor; split.
+ eapply exec_Lload.
+ rewrite <- EV. apply eval_addressing_preserved. exact symbols_preserved.
+ exact Htload. eauto.
+ econstructor; eauto using locmap_set_lessdef, locmap_undef_regs_lessdef.
+ }
+ {
+ left; simpl; econstructor; split.
+ eapply exec_Lload_notrap2.
+ rewrite <- EV. apply eval_addressing_preserved. exact symbols_preserved.
+ exact Htload. eauto.
+ econstructor; eauto using locmap_set_lessdef, locmap_undef_regs_lessdef.
+ }
- (* Lgetstack *)
left; simpl; econstructor; split.
econstructor; eauto.
diff --git a/backend/Unusedglob.v b/backend/Unusedglob.v
index 8ac7c4ce..1b5f2547 100644
--- a/backend/Unusedglob.v
+++ b/backend/Unusedglob.v
@@ -46,7 +46,7 @@ Definition ref_instruction (i: instruction) : list ident :=
match i with
| Inop _ => nil
| Iop op _ _ _ => globals_operation op
- | Iload _ addr _ _ _ => globals_addressing addr
+ | Iload _ _ addr _ _ _ => globals_addressing addr
| Istore _ addr _ _ _ => globals_addressing addr
| Icall _ (inl r) _ _ _ => nil
| Icall _ (inr id) _ _ _ => id :: nil
diff --git a/backend/Unusedglobproof.v b/backend/Unusedglobproof.v
index 680daba7..fa120b6d 100644
--- a/backend/Unusedglobproof.v
+++ b/backend/Unusedglobproof.v
@@ -915,7 +915,7 @@ Proof.
/\ Val.inject j a ta).
{ apply eval_addressing_inj with (ge1 := ge) (sp1 := Vptr sp0 Ptrofs.zero) (vl1 := rs##args).
intros. apply symbol_address_inject. eapply match_stacks_preserves_globals; eauto.
- apply KEPT. red. exists pc, (Iload chunk addr args dst pc'); auto.
+ apply KEPT. red. exists pc, (Iload trap chunk addr args dst pc'); auto.
econstructor; eauto.
apply regs_inject; auto.
assumption. }
@@ -924,6 +924,36 @@ Proof.
econstructor; split. eapply exec_Iload; eauto.
econstructor; eauto. apply set_reg_inject; auto.
+- (* load notrap1 *)
+ assert (eval_addressing tge (Vptr tsp Ptrofs.zero) addr trs##args = None).
+ { eapply eval_addressing_inj_none.
+ intros. apply symbol_address_inject. eapply match_stacks_preserves_globals; eauto.
+ apply KEPT. red. exists pc, (Iload NOTRAP chunk addr args dst pc'); auto.
+ econstructor; eauto.
+ rewrite Ptrofs.add_zero; reflexivity.
+ apply regs_inject; auto.
+ eassumption.
+ assumption. }
+
+ econstructor; split. eapply exec_Iload_notrap1; eauto.
+ econstructor; eauto. apply set_reg_inject; auto.
+
+- (* load notrap2 *)
+ assert (A: exists ta,
+ eval_addressing tge (Vptr tsp Ptrofs.zero) addr trs##args = Some ta
+ /\ Val.inject j a ta).
+ { apply eval_addressing_inj with (ge1 := ge) (sp1 := Vptr sp0 Ptrofs.zero) (vl1 := rs##args).
+ intros. apply symbol_address_inject. eapply match_stacks_preserves_globals; eauto.
+ apply KEPT. red. exists pc, (Iload NOTRAP chunk addr args dst pc'); auto.
+ econstructor; eauto.
+ apply regs_inject; auto.
+ assumption. }
+ destruct A as (ta & B & C).
+ destruct (Mem.loadv chunk tm ta) eqn:Echunk2.
+ + econstructor; split. eapply exec_Iload; eauto.
+ econstructor; eauto. apply set_reg_inject; auto.
+ + econstructor; split. eapply exec_Iload_notrap2; eauto.
+ econstructor; eauto. apply set_reg_inject; auto.
- (* store *)
assert (A: exists ta,
eval_addressing tge (Vptr tsp Ptrofs.zero) addr trs##args = Some ta
diff --git a/backend/ValueAnalysis.v b/backend/ValueAnalysis.v
index 2b233900..9a33768c 100644
--- a/backend/ValueAnalysis.v
+++ b/backend/ValueAnalysis.v
@@ -139,9 +139,14 @@ 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 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
+
+ (* TODO: maybe a case analysis on the results of loadv? *)
+
+ | 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 +1273,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/backend/XTL.ml b/backend/XTL.ml
index f10efeed..c496fafb 100644
--- a/backend/XTL.ml
+++ b/backend/XTL.ml
@@ -30,7 +30,7 @@ type instruction =
| Xspill of var * var
| Xparmove of var list * var list * var * var
| Xop of operation * var list * var
- | Xload of memory_chunk * addressing * var list * var
+ | Xload of trapping_mode * memory_chunk * addressing * var list * var
| Xstore of memory_chunk * addressing * var list * var
| Xcall of signature * (var, ident) sum * var list * var list
| Xtailcall of signature * (var, ident) sum * var list
@@ -159,7 +159,7 @@ let type_instr = function
let (targs, tres) = type_of_operation op in
set_vars_type args targs;
set_var_type res tres
- | Xload(chunk, addr, args, dst) ->
+ | Xload(trap, chunk, addr, args, dst) ->
set_vars_type args (type_of_addressing addr);
set_var_type dst (type_of_chunk chunk)
| Xstore(chunk, addr, args, src) ->
diff --git a/backend/XTL.mli b/backend/XTL.mli
index 54988d4b..b4b77fab 100644
--- a/backend/XTL.mli
+++ b/backend/XTL.mli
@@ -31,7 +31,7 @@ type instruction =
| Xspill of var * var
| Xparmove of var list * var list * var * var
| Xop of operation * var list * var
- | Xload of memory_chunk * addressing * var list * var
+ | Xload of trapping_mode * memory_chunk * addressing * var list * var
| Xstore of memory_chunk * addressing * var list * var
| Xcall of signature * (var, ident) sum * var list * var list
| Xtailcall of signature * (var, ident) sum * var list