(* *********************************************************************)
(* *)
(* The Compcert verified compiler *)
(* *)
(* Xavier Leroy, INRIA Paris-Rocquencourt *)
(* *)
(* Copyright Institut National de Recherche en Informatique et en *)
(* Automatique. All rights reserved. This file is distributed *)
(* under the terms of the INRIA Non-Commercial License Agreement. *)
(* *)
(* *********************************************************************)
(** Correctness of instruction selection *)
Require Import FunInd. Require Import Coqlib Maps. Require Import AST Linking Errors Integers. Require Import Values Memory Builtins Events Globalenvs Smallstep. Require Import Switch Cminor Op CminorSel Cminortyping. Require Import SelectOp SelectDiv SplitLong SelectLong Selection. Require Import SelectOpproof SelectDivproof SplitLongproof SelectLongproof. Local Open Scope cminorsel_scope. Local Open Scope error_monad_scope. (** * Relational specification of instruction selection *)
Definition match_fundef (cunit: Cminor.program) (f: Cminor.fundef) (tf: CminorSel.fundef) : Prop :=
exists hf, helper_functions_declared cunit hf /\ sel_fundef (prog_defmap cunit) hf f = OK tf. Definition match_prog (p: Cminor.program) (tp: CminorSel.program) :=
match_program match_fundef eq p tp.
Proof. Admitted.
Proof. intros. set (f := fun m id gd => if globdef_of_interest gd then PTree.set id gd m else m) in *. set (P := fun m m' => m'!id = Some gd -> m!id = Some gd). assert (X: P dm (PTree.fold f dm (PTree.empty _))). apply PTree_Properties.fold_rec. - unfold P; intros. rewrite <- H0; auto. - red. rewrite ! PTree.gempty. auto. - unfold P; intros. rewrite PTree.gsspec. unfold f in H3. destruct (globdef_of_interest v). + rewrite PTree.gsspec in H3. destruct (peq id k); auto. + apply H2 in H3. destruct (peq id k). congruence. auto. apply X. auto. Qed.
Proof. Admitted.
Proof. intros. set (P := fun (m: PTree.t globdef) res => res = Some id -> m!id = Some(Gfun(External (EF_runtime name sg)))). assert (P globs (PTree.fold (lookup_helper_aux name sg) globs None)). apply PTree_Properties.fold_rec; red; intros. - rewrite <- H0. apply H1; auto. - discriminate. - assert (EITHER: k = id /\ v = Gfun (External (EF_runtime name sg))
\/ a = Some id). unfold lookup_helper_aux in H3. destruct v; auto. destruct f; auto. destruct e; auto. destruct (String.string_dec name name0); auto. destruct (signature_eq sg sg0); auto. inversion H3. left; split; auto. repeat f_equal; auto. destruct EITHER as [[X Y] | X]. subst k v. apply PTree.gss. apply H2 in X. rewrite PTree.gso by congruence. auto. red in H0. unfold lookup_helper in H. destruct (PTree.fold (lookup_helper_aux name sg) globs None); inv H. auto. Qed.
Proof. intros. assert (X: forall id name sg, helper_declared p id name sg -> helper_declared p' id name sg). unfold helper_declared; intros. destruct (prog_defmap_linkorder _ _ _ _ H0 H1) as (gd & P & Q). inv Q. inv H3. auto. red in H. decompose [Logic.and] H; clear H. red; auto 20. Qed.
(** * Correctness of the instruction selection functions for expressions *)
Section PRESERVATION. Variable prog: Cminor.program. Variable tprog: CminorSel.program. Let ge := Genv.globalenv prog. Let tge := Genv.globalenv tprog. Hypothesis TRANSF: match_prog prog tprog.
Proof. Admitted.
Proof. red; intros. destruct TRANSF as [A _]. exploit list_forall2_in_left; eauto. intros ((i' & gd') & B & (C & D)). simpl in *. inv D. destruct H2 as (hf & P & Q). destruct f; monadInv Q. - monadInv EQ. econstructor; apply type_function_sound; eauto. - constructor. Qed.
Proof. intros. destruct H as (hf & P & Q). destruct f; monadInv Q; auto. monadInv EQ; auto. Qed.
Proof. intros. monadInv H. simpl. eauto. Qed.
Proof. intros. monadInv H. auto. Qed.
Proof. Admitted.
Proof. assert (X: forall id name sg, helper_declared prog id name sg -> helper_declared tprog id name sg). unfold helper_declared; intros. generalize (match_program_defmap _ _ _ _ _ TRANSF id). unfold Cminor.fundef; rewrite H; intros R; inv R. inv H2. destruct H4 as (cu & A & B). monadInv B. auto. unfold helper_functions_declared; intros. decompose [Logic.and] H; clear H. auto 20. Qed.
Proof. unfold classify_call; intros. destruct (expr_is_addrof_ident a) as [id|] eqn:EA; auto. exploit expr_is_addrof_ident_correct; eauto. intros EQ; subst a. inv H0. inv H3. unfold Genv.symbol_address in *. destruct (Genv.find_symbol ge id) as [b|] eqn:FS; try discriminate. rewrite Genv.find_funct_find_funct_ptr in H1. assert (DFL: exists b1, Genv.find_symbol ge id = Some b1 /\ Vptr b Ptrofs.zero = Vptr b1 Ptrofs.zero) by (exists b; auto). unfold globdef; destruct (prog_defmap unit)!id as [[[f|ef] |gv] |] eqn:G; auto. destruct (ef_inline ef) eqn:INLINE; auto. destruct (prog_defmap_linkorder _ _ _ _ H G) as (gd & P & Q). inv Q. inv H2. - apply Genv.find_def_symbol in P. destruct P as (b' & X & Y). fold ge in X, Y. rewrite <- Genv.find_funct_ptr_iff in Y. congruence. - simpl in INLINE. discriminate. Qed.
(** Translation of [switch] statements *)
Inductive Rint: Z -> val -> Prop :=
| Rint_intro: forall n, Rint (Int.unsigned n) (Vint n). Inductive Rlong: Z -> val -> Prop :=
| Rlong_intro: forall n, Rlong (Int64.unsigned n) (Vlong n). Section SEL_SWITCH. Variable make_cmp_eq: expr -> Z -> expr. Variable make_cmp_ltu: expr -> Z -> expr. Variable make_sub: expr -> Z -> expr. Variable make_to_int: expr -> expr. Variable modulus: Z. Variable R: Z -> val -> Prop. Hypothesis eval_make_cmp_eq: forall sp e m le a v i n,
eval_expr tge sp e m le a v -> R i v -> 0 <= n < modulus ->
eval_expr tge sp e m le (make_cmp_eq a n) (Val.of_bool (zeq i n)). Hypothesis eval_make_cmp_ltu: forall sp e m le a v i n,
eval_expr tge sp e m le a v -> R i v -> 0 <= n < modulus ->
eval_expr tge sp e m le (make_cmp_ltu a n) (Val.of_bool (zlt i n)). Hypothesis eval_make_sub: forall sp e m le a v i n,
eval_expr tge sp e m le a v -> R i v -> 0 <= n < modulus ->
exists v', eval_expr tge sp e m le (make_sub a n) v' /\ R ((i - n) mod modulus) v'. Hypothesis eval_make_to_int: forall sp e m le a v i,
eval_expr tge sp e m le a v -> R i v ->
exists v', eval_expr tge sp e m le (make_to_int a) v' /\ Rint (i mod Int.modulus) v'.
Proof. Admitted.
Proof. intros until x; intros Ri. induction t; simpl; intros until le; intros WF ARG MATCH. - (* base case *)
inv MATCH. constructor. - (* eq test *)
inv WF. assert (eval_expr tge sp e m le (make_cmp_eq (Eletvar arg) key) (Val.of_bool (zeq i key))). eapply eval_make_cmp_eq; eauto. constructor; auto. eapply eval_XEcondition with (va := zeq i key). eapply eval_condexpr_of_expr; eauto. destruct (zeq i key); constructor; auto. destruct (zeq i key); simpl. + inv MATCH. constructor. + eapply IHt; eauto. - (* lt test *)
inv WF. assert (eval_expr tge sp e m le (make_cmp_ltu (Eletvar arg) key) (Val.of_bool (zlt i key))). eapply eval_make_cmp_ltu; eauto. constructor; auto. eapply eval_XEcondition with (va := zlt i key). eapply eval_condexpr_of_expr; eauto. destruct (zlt i key); constructor; auto. destruct (zlt i key); simpl. + eapply IHt1; eauto. + eapply IHt2; eauto. - (* jump table *)
inv WF. exploit (eval_make_sub sp e m le). eapply eval_Eletvar. eauto. eauto. instantiate (1 := ofs). auto. intros (v' & A & B). set (i' := (i - ofs) mod modulus) in *. assert (eval_expr tge sp e m (v' :: le) (make_cmp_ltu (Eletvar O) sz) (Val.of_bool (zlt i' sz))). eapply eval_make_cmp_ltu; eauto. constructor; auto. econstructor. eauto. eapply eval_XEcondition with (va := zlt i' sz). eapply eval_condexpr_of_expr; eauto. destruct (zlt i' sz); constructor; auto. destruct (zlt i' sz); simpl. + exploit (eval_make_to_int sp e m (v' :: le) (Eletvar O)). constructor. simpl; eauto. eauto. intros (v'' & C & D). inv D. econstructor; eauto. congruence. + eapply IHt; eauto. Qed.
Proof. assert (INTCONST: forall n sp e m le,
eval_expr tge sp e m le (Eop (Ointconst n) Enil) (Vint n)). intros. econstructor. constructor. auto. intros. eapply sel_switch_correct with (R := Rint); eauto. - intros until n; intros EVAL R RANGE. exploit eval_comp. eexact EVAL. apply (INTCONST (Int.repr n)). instantiate (1 := Ceq). intros (vb & A & B). inv R. unfold Val.cmp in B. simpl in B. revert B. predSpec Int.eq Int.eq_spec n0 (Int.repr n); intros B; inv B. rewrite Int.unsigned_repr. unfold proj_sumbool; rewrite zeq_true; auto. unfold Int.max_unsigned; omega. unfold proj_sumbool; rewrite zeq_false; auto. red; intros; elim H1. rewrite <- (Int.repr_unsigned n0). congruence. - intros until n; intros EVAL R RANGE. exploit eval_compu. eexact EVAL. apply (INTCONST (Int.repr n)). instantiate (1 := Clt). intros (vb & A & B). inv R. unfold Val.cmpu in B. simpl in B. unfold Int.ltu in B. rewrite Int.unsigned_repr in B. destruct (zlt (Int.unsigned n0) n); inv B; auto. unfold Int.max_unsigned; omega. - intros until n; intros EVAL R RANGE. exploit eval_sub. eexact EVAL. apply (INTCONST (Int.repr n)). intros (vb & A & B). inv R. simpl in B. inv B. econstructor; split; eauto. replace ((Int.unsigned n0 - n) mod Int.modulus)
with (Int.unsigned (Int.sub n0 (Int.repr n))). constructor. unfold Int.sub. rewrite Int.unsigned_repr_eq. f_equal. f_equal. apply Int.unsigned_repr. unfold Int.max_unsigned; omega. - intros until i0; intros EVAL R. exists v; split; auto. inv R. rewrite Z.mod_small by (apply Int.unsigned_range). constructor. - constructor. - apply Int.unsigned_range. Qed.
Proof. Admitted.
Proof. intros. eapply sel_switch_correct with (R := Rlong); eauto. - intros until n; intros EVAL R RANGE. eapply eval_cmpl. eexact EVAL. apply eval_longconst with (n := Int64.repr n). inv R. unfold Val.cmpl. simpl. f_equal; f_equal. unfold Int64.eq. rewrite Int64.unsigned_repr. destruct (zeq (Int64.unsigned n0) n); auto. unfold Int64.max_unsigned; omega. - intros until n; intros EVAL R RANGE. eapply eval_cmplu; auto. eexact EVAL. apply eval_longconst with (n := Int64.repr n). inv R. unfold Val.cmplu. simpl. f_equal; f_equal. unfold Int64.ltu. rewrite Int64.unsigned_repr. destruct (zlt (Int64.unsigned n0) n); auto. unfold Int64.max_unsigned; omega. - intros until n; intros EVAL R RANGE. exploit eval_subl; auto; try apply HF'. eexact EVAL. apply eval_longconst with (n := Int64.repr n). intros (vb & A & B). inv R. simpl in B. inv B. econstructor; split; eauto. replace ((Int64.unsigned n0 - n) mod Int64.modulus)
with (Int64.unsigned (Int64.sub n0 (Int64.repr n))). constructor. unfold Int64.sub. rewrite Int64.unsigned_repr_eq. f_equal. f_equal. apply Int64.unsigned_repr. unfold Int64.max_unsigned; omega. - intros until i0; intros EVAL R. exploit eval_lowlong. eexact EVAL. intros (vb & A & B). inv R. simpl in B. inv B. econstructor; split; eauto. replace (Int64.unsigned n mod Int.modulus) with (Int.unsigned (Int64.loword n)). constructor. unfold Int64.loword. apply Int.unsigned_repr_eq. - constructor. - apply Int64.unsigned_range. Qed.
End SEL_SWITCH_INT. (** Compatibility of evaluation functions with the "less defined than" relation. *)
Ltac TrivialExists :=
match goal with
| [ |- exists v, Some ?x = Some v /\ _ ] => exists x; split; auto
| _ => idtac
end.
Proof. Admitted.
Proof. intros until v; intros EV LD. inv LD. exists v; auto. destruct op; simpl in *; inv EV; TrivialExists. Qed.
Proof. Admitted.
Proof. intros until m'; intros EV LD1 LD2 ME. assert (exists v', eval_binop op v1' v2' m = Some v' /\ Val.lessdef v v'). inv LD1. inv LD2. exists v; auto. destruct op; destruct v1'; simpl in *; inv EV; TrivialExists. destruct op; simpl in *; inv EV; TrivialExists. assert (CMPU: forall c,
eval_binop (Ocmpu c) v1 v2 m = Some v ->
exists v' : val, eval_binop (Ocmpu c) v1' v2' m' = Some v' /\ Val.lessdef v v'). intros c A. simpl in *. inv A. econstructor; split. eauto. apply Val.of_optbool_lessdef. intros. apply Val.cmpu_bool_lessdef with (Mem.valid_pointer m) v1 v2; auto. intros; eapply Mem.valid_pointer_extends; eauto. assert (CMPLU: forall c,
eval_binop (Ocmplu c) v1 v2 m = Some v ->
exists v' : val, eval_binop (Ocmplu c) v1' v2' m' = Some v' /\ Val.lessdef v v'). intros c A. simpl in *. unfold Val.cmplu in *. destruct (Val.cmplu_bool (Mem.valid_pointer m) c v1 v2) as [b|] eqn:C; simpl in A; inv A. eapply Val.cmplu_bool_lessdef with (valid_ptr' := (Mem.valid_pointer m')) in C;
eauto using Mem.valid_pointer_extends. rewrite C. exists (Val.of_bool b); auto. destruct op; auto. Qed.
(** * Semantic preservation for instruction selection. *)
(** Relationship between the local environments. *)
Definition env_lessdef (e1 e2: env) : Prop :=
forall id v1, e1!id = Some v1 -> exists v2, e2!id = Some v2 /\ Val.lessdef v1 v2.
Proof. intros. unfold sel_builtin_default. exploit sel_builtin_args_correct; eauto. intros (vl' & A & B). exploit external_call_mem_extends; eauto. intros (v' & m2' & D & E & F & _). econstructor; exists m2'; split. econstructor. eexact A. eapply external_call_symbols_preserved. eexact senv_preserved. eexact D. split; auto. apply sel_builtin_res_correct; auto. Qed.
Proof. Admitted.
Proof. intros. exploit sel_exprlist_correct; eauto. intros (vl' & A & B). exploit external_call_mem_extends; eauto. intros (v' & m2' & D & E & F & _). unfold sel_builtin. destruct optid as [id|]; eauto using sel_builtin_default_correct. destruct ef; eauto using sel_builtin_default_correct. destruct (lookup_builtin_function name sg) as [bf|] eqn:LKUP; eauto using sel_builtin_default_correct. destruct (sel_known_builtin bf (sel_exprlist al)) as [a|] eqn:SKB; eauto using sel_builtin_default_correct. simpl in D. red in D. rewrite LKUP in D. inv D. exploit eval_sel_known_builtin; eauto. intros (v'' & U & V). econstructor; exists m2'; split. econstructor. eexact U. split; auto. apply set_var_lessdef; auto. apply Val.lessdef_trans with v'; auto. Qed.
Proof. intros. destruct (sel_select_opt ty cond a1 a2) as [a'|] eqn:SSO; simpl in H1; inv H1. destruct (eval_safe_expr ge f sp e m a1) as (v1 & EV1); auto. destruct (eval_safe_expr ge f sp e m a2) as (v2 & EV2); auto. assert (TY1: Val.has_type v1 ty) by (eapply wt_eval_expr; eauto). assert (TY2: Val.has_type v2 ty) by (eapply wt_eval_expr; eauto). exploit sel_select_opt_correct; eauto. intros (v' & EV' & LD). exists a', v1, v2, v'; intuition eauto. apply Val.lessdef_trans with (Val.select (Some b) v1 v2 ty). simpl. rewrite Val.normalize_idem; auto. destruct b; auto. assumption. Qed.
Proof. Admitted.
Proof. unfold if_conversion; intros until m'; intros IFC DE WTE WT1 WT2 EVC BOV ELD MEXT. set (s0 := if b then ifso else ifnot). set (ki := known_id f) in *. destruct (classify_stmt ifso) eqn:IFSO; try discriminate;
destruct (classify_stmt ifnot) eqn:IFNOT; try discriminate;
unfold if_conversion_base in IFC. - destruct (is_known ki id && safe_expr ki (Cminor.Evar id) && safe_expr ki a
&& if_conversion_heuristic cond (Cminor.Evar id) a (env id)) eqn:B; inv IFC. InvBooleans. exploit (eval_select_safe_exprs (Cminor.Evar id) a); eauto. constructor. eapply classify_stmt_wt; eauto. intros (a' & v1 & v2 & v' & A & B & C & D & E). exists (PTree.set id (if b then v1 else v2) e), (PTree.set id v' e'). split. subst s. constructor; auto. split. unfold s0; destruct b. rewrite PTree.gsident by (inv B; auto). apply classify_stmt_sound_1; auto. eapply classify_stmt_sound_2; eauto. apply set_var_lessdef; auto. - destruct (is_known ki id && safe_expr ki a && safe_expr ki (Cminor.Evar id)
&& if_conversion_heuristic cond a (Cminor.Evar id) (env id)) eqn:B; inv IFC. InvBooleans. exploit (eval_select_safe_exprs a (Cminor.Evar id)); eauto. eapply classify_stmt_wt; eauto. constructor. intros (a' & v1 & v2 & v' & A & B & C & D & E). exists (PTree.set id (if b then v1 else v2) e), (PTree.set id v' e'). split. subst s. constructor; auto. split. unfold s0; destruct b. eapply classify_stmt_sound_2; eauto. rewrite PTree.gsident by (inv C; auto). apply classify_stmt_sound_1; auto. apply set_var_lessdef; auto. - destruct (ident_eq id id0); try discriminate. subst id0. destruct (is_known ki id && safe_expr ki a && safe_expr ki a0
&& if_conversion_heuristic cond a a0 (env id)) eqn:B; inv IFC. InvBooleans. exploit (eval_select_safe_exprs a a0); eauto. eapply classify_stmt_wt; eauto. eapply classify_stmt_wt; eauto. intros (a' & v1 & v2 & v' & A & B & C & D & E). exists (PTree.set id (if b then v1 else v2) e), (PTree.set id v' e'). split. subst s. constructor; auto. split. unfold s0; destruct b; eapply classify_stmt_sound_2; eauto. apply set_var_lessdef; auto. Qed.
End EXPRESSIONS. (** Semantic preservation for functions and statements. *)
Inductive match_cont: Cminor.program -> helper_functions -> known_idents -> typenv -> Cminor.cont -> CminorSel.cont -> Prop :=
| match_cont_seq: forall cunit hf ki env s s' k k',
sel_stmt (prog_defmap cunit) ki env s = OK s' ->
match_cont cunit hf ki env k k' ->
match_cont cunit hf ki env (Cminor.Kseq s k) (Kseq s' k')
| match_cont_block: forall cunit hf ki env k k',
match_cont cunit hf ki env k k' ->
match_cont cunit hf ki env (Cminor.Kblock k) (Kblock k')
| match_cont_other: forall cunit hf ki env k k',
match_call_cont k k' ->
match_cont cunit hf ki env k k'
with match_call_cont: Cminor.cont -> CminorSel.cont -> Prop :=
| match_cont_stop:
match_call_cont Cminor.Kstop Kstop
| match_cont_call: forall cunit hf env id f sp e k f' e' k',
linkorder cunit prog ->
helper_functions_declared cunit hf ->
sel_function (prog_defmap cunit) hf f = OK f' ->
type_function f = OK env ->
match_cont cunit hf (known_id f) env k k' ->
env_lessdef e e' ->
match_call_cont (Cminor.Kcall id f sp e k) (Kcall id f' sp e' k'). Inductive match_states: Cminor.state -> CminorSel.state -> Prop :=
| match_state: forall cunit hf f f' s k s' k' sp e m e' m' env
(LINK: linkorder cunit prog)
(HF: helper_functions_declared cunit hf)
(TF: sel_function (prog_defmap cunit) hf f = OK f')
(TYF: type_function f = OK env)
(TS: sel_stmt (prog_defmap cunit) (known_id f) env s = OK s')
(MC: match_cont cunit hf (known_id f) env k k')
(LD: env_lessdef e e')
(ME: Mem.extends m m'),
match_states
(Cminor.State f s k sp e m)
(State f' s' k' sp e' m')
| match_callstate: forall cunit f f' args args' k k' m m'
(LINK: linkorder cunit prog)
(TF: match_fundef cunit f f')
(MC: match_call_cont k k')
(LD: Val.lessdef_list args args')
(ME: Mem.extends m m'),
match_states
(Cminor.Callstate f args k m)
(Callstate f' args' k' m')
| match_returnstate: forall v v' k k' m m'
(MC: match_call_cont k k')
(LD: Val.lessdef v v')
(ME: Mem.extends m m'),
match_states
(Cminor.Returnstate v k m)
(Returnstate v' k' m')
| match_builtin_1: forall cunit hf ef args optid f sp e k m al f' e' k' m' env
(LINK: linkorder cunit prog)
(HF: helper_functions_declared cunit hf)
(TF: sel_function (prog_defmap cunit) hf f = OK f')
(TYF: type_function f = OK env)
(MC: match_cont cunit hf (known_id f) env k k')
(EA: Cminor.eval_exprlist ge sp e m al args)
(LDE: env_lessdef e e')
(ME: Mem.extends m m'),
match_states
(Cminor.Callstate (External ef) args (Cminor.Kcall optid f sp e k) m)
(State f' (sel_builtin optid ef al) k' sp e' m')
| match_builtin_2: forall cunit hf v v' optid f sp e k m f' e' m' k' env
(LINK: linkorder cunit prog)
(HF: helper_functions_declared cunit hf)
(TF: sel_function (prog_defmap cunit) hf f = OK f')
(TYF: type_function f = OK env)
(MC: match_cont cunit hf (known_id f) env k k')
(LDV: Val.lessdef v v')
(LDE: env_lessdef (set_optvar optid v e) e')
(ME: Mem.extends m m'),
match_states
(Cminor.Returnstate v (Cminor.Kcall optid f sp e k) m)
(State f' Sskip k' sp e' m').