From 7a5aef7aa4b669a722ecce3da6e7f7c646a392cc Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Sun, 17 Jun 2018 17:48:01 +0200 Subject: Treat Outgoing stack slots as caller-save in LTL/Linear semantics (#237) * Outgoing stack slots are set to Vundef on return from a function call, modeling the fact that the callee could write into those stack slots. (CompCert-generated code does not do this, but code generated by other compilers sometimes does.) * Adapt Stackingproof to this new semantics. This requires tighter reasoning on how Linear's locsets are related at call points and at return points. * Most of this reasoning was moved from Stackingproof to Lineartyping, because it can be expressed purely in terms of the Linear semantics, and tracked through the wt_state predicate. * Factor out and into Conventions.v: the notion of callee-save locations, the "agree_callee_save" predicate, and useful lemmas on Locmap.setpair. Now the same "agree_callee_save" predicate is used in Allocproof and in Stackingproof. --- backend/Allocproof.v | 18 +++------- backend/Conventions.v | 53 +++++++++++++++++++++++++++++ backend/LTL.v | 10 ++++-- backend/Lineartyping.v | 63 +++++++++++++++++++++++++++++++--- backend/Stackingproof.v | 88 +++++++++++------------------------------------- backend/Tunnelingproof.v | 6 ++-- 6 files changed, 145 insertions(+), 93 deletions(-) (limited to 'backend') diff --git a/backend/Allocproof.v b/backend/Allocproof.v index 068e0de0..1804f46b 100644 --- a/backend/Allocproof.v +++ b/backend/Allocproof.v @@ -1317,15 +1317,6 @@ Proof. eauto. Qed. -Definition callee_save_loc (l: loc) := - match l with - | R r => is_callee_save r = true - | S sl ofs ty => sl <> Outgoing - end. - -Definition agree_callee_save (ls1 ls2: locset) : Prop := - forall l, callee_save_loc l -> ls1 l = ls2 l. - Lemma return_regs_agree_callee_save: forall caller callee, agree_callee_save caller (return_regs caller callee). @@ -2476,11 +2467,10 @@ Proof. rewrite Locmap.gss. rewrite Locmap.gso by (red; auto). rewrite Locmap.gss. rewrite val_longofwords_eq_1 by auto. auto. red; intros. rewrite (AG l H0). - rewrite Locmap.gpo. - unfold undef_caller_save_regs. destruct l; auto. simpl in H0; rewrite H0; auto. - assert (X: forall r, is_callee_save r = false -> Loc.diff l (R r)). - { intros. destruct l; simpl in *. congruence. auto. } - generalize (loc_result_caller_save (ef_sig ef)). destruct (loc_result (ef_sig ef)); simpl; intuition auto. + rewrite locmap_get_set_loc_result_callee_save by auto. + unfold undef_caller_save_regs. destruct l; simpl in H0. + rewrite H0; auto. + destruct sl; auto; congruence. eapply external_call_well_typed; eauto. (* return *) diff --git a/backend/Conventions.v b/backend/Conventions.v index bdc4c8b6..989bfa05 100644 --- a/backend/Conventions.v +++ b/backend/Conventions.v @@ -103,3 +103,56 @@ Proof. generalize (loc_arguments_bounded _ _ _ H0). generalize (typesize_pos ty). omega. Qed. + + +(** * Callee-save locations *) + +(** We classify locations as either +- callee-save, i.e. preserved across function calls: + callee-save registers, [Local] and [Incoming] stack slots; +- caller-save, i.e. possibly modified by a function call: + non-callee-save registers, [Outgoing] stack slots. + +Concerning [Outgoing] stack slots: several ABIs allow a function to modify +the stack slots used for passing parameters to this function. +The code currently generated by CompCert never does so, but the code +generated by other compilers often does so (e.g. GCC for x86-32). +Hence, CompCert-generated code must not assume that [Outgoing] stack slots +are preserved across function calls, because they might not be preserved +if the called function was compiled by another compiler. +*) + +Definition callee_save_loc (l: loc) := + match l with + | R r => is_callee_save r = true + | S sl ofs ty => sl <> Outgoing + end. + +Hint Unfold callee_save_loc. + +Definition agree_callee_save (ls1 ls2: Locmap.t) : Prop := + forall l, callee_save_loc l -> ls1 l = ls2 l. + +(** * Assigning result locations *) + +(** Useful lemmas to reason about the result of an external call. *) + +Lemma locmap_get_set_loc_result: + forall sg v rs l, + match l with R r => is_callee_save r = true | S _ _ _ => True end -> + Locmap.setpair (loc_result sg) v rs l = rs l. +Proof. + intros. apply Locmap.gpo. + assert (X: forall r, is_callee_save r = false -> Loc.diff l (R r)). + { intros. destruct l; simpl. congruence. auto. } + generalize (loc_result_caller_save sg). destruct (loc_result sg); simpl; intuition auto. +Qed. + +Lemma locmap_get_set_loc_result_callee_save: + forall sg v rs l, + callee_save_loc l -> + Locmap.setpair (loc_result sg) v rs l = rs l. +Proof. + intros. apply locmap_get_set_loc_result. + red in H; destruct l; auto. +Qed. diff --git a/backend/LTL.v b/backend/LTL.v index 6dd1abd6..5e7eec8c 100644 --- a/backend/LTL.v +++ b/backend/LTL.v @@ -96,24 +96,28 @@ Definition call_regs (caller: locset) : locset := - Callee-save machine registers have the same values as in the caller before the call. - Caller-save machine registers have the same values as in the callee. -- Stack slots have the same values as in the caller. +- Local and Incoming stack slots have the same values as in the caller. +- Outgoing stack slots are set to Vundef to reflect the fact that they + may have been changed by the callee. *) Definition return_regs (caller callee: locset) : locset := fun (l: loc) => match l with | R r => if is_callee_save r then caller (R r) else callee (R r) + | S Outgoing ofs ty => Vundef | S sl ofs ty => caller (S sl ofs ty) end. (** [undef_caller_save_regs ls] models the effect of calling - an external function: caller-save registers can change unpredictably, - hence we set them to [Vundef]. *) + an external function: caller-save registers and outgoing locations + can change unpredictably, hence we set them to [Vundef]. *) Definition undef_caller_save_regs (ls: locset) : locset := fun (l: loc) => match l with | R r => if is_callee_save r then ls (R r) else Vundef + | S Outgoing ofs ty => Vundef | S sl ofs ty => ls (S sl ofs ty) end. diff --git a/backend/Lineartyping.v b/backend/Lineartyping.v index fd252dc6..fc163719 100644 --- a/backend/Lineartyping.v +++ b/backend/Lineartyping.v @@ -146,15 +146,18 @@ Lemma wt_return_regs: wt_locset caller -> wt_locset callee -> wt_locset (return_regs caller callee). Proof. intros; red; intros. - unfold return_regs. destruct l; auto. destruct (is_callee_save r); auto. + unfold return_regs. destruct l. +- destruct (is_callee_save r); auto. +- destruct sl; auto; red; auto. Qed. Lemma wt_undef_caller_save_regs: forall ls, wt_locset ls -> wt_locset (undef_caller_save_regs ls). Proof. intros; red; intros. unfold undef_caller_save_regs. - destruct l; auto. - destruct (is_callee_save r); auto; simpl; auto. + destruct l. + destruct (is_callee_save r); auto; simpl; auto. + destruct sl; auto; red; auto. Qed. Lemma wt_init: @@ -205,6 +208,24 @@ Proof. auto. Qed. +(** In addition to type preservation during evaluation, we also show + properties of the environment [ls] at call points and at return points. + These properties are used in the proof of the [Stacking] pass. + For call points, we have that the current environment [ls] and the + one from the top call stack agree on the [Outgoing] locations + used for parameter passing. *) + +Definition agree_outgoing_arguments (sg: signature) (ls pls: locset) : Prop := + forall ty ofs, + In (S Outgoing ofs ty) (regs_of_rpairs (loc_arguments sg)) -> + ls (S Outgoing ofs ty) = pls (S Outgoing ofs ty). + +(** For return points, we have that all [Outgoing] stack locations have + been set to [Vundef]. *) + +Definition outgoing_undef (ls: locset) : Prop := + forall ty ofs, ls (S Outgoing ofs ty) = Vundef. + (** Soundness of the type system *) Definition wt_fundef (fd: fundef) := @@ -241,11 +262,15 @@ Inductive wt_state: state -> Prop := | wt_call_state: forall s fd rs m (WTSTK: wt_callstack s) (WTFD: wt_fundef fd) - (WTRS: wt_locset rs), + (WTRS: wt_locset rs) + (AGCS: agree_callee_save rs (parent_locset s)) + (AGARGS: agree_outgoing_arguments (funsig fd) rs (parent_locset s)), wt_state (Callstate s fd rs m) | wt_return_state: forall s rs m (WTSTK: wt_callstack s) - (WTRS: wt_locset rs), + (WTRS: wt_locset rs) + (AGCS: agree_callee_save rs (parent_locset s)) + (UOUT: outgoing_undef rs), wt_state (Returnstate s rs m). (** Preservation of state typing by transitions *) @@ -315,11 +340,15 @@ Local Opaque mreg_type. simpl in *; InvBooleans. econstructor; eauto. econstructor; eauto. eapply wt_find_function; eauto. + red; simpl; auto. + red; simpl; auto. - (* tailcall *) simpl in *; InvBooleans. econstructor; eauto. eapply wt_find_function; eauto. apply wt_return_regs; auto. apply wt_parent_locset; auto. + red; simpl; intros. destruct l; simpl in *. rewrite H3; auto. destruct sl; auto; congruence. + red; simpl; intros. apply zero_size_arguments_tailcall_possible in H. apply H in H3. contradiction. - (* builtin *) simpl in *; InvBooleans. econstructor; eauto. @@ -342,6 +371,8 @@ Local Opaque mreg_type. simpl in *. InvBooleans. econstructor; eauto. apply wt_return_regs; auto. apply wt_parent_locset; auto. + red; simpl; intros. destruct l; simpl in *. rewrite H0; auto. destruct sl; auto; congruence. + red; simpl; intros. auto. - (* internal function *) simpl in WTFD. econstructor. eauto. eauto. eauto. @@ -350,6 +381,10 @@ Local Opaque mreg_type. econstructor. auto. apply wt_setpair. eapply external_call_well_typed; eauto. apply wt_undef_caller_save_regs; auto. + red; simpl; intros. destruct l; simpl in *. + rewrite locmap_get_set_loc_result by auto. simpl. rewrite H; auto. + rewrite locmap_get_set_loc_result by auto. simpl. destruct sl; auto; congruence. + red; simpl; intros. rewrite locmap_get_set_loc_result by auto. auto. - (* return *) inv WTSTK. econstructor; eauto. Qed. @@ -361,6 +396,8 @@ Proof. unfold ge0 in H1. exploit Genv.find_funct_ptr_inversion; eauto. intros [id IN]. eapply wt_prog; eauto. apply wt_init. + red; auto. + red; auto. Qed. End SOUNDNESS. @@ -406,3 +443,19 @@ Lemma wt_callstate_wt_regs: Proof. intros. inv H. apply WTRS. Qed. + +Lemma wt_callstate_agree: + forall s f rs m, + wt_state (Callstate s f rs m) -> + agree_callee_save rs (parent_locset s) /\ agree_outgoing_arguments (funsig f) rs (parent_locset s). +Proof. + intros. inv H; auto. +Qed. + +Lemma wt_returnstate_agree: + forall s rs m, + wt_state (Returnstate s rs m) -> + agree_callee_save rs (parent_locset s) /\ outgoing_undef rs. +Proof. + intros. inv H; auto. +Qed. diff --git a/backend/Stackingproof.v b/backend/Stackingproof.v index 5dbc4532..ffd9b227 100644 --- a/backend/Stackingproof.v +++ b/backend/Stackingproof.v @@ -294,12 +294,13 @@ Qed. Lemma contains_locations_exten: forall ls ls' j sp pos bound sl, - (forall ofs ty, ls' (S sl ofs ty) = ls (S sl ofs ty)) -> + (forall ofs ty, Val.lessdef (ls' (S sl ofs ty)) (ls (S sl ofs ty))) -> massert_imp (contains_locations j sp pos bound sl ls) (contains_locations j sp pos bound sl ls'). Proof. intros; split; simpl; intros; auto. - intuition auto. rewrite H. eauto. + intuition auto. exploit H5; eauto. intros (v & A & B). exists v; split; auto. + specialize (H ofs ty). inv H. congruence. auto. Qed. Lemma contains_locations_incr: @@ -481,7 +482,8 @@ Qed. Lemma frame_contents_exten: forall ls ls0 ls' ls0' j sp parent retaddr P m, - (forall sl ofs ty, ls' (S sl ofs ty) = ls (S sl ofs ty)) -> + (forall ofs ty, Val.lessdef (ls' (S Local ofs ty)) (ls (S Local ofs ty))) -> + (forall ofs ty, Val.lessdef (ls' (S Outgoing ofs ty)) (ls (S Outgoing ofs ty))) -> (forall r, In r b.(used_callee_save) -> ls0' (R r) = ls0 (R r)) -> m |= frame_contents j sp ls ls0 parent retaddr ** P -> m |= frame_contents j sp ls' ls0' parent retaddr ** P. @@ -573,16 +575,6 @@ Record agree_locs (ls ls0: locset) : Prop := ls (S Incoming ofs ty) = ls0 (S Outgoing ofs ty) }. -(** Auxiliary predicate used at call points *) - -Definition agree_callee_save (ls ls0: locset) : Prop := - forall l, - match l with - | R r => is_callee_save r = true - | S _ _ _ => True - end -> - ls l = ls0 l. - (** ** Properties of [agree_regs]. *) (** Values of registers *) @@ -810,51 +802,7 @@ Lemma agree_locs_return: Proof. intros. red in H0. inv H; constructor; auto; intros. - rewrite H0; auto. unfold mreg_within_bounds in H. tauto. -- rewrite H0; auto. -Qed. - -(** Preservation at tailcalls (when [ls0] is changed but not [ls]). *) - -Lemma agree_locs_tailcall: - forall ls ls0 ls0', - agree_locs ls ls0 -> - agree_callee_save ls0 ls0' -> - agree_locs ls ls0'. -Proof. - intros. red in H0. inv H; constructor; auto; intros. -- rewrite <- H0; auto. unfold mreg_within_bounds in H. tauto. -- rewrite <- H0; auto. -Qed. - -(** ** Properties of [agree_callee_save]. *) - -Lemma agree_callee_save_return_regs: - forall ls1 ls2, - agree_callee_save (return_regs ls1 ls2) ls1. -Proof. - intros; red; intros. - unfold return_regs. destruct l; auto. rewrite H; auto. -Qed. - -Lemma agree_callee_save_undef_caller_save_regs: - forall ls1 ls2, - agree_callee_save ls1 ls2 -> - agree_callee_save (LTL.undef_caller_save_regs ls1) ls2. -Proof. - intros; red; intros. unfold LTL.undef_caller_save_regs. - destruct l; auto. - rewrite H0; auto. -Qed. - -Lemma agree_callee_save_set_result: - forall sg v ls1 ls2, - agree_callee_save ls1 ls2 -> - agree_callee_save (Locmap.setpair (loc_result sg) v ls1) ls2. -Proof. - intros; red; intros. rewrite Locmap.gpo. apply H; auto. - assert (X: forall r, is_callee_save r = false -> Loc.diff l (R r)). - { intros. destruct l; auto. simpl; congruence. } - generalize (loc_result_caller_save sg). destruct (loc_result sg); simpl; intuition auto. +- rewrite <- agree_incoming0 by auto. apply H0. congruence. Qed. (** ** Properties of destroyed registers. *) @@ -1091,6 +1039,7 @@ Lemma function_prologue_correct: forall j ls ls0 ls1 rs rs1 m1 m1' m2 sp parent ra cs fb k P, agree_regs j ls rs -> agree_callee_save ls ls0 -> + agree_outgoing_arguments (Linear.fn_sig f) ls ls0 -> (forall r, Val.has_type (ls (R r)) (mreg_type r)) -> ls1 = LTL.undef_regs destroyed_at_function_entry (LTL.call_regs ls) -> rs1 = undef_regs destroyed_at_function_entry rs -> @@ -1110,7 +1059,7 @@ Lemma function_prologue_correct: /\ j' sp = Some(sp', fe.(fe_stack_data)) /\ inject_incr j j'. Proof. - intros until P; intros AGREGS AGCS WTREGS LS1 RS1 ALLOC TYPAR TYRA SEP. + intros until P; intros AGREGS AGCS AGARGS WTREGS LS1 RS1 ALLOC TYPAR TYRA SEP. rewrite unfold_transf_function. unfold fn_stacksize, fn_link_ofs, fn_retaddr_ofs. (* Stack layout info *) @@ -1194,7 +1143,7 @@ Local Opaque b fe. split. rewrite LS1. apply agree_locs_undef_locs; [|reflexivity]. constructor; intros. unfold call_regs. apply AGCS. unfold mreg_within_bounds in H; tauto. - unfold call_regs. apply AGCS. auto. + unfold call_regs. apply AGARGS. apply incoming_slot_in_parameters; auto. split. exact SEPFINAL. split. exact SAME. exact INCR. Qed. @@ -1345,7 +1294,7 @@ Proof. apply CS; auto. rewrite NCS by auto. apply AGR. split. red; unfold return_regs; intros. - destruct l; auto. rewrite H; auto. + destruct l. rewrite H; auto. destruct sl; auto; contradiction. assumption. Qed. @@ -1639,6 +1588,7 @@ Variable ls: locset. Variable rs: regset. Hypothesis AGR: agree_regs j ls rs. Hypothesis AGCS: agree_callee_save ls (parent_locset cs). +Hypothesis AGARGS: agree_outgoing_arguments sg ls (parent_locset cs). Variable m': mem. Hypothesis SEP: m' |= stack_contents j cs cs'. @@ -1661,7 +1611,7 @@ Proof. assert (slot_within_bounds (function_bounds f) Outgoing pos ty) by eauto. exploit frame_get_outgoing; eauto. intros (v & A & B). exists v; split. - constructor. exact A. red in AGCS. rewrite AGCS; auto. + constructor. exact A. rewrite AGARGS by auto. exact B. Qed. Lemma transl_external_argument_2: @@ -1836,7 +1786,6 @@ Inductive match_states: Linear.state -> Mach.state -> Prop := (TRANSL: transf_fundef f = OK tf) (FIND: Genv.find_funct_ptr tge fb = Some tf) (AGREGS: agree_regs j ls rs) - (AGLOCS: agree_callee_save ls (parent_locset cs)) (SEP: m' |= stack_contents j cs cs' ** minjection j m ** globalenv_inject ge j), @@ -1846,7 +1795,6 @@ Inductive match_states: Linear.state -> Mach.state -> Prop := forall cs ls m cs' rs m' j sg (STACKS: match_stacks j cs cs' sg) (AGREGS: agree_regs j ls rs) - (AGLOCS: agree_callee_save ls (parent_locset cs)) (SEP: m' |= stack_contents j cs cs' ** minjection j m ** globalenv_inject ge j), @@ -2004,9 +1952,8 @@ Proof. econstructor; eauto with coqlib. apply Val.Vptr_has_type. intros; red. - apply Z.le_trans with (size_arguments (Linear.funsig f')); auto. + apply Z.le_trans with (size_arguments (Linear.funsig f')); auto. apply loc_arguments_bounded; auto. - simpl; red; auto. simpl. rewrite sep_assoc. exact SEP. - (* Ltailcall *) @@ -2106,6 +2053,7 @@ Proof. destruct (transf_function f) as [tfn|] eqn:TRANSL; simpl; try congruence. intros EQ; inversion EQ; clear EQ; subst tf. rewrite sep_comm, sep_assoc in SEP. + exploit wt_callstate_agree; eauto. intros [AGCS AGARGS]. exploit function_prologue_correct; eauto. red; intros; eapply wt_callstate_wt_regs; eauto. eapply match_stacks_type_sp; eauto. @@ -2124,6 +2072,7 @@ Proof. - (* external function *) simpl in TRANSL. inversion TRANSL; subst tf. + exploit wt_callstate_agree; eauto. intros [AGCS AGARGS]. exploit transl_external_arguments; eauto. apply sep_proj1 in SEP; eauto. intros [vl [ARGS VINJ]]. rewrite sep_comm, sep_assoc in SEP. exploit external_call_parallel_rule; eauto. @@ -2136,17 +2085,19 @@ Proof. apply agree_regs_set_pair. apply agree_regs_undef_caller_save_regs. apply agree_regs_inject_incr with j; auto. auto. - apply agree_callee_save_set_result. apply agree_callee_save_undef_caller_save_regs. auto. apply stack_contents_change_meminj with j; auto. rewrite sep_comm, sep_assoc; auto. - (* return *) - inv STACKS. simpl in AGLOCS. simpl in SEP. rewrite sep_assoc in SEP. + inv STACKS. exploit wt_returnstate_agree; eauto. intros [AGCS OUTU]. + simpl in AGCS. simpl in SEP. rewrite sep_assoc in SEP. econstructor; split. apply plus_one. apply exec_return. econstructor; eauto. apply agree_locs_return with rs0; auto. apply frame_contents_exten with rs0 (parent_locset s); auto. + intros; apply Val.lessdef_same; apply AGCS; red; congruence. + intros; rewrite (OUTU ty ofs); auto. Qed. Lemma transf_initial_states: @@ -2164,7 +2115,6 @@ Proof. eapply match_states_call with (j := j); eauto. constructor. red; intros. rewrite H3, loc_arguments_main in H. contradiction. red; simpl; auto. - red; simpl; auto. simpl. rewrite sep_pure. split; auto. split;[|split]. eapply Genv.initmem_inject; eauto. simpl. exists (Mem.nextblock m0); split. apply Ple_refl. diff --git a/backend/Tunnelingproof.v b/backend/Tunnelingproof.v index fd97ce33..4f95ac9b 100644 --- a/backend/Tunnelingproof.v +++ b/backend/Tunnelingproof.v @@ -339,7 +339,9 @@ Lemma locmap_undef_caller_save_regs_lessdef: locmap_lessdef ls1 ls2 -> locmap_lessdef (undef_caller_save_regs ls1) (undef_caller_save_regs ls2). Proof. intros; red; intros. unfold undef_caller_save_regs. - destruct l; auto. destruct (Conventions1.is_callee_save r); auto. + destruct l. +- destruct (Conventions1.is_callee_save r); auto. +- destruct sl; auto. Qed. Lemma find_function_translated: @@ -371,7 +373,7 @@ Lemma return_regs_lessdef: Proof. intros; red; intros. destruct l; simpl. - destruct (Conventions1.is_callee_save r); auto. -- auto. +- destruct sl; auto. Qed. (** To preserve non-terminating behaviours, we show that the transformed -- cgit