From 9ab3738ae87a554fb742420b8c81ced4cd3c66c7 Mon Sep 17 00:00:00 2001 From: Bernhard Schommer Date: Tue, 8 Sep 2020 13:56:01 +0200 Subject: Changed cc_varargs to an option type Instead of being a simple boolean we now use an option type to record the number of fixed (non-vararg) arguments. Hence, `None` means not vararg, and `Some n` means `n` fixed arguments followed with varargs. --- riscV/Asmexpand.ml | 6 +++--- riscV/Conventions1.v | 3 ++- 2 files changed, 5 insertions(+), 4 deletions(-) (limited to 'riscV') diff --git a/riscV/Asmexpand.ml b/riscV/Asmexpand.ml index 810514a3..80e33b2b 100644 --- a/riscV/Asmexpand.ml +++ b/riscV/Asmexpand.ml @@ -100,7 +100,7 @@ let rec fixup_variadic_call ri rf tyl = end let fixup_call sg = - if sg.sig_cc.cc_vararg then fixup_variadic_call 0 0 sg.sig_args + if (sg.sig_cc.cc_vararg <> None) then fixup_variadic_call 0 0 sg.sig_args (* Handling of annotations *) @@ -588,7 +588,7 @@ let expand_instruction instr = | Pallocframe (sz, ofs) -> let sg = get_current_function_sig() in emit (Pmv (X30, X2)); - if sg.sig_cc.cc_vararg then begin + if (sg.sig_cc.cc_vararg <> None) then begin let n = arguments_size sg in let extra_sz = if n >= 8 then 0 else align ((8 - n) * wordsize) 16 in let full_sz = Z.add sz (Z.of_uint extra_sz) in @@ -606,7 +606,7 @@ let expand_instruction instr = | Pfreeframe (sz, ofs) -> let sg = get_current_function_sig() in let extra_sz = - if sg.sig_cc.cc_vararg then begin + if (sg.sig_cc.cc_vararg <> None) then begin let n = arguments_size sg in if n >= 8 then 0 else align ((8 - n) * wordsize) 16 end else 0 in diff --git a/riscV/Conventions1.v b/riscV/Conventions1.v index 17326139..74c99d07 100644 --- a/riscV/Conventions1.v +++ b/riscV/Conventions1.v @@ -281,7 +281,8 @@ Fixpoint loc_arguments_rec (va: bool) when calling a function with signature [s]. *) Definition loc_arguments (s: signature) : list (rpair loc) := - loc_arguments_rec s.(sig_cc).(cc_vararg) s.(sig_args) 0 0 0. + let va := match s.(sig_cc).(cc_vararg) with Some _ => true | None => false end in + loc_arguments_rec va s.(sig_args) 0 0 0. (** Argument locations are either non-temporary registers or [Outgoing] stack slots at nonnegative offsets. *) -- cgit From 2076a3bb31daf5bc3663fce25de6d837ace3f952 Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Tue, 22 Dec 2020 15:06:19 +0100 Subject: RISC-V: revised calling conventions for variadic functions Fixed (non-variadic) arguments follow the standard calling conventions. It's only the variadic arguments that need special treatment. --- riscV/Asmexpand.ml | 96 ++++++++++++++++++++++++++++++++++------------------ riscV/Conventions1.v | 72 ++++++++++++++++++++++----------------- 2 files changed, 105 insertions(+), 63 deletions(-) (limited to 'riscV') diff --git a/riscV/Asmexpand.ml b/riscV/Asmexpand.ml index 80e33b2b..f01ecb23 100644 --- a/riscV/Asmexpand.ml +++ b/riscV/Asmexpand.ml @@ -57,50 +57,59 @@ let expand_storeind_ptr src base ofs = registers. *) -(* Fix-up code around calls to variadic functions. Floating-point arguments - residing in FP registers need to be moved to integer registers. *) +(* Fix-up code around calls to variadic functions. + Floating-point variadic arguments residing in FP registers need to + be moved to integer registers. *) let int_param_regs = [| X10; X11; X12; X13; X14; X15; X16; X17 |] let float_param_regs = [| F10; F11; F12; F13; F14; F15; F16; F17 |] -let rec fixup_variadic_call ri rf tyl = +let rec fixup_variadic_call fixed ri rf tyl = if ri < 8 then match tyl with | [] -> () | (Tint | Tany32) :: tyl -> - fixup_variadic_call (ri + 1) rf tyl + fixup_variadic_call (fixed - 1) (ri + 1) rf tyl | Tsingle :: tyl -> - let rs = float_param_regs.(rf) - and rd = int_param_regs.(ri) in - emit (Pfmvxs(rd, rs)); - fixup_variadic_call (ri + 1) (rf + 1) tyl + if fixed <= 0 then begin + let rs = float_param_regs.(rf) + and rd = int_param_regs.(ri) in + emit (Pfmvxs(rd, rs)) + end; + fixup_variadic_call (fixed - 1) (ri + 1) (rf + 1) tyl | Tlong :: tyl -> let ri' = if Archi.ptr64 then ri + 1 else align ri 2 + 2 in - fixup_variadic_call ri' rf tyl + fixup_variadic_call (fixed - 1) ri' rf tyl | (Tfloat | Tany64) :: tyl -> if Archi.ptr64 then begin - let rs = float_param_regs.(rf) - and rd = int_param_regs.(ri) in - emit (Pfmvxd(rd, rs)); - fixup_variadic_call (ri + 1) (rf + 1) tyl + if fixed <= 0 then begin + let rs = float_param_regs.(rf) + and rd = int_param_regs.(ri) in + emit (Pfmvxd(rd, rs)) + end; + fixup_variadic_call (fixed - 1) (ri + 1) (rf + 1) tyl end else begin let ri = align ri 2 in if ri < 8 then begin - let rs = float_param_regs.(rf) - and rd1 = int_param_regs.(ri) - and rd2 = int_param_regs.(ri + 1) in - emit (Paddiw(X2, X X2, Integers.Int.neg _16)); - emit (Pfsd(rs, X2, Ofsimm _0)); - emit (Plw(rd1, X2, Ofsimm _0)); - emit (Plw(rd2, X2, Ofsimm _4)); - emit (Paddiw(X2, X X2, _16)); - fixup_variadic_call (ri + 2) (rf + 1) tyl + if fixed <= 0 then begin + let rs = float_param_regs.(rf) + and rd1 = int_param_regs.(ri) + and rd2 = int_param_regs.(ri + 1) in + emit (Paddiw(X2, X X2, Integers.Int.neg _16)); + emit (Pfsd(rs, X2, Ofsimm _0)); + emit (Plw(rd1, X2, Ofsimm _0)); + emit (Plw(rd2, X2, Ofsimm _4)); + emit (Paddiw(X2, X X2, _16)) + end; + fixup_variadic_call (fixed - 1) (ri + 2) (rf + 1) tyl end end let fixup_call sg = - if (sg.sig_cc.cc_vararg <> None) then fixup_variadic_call 0 0 sg.sig_args + match sg.sig_cc.cc_vararg with + | None -> () + | Some fixed -> fixup_variadic_call (Z.to_int fixed) 0 0 sg.sig_args (* Handling of annotations *) @@ -305,18 +314,41 @@ let expand_builtin_vstore chunk args = (* Handling of varargs *) -(* Size in words of the arguments to a function. This includes both - arguments passed in registers and arguments passed on stack. *) +(* Number of integer registers, FP registers, and stack words + used to pass the (fixed) arguments to a function. *) + +let rec args_size ri rf ofs = function + | [] -> (ri, rf, ofs) + | (Tint | Tany32) :: l -> + if ri < 8 + then args_size (ri + 1) rf ofs l + else args_size ri rf (ofs + 1) l + | Tsingle :: l -> + if rf < 8 + then args_size ri (rf + 1) ofs l + else args_size ri rf (ofs + 1) l + | Tlong :: l -> + if Archi.ptr64 then + if ri < 8 + then args_size (ri + 1) rf ofs l + else args_size ri rf (ofs + 1) l + else + if ri < 7 then args_size (ri + 2) rf ofs l + else if ri = 7 then args_size (ri + 1) rf (ofs + 1) l + else args_size ri rf (align ofs 2 + 2) l + | (Tfloat | Tany64) :: l -> + if rf < 8 + then args_size ri (rf + 1) ofs l + else let ofs' = if Archi.ptr64 then ofs + 1 else align ofs 2 + 2 in + args_size ri rf ofs' l -let rec args_size sz = function - | [] -> sz - | (Tint | Tsingle | Tany32) :: l -> - args_size (sz + 1) l - | (Tlong | Tfloat | Tany64) :: l -> - args_size (if Archi.ptr64 then sz + 1 else align sz 2 + 2) l +(* Size in words of the arguments to a function. This includes both + arguments passed in integer registers and arguments passed on stack, + but not arguments passed in FP registers. *) let arguments_size sg = - args_size 0 sg.sig_args + let (ri, _, ofs) = args_size 0 0 0 sg.sig_args in + ri + ofs let save_arguments first_reg base_ofs = for i = first_reg to 7 do diff --git a/riscV/Conventions1.v b/riscV/Conventions1.v index 74c99d07..0fc8295e 100644 --- a/riscV/Conventions1.v +++ b/riscV/Conventions1.v @@ -172,25 +172,28 @@ Qed. (** ** Location of function arguments *) (** The RISC-V ABI states the following conventions for passing arguments - to a function: + to a function. First for non-variadic functions: -- RV64, not variadic: pass the first 8 integer arguments in - integer registers (a1...a8: int_param_regs), the first 8 FP arguments - in FP registers (fa1...fa8: float_param_regs), and the remaining - arguments on the stack, in 8-byte slots. +- RV64: pass the first 8 integer arguments in integer registers + (a1...a8: int_param_regs), the first 8 FP arguments in FP registers + (fa1...fa8: float_param_regs), and the remaining arguments on the + stack, in 8-byte slots. -- RV32, not variadic: same, but arguments of 64-bit integer type - are passed in two consecutive integer registers (a(i), a(i+1)) - or in a(8) and on a 32-bit word on the stack. Stack-allocated - arguments are aligned to their natural alignment. +- RV32: same, but arguments of 64-bit integer type are passed in two + consecutive integer registers (a(i), a(i+1)) or in a(8) and on a + 32-bit word on the stack. Stack-allocated arguments are aligned to + their natural alignment. -- RV64, variadic: pass the first 8 arguments in integer registers - (a1...a8), including FP arguments; pass the remaining arguments on - the stack, in 8-byte slots. +For variadic functions, the fixed arguments are passed as described +above, then the variadic arguments receive special treatment: -- RV32, variadic: same, but arguments of 64-bit types (integers as well +- RV64: all variadic arguments, including FP arguments, + are passed in the remaining integer registers (a1...a8), + then on the stack, in 8-byte slots. + +- RV32: likewise, but arguments of 64-bit types (integers as well as floats) are passed in two consecutive aligned integer registers - (a(2i), a(2i+1)). + (a(2i), a(2i+1)), or on the stack, in aligned 8-byte slots. The passing of FP arguments to variadic functions in integer registers doesn't quite fit CompCert's model. We do our best by passing the FP @@ -253,36 +256,43 @@ Definition split_long_arg (va: bool) (ri rf ofs: Z) rec ri rf (ofs + 2) end. -Fixpoint loc_arguments_rec (va: bool) - (tyl: list typ) (ri rf ofs: Z) {struct tyl} : list (rpair loc) := +Fixpoint loc_arguments_rec + (tyl: list typ) (fixed ri rf ofs: Z) {struct tyl} : list (rpair loc) := match tyl with | nil => nil | (Tint | Tany32) as ty :: tys => (* pass in one integer register or on stack *) - int_arg ri rf ofs ty (loc_arguments_rec va tys) + int_arg ri rf ofs ty (loc_arguments_rec tys (fixed - 1)) | Tsingle as ty :: tys => (* pass in one FP register or on stack. If vararg, reserve 1 integer register. *) - float_arg va ri rf ofs ty (loc_arguments_rec va tys) + float_arg (zle fixed 0) ri rf ofs ty (loc_arguments_rec tys (fixed - 1)) | Tlong as ty :: tys => if Archi.ptr64 then (* pass in one integer register or on stack *) - int_arg ri rf ofs ty (loc_arguments_rec va tys) + int_arg ri rf ofs ty (loc_arguments_rec tys (fixed - 1)) else (* pass in register pair or on stack; align register pair if vararg *) - split_long_arg va ri rf ofs(loc_arguments_rec va tys) + split_long_arg (zle fixed 0) ri rf ofs(loc_arguments_rec tys (fixed - 1)) | (Tfloat | Tany64) as ty :: tys => (* pass in one FP register or on stack. If vararg, reserve 1 or 2 integer registers. *) - float_arg va ri rf ofs ty (loc_arguments_rec va tys) + float_arg (zle fixed 0) ri rf ofs ty (loc_arguments_rec tys (fixed - 1)) + end. + +(** Number of fixed arguments for a function with signature [s]. *) + +Definition fixed_arguments (s: signature) : Z := + match s.(sig_cc).(cc_vararg) with + | Some n => n + | None => list_length_z s.(sig_args) end. (** [loc_arguments s] returns the list of locations where to store arguments when calling a function with signature [s]. *) Definition loc_arguments (s: signature) : list (rpair loc) := - let va := match s.(sig_cc).(cc_vararg) with Some _ => true | None => false end in - loc_arguments_rec va s.(sig_args) 0 0 0. + loc_arguments_rec s.(sig_args) (fixed_arguments s) 0 0 0. (** Argument locations are either non-temporary registers or [Outgoing] stack slots at nonnegative offsets. *) @@ -362,20 +372,20 @@ Proof. + subst p; repeat split; auto using Z.divide_1_l. omega. + eapply OF; [idtac|eauto]. omega. } - cut (forall va tyl ri rf ofs, ofs >= 0 -> OK (loc_arguments_rec va tyl ri rf ofs)). + cut (forall tyl fixed ri rf ofs, ofs >= 0 -> OK (loc_arguments_rec tyl fixed ri rf ofs)). unfold OK. eauto. induction tyl as [ | ty1 tyl]; intros until ofs; intros OO; simpl. - red; simpl; tauto. - destruct ty1. -+ (* int *) apply A; auto. -+ (* float *) apply B; auto. ++ (* int *) apply A; unfold OKF; auto. ++ (* float *) apply B; unfold OKF; auto. + (* long *) destruct Archi.ptr64. - apply A; auto. - apply C; auto. -+ (* single *) apply B; auto. -+ (* any32 *) apply A; auto. -+ (* any64 *) apply B; auto. + apply A; unfold OKF; auto. + apply C; unfold OKF; auto. ++ (* single *) apply B; unfold OKF; auto. ++ (* any32 *) apply A; unfold OKF; auto. ++ (* any64 *) apply B; unfold OKF; auto. Qed. Lemma loc_arguments_acceptable: -- cgit From aba0e740f25ffa5c338dfa76cab71144802cebc2 Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Sun, 21 Jun 2020 18:22:00 +0200 Subject: Replace `omega` tactic with `lia` Since Coq 8.12, `omega` is flagged as deprecated and scheduled for removal. Also replace CompCert's homemade tactics `omegaContradiction`, `xomega`, and `xomegaContradiction` with `lia` and `extlia`. Turn back on the deprecation warning for uses of `omega`. Make the proof of `Ctypes.sizeof_pos` more robust to variations in `lia`. --- riscV/Asm.v | 2 +- riscV/Asmgenproof.v | 12 ++++++------ riscV/Asmgenproof1.v | 20 +++++++++---------- riscV/ConstpropOpproof.v | 2 +- riscV/Conventions1.v | 22 ++++++++++----------- riscV/NeedOp.v | 4 ++-- riscV/SelectLongproof.v | 6 +++--- riscV/SelectOpproof.v | 34 ++++++++++++++++---------------- riscV/Stacklayout.v | 50 ++++++++++++++++++++++++------------------------ 9 files changed, 76 insertions(+), 76 deletions(-) (limited to 'riscV') diff --git a/riscV/Asm.v b/riscV/Asm.v index dc410a3b..30b128ec 100644 --- a/riscV/Asm.v +++ b/riscV/Asm.v @@ -1157,7 +1157,7 @@ Ltac Equalities := split. auto. intros. destruct B; auto. subst. auto. - (* trace length *) red; intros. inv H; simpl. - omega. + lia. eapply external_call_trace_length; eauto. eapply external_call_trace_length; eauto. - (* initial states *) diff --git a/riscV/Asmgenproof.v b/riscV/Asmgenproof.v index 5ec57886..ab07d071 100644 --- a/riscV/Asmgenproof.v +++ b/riscV/Asmgenproof.v @@ -67,7 +67,7 @@ Lemma transf_function_no_overflow: transf_function f = OK tf -> list_length_z tf.(fn_code) <= Ptrofs.max_unsigned. Proof. intros. monadInv H. destruct (zlt Ptrofs.max_unsigned (list_length_z x.(fn_code))); inv EQ0. - omega. + lia. Qed. Lemma exec_straight_exec: @@ -432,8 +432,8 @@ Proof. split. unfold goto_label. rewrite P. rewrite H1. auto. split. rewrite Pregmap.gss. constructor; auto. rewrite Ptrofs.unsigned_repr. replace (pos' - 0) with pos' in Q. - auto. omega. - generalize (transf_function_no_overflow _ _ H0). omega. + auto. lia. + generalize (transf_function_no_overflow _ _ H0). lia. intros. apply Pregmap.gso; auto. Qed. @@ -948,10 +948,10 @@ Local Transparent destroyed_by_op. rewrite <- (sp_val _ _ _ AG). rewrite chunk_of_Tptr in F. rewrite F. reflexivity. reflexivity. eexact U. } - exploit exec_straight_steps_2; eauto using functions_transl. omega. constructor. + exploit exec_straight_steps_2; eauto using functions_transl. lia. constructor. intros (ofs' & X & Y). left; exists (State rs3 m3'); split. - eapply exec_straight_steps_1; eauto. omega. constructor. + eapply exec_straight_steps_1; eauto. lia. constructor. econstructor; eauto. rewrite X; econstructor; eauto. apply agree_exten with rs2; eauto with asmgen. @@ -980,7 +980,7 @@ Local Transparent destroyed_at_function_entry. - (* return *) inv STACKS. simpl in *. - right. split. omega. split. auto. + right. split. lia. split. auto. rewrite <- ATPC in H5. econstructor; eauto. congruence. Qed. diff --git a/riscV/Asmgenproof1.v b/riscV/Asmgenproof1.v index c20c4e49..253e769f 100644 --- a/riscV/Asmgenproof1.v +++ b/riscV/Asmgenproof1.v @@ -35,7 +35,7 @@ Proof. - set (m := Int.sub n lo). assert (A: eqmod (two_p 12) (Int.unsigned lo) (Int.unsigned n)) by (apply Int.eqmod_sign_ext'; compute; auto). assert (B: eqmod (two_p 12) (Int.unsigned n - Int.unsigned lo) 0). - { replace 0 with (Int.unsigned n - Int.unsigned n) by omega. + { replace 0 with (Int.unsigned n - Int.unsigned n) by lia. auto using eqmod_sub, eqmod_refl. } assert (C: eqmod (two_p 12) (Int.unsigned m) 0). { apply eqmod_trans with (Int.unsigned n - Int.unsigned lo); auto. @@ -45,7 +45,7 @@ Proof. { apply eqmod_mod_eq in C. unfold Int.modu. change (Int.unsigned (Int.repr 4096)) with (two_p 12). rewrite C. reflexivity. - apply two_p_gt_ZERO; omega. } + apply two_p_gt_ZERO; lia. } rewrite <- (Int.divu_pow2 m (Int.repr 4096) (Int.repr 12)) by auto. rewrite Int.shl_mul_two_p. change (two_p (Int.unsigned (Int.repr 12))) with 4096. @@ -684,18 +684,18 @@ Proof. unfold Val.cmp; destruct (rs#r1); simpl; auto. rewrite B1. unfold Int.lt. rewrite zlt_false. auto. change (Int.signed (Int.repr Int.max_signed)) with Int.max_signed. - generalize (Int.signed_range i); omega. + generalize (Int.signed_range i); lia. * exploit (opimm32_correct Psltw Psltiw (Val.cmp Clt)); eauto. intros (rs1 & A1 & B1 & C1). exists rs1; split. eexact A1. split; auto. rewrite B1. unfold Val.cmp; simpl; destruct (rs#r1); simpl; auto. unfold Int.lt. replace (Int.signed (Int.add n Int.one)) with (Int.signed n + 1). destruct (zlt (Int.signed n) (Int.signed i)). - rewrite zlt_false by omega. auto. - rewrite zlt_true by omega. auto. + rewrite zlt_false by lia. auto. + rewrite zlt_true by lia. auto. rewrite Int.add_signed. symmetry; apply Int.signed_repr. assert (Int.signed n <> Int.max_signed). { red; intros E. elim H1. rewrite <- (Int.repr_signed n). rewrite E. auto. } - generalize (Int.signed_range n); omega. + generalize (Int.signed_range n); lia. + apply DFL. + apply DFL. Qed. @@ -782,18 +782,18 @@ Proof. unfold Val.cmpl; destruct (rs#r1); simpl; auto. rewrite B1. unfold Int64.lt. rewrite zlt_false. auto. change (Int64.signed (Int64.repr Int64.max_signed)) with Int64.max_signed. - generalize (Int64.signed_range i); omega. + generalize (Int64.signed_range i); lia. * exploit (opimm64_correct Psltl Psltil (fun v1 v2 => Val.maketotal (Val.cmpl Clt v1 v2))); eauto. intros (rs1 & A1 & B1 & C1). exists rs1; split. eexact A1. split; auto. rewrite B1. unfold Val.cmpl; simpl; destruct (rs#r1); simpl; auto. unfold Int64.lt. replace (Int64.signed (Int64.add n Int64.one)) with (Int64.signed n + 1). destruct (zlt (Int64.signed n) (Int64.signed i)). - rewrite zlt_false by omega. auto. - rewrite zlt_true by omega. auto. + rewrite zlt_false by lia. auto. + rewrite zlt_true by lia. auto. rewrite Int64.add_signed. symmetry; apply Int64.signed_repr. assert (Int64.signed n <> Int64.max_signed). { red; intros E. elim H1. rewrite <- (Int64.repr_signed n). rewrite E. auto. } - generalize (Int64.signed_range n); omega. + generalize (Int64.signed_range n); lia. + apply DFL. + apply DFL. Qed. diff --git a/riscV/ConstpropOpproof.v b/riscV/ConstpropOpproof.v index 765aa035..8445e55f 100644 --- a/riscV/ConstpropOpproof.v +++ b/riscV/ConstpropOpproof.v @@ -333,7 +333,7 @@ Proof. Int.bit_solve. destruct (zlt i0 n0). replace (Int.testbit n i0) with (negb (Int.testbit Int.zero i0)). rewrite Int.bits_zero. simpl. rewrite andb_true_r. auto. - rewrite <- EQ. rewrite Int.bits_zero_ext by omega. rewrite zlt_true by auto. + rewrite <- EQ. rewrite Int.bits_zero_ext by lia. rewrite zlt_true by auto. rewrite Int.bits_not by auto. apply negb_involutive. rewrite H6 by auto. auto. econstructor; split; eauto. auto. diff --git a/riscV/Conventions1.v b/riscV/Conventions1.v index 0fc8295e..5dd50a4c 100644 --- a/riscV/Conventions1.v +++ b/riscV/Conventions1.v @@ -320,14 +320,14 @@ Proof. assert (AL: forall ofs ty, ofs >= 0 -> align ofs (typesize ty) >= 0). { intros. assert (ofs <= align ofs (typesize ty)) by (apply align_le; apply typesize_pos). - omega. } + lia. } assert (ALD: forall ofs ty, ofs >= 0 -> (typealign ty | align ofs (typesize ty))). { intros. eapply Z.divide_trans. apply typealign_typesize. apply align_divides. apply typesize_pos. } assert (SK: (if Archi.ptr64 then 2 else 1) > 0). - { destruct Archi.ptr64; omega. } + { destruct Archi.ptr64; lia. } assert (SKK: forall ty, (if Archi.ptr64 then 2 else typesize ty) > 0). - { intros. destruct Archi.ptr64. omega. apply typesize_pos. } + { intros. destruct Archi.ptr64. lia. apply typesize_pos. } assert (A: forall ri rf ofs ty f, OKF f -> ofs >= 0 -> OK (int_arg ri rf ofs ty f)). { intros until f; intros OF OO; red; unfold int_arg; intros. @@ -336,7 +336,7 @@ Proof. - eapply OF; eauto. - subst p; simpl. auto using align_divides, typealign_pos. - eapply OF; [idtac|eauto]. - generalize (AL ofs ty OO) (SKK ty); omega. + generalize (AL ofs ty OO) (SKK ty); lia. } assert (B: forall va ri rf ofs ty f, OKF f -> ofs >= 0 -> OK (float_arg va ri rf ofs ty f)). @@ -347,12 +347,12 @@ Proof. + subst p; simpl. apply CSF. eapply list_nth_z_in; eauto. + eapply OF; eauto. + subst p; repeat split; auto. - + eapply OF; [idtac|eauto]. generalize (AL ofs ty OO) (SKK ty); omega. + + eapply OF; [idtac|eauto]. generalize (AL ofs ty OO) (SKK ty); lia. + subst p; simpl. apply CSF. eapply list_nth_z_in; eauto. + eapply OF; eauto. - destruct H. + subst p; repeat split; auto. - + eapply OF; [idtac|eauto]. generalize (AL ofs ty OO) (SKK ty); omega. + + eapply OF; [idtac|eauto]. generalize (AL ofs ty OO) (SKK ty); lia. } assert (C: forall va ri rf ofs f, OKF f -> ofs >= 0 -> OK (split_long_arg va ri rf ofs f)). @@ -364,13 +364,13 @@ Proof. [destruct (list_nth_z int_param_regs (ri'+1)) as [r2|] eqn:NTH2 | idtac]. - red; simpl; intros; destruct H. + subst p; split; apply CSI; eauto using list_nth_z_in. - + eapply OF; [idtac|eauto]. omega. + + eapply OF; [idtac|eauto]. lia. - red; simpl; intros; destruct H. + subst p; split. split; auto using Z.divide_1_l. apply CSI; eauto using list_nth_z_in. - + eapply OF; [idtac|eauto]. omega. + + eapply OF; [idtac|eauto]. lia. - red; simpl; intros; destruct H. - + subst p; repeat split; auto using Z.divide_1_l. omega. - + eapply OF; [idtac|eauto]. omega. + + subst p; repeat split; auto using Z.divide_1_l. lia. + + eapply OF; [idtac|eauto]. lia. } cut (forall tyl fixed ri rf ofs, ofs >= 0 -> OK (loc_arguments_rec tyl fixed ri rf ofs)). unfold OK. eauto. @@ -392,7 +392,7 @@ Lemma loc_arguments_acceptable: forall (s: signature) (p: rpair loc), In p (loc_arguments s) -> forall_rpair loc_argument_acceptable p. Proof. - unfold loc_arguments; intros. eapply loc_arguments_rec_charact; eauto. omega. + unfold loc_arguments; intros. eapply loc_arguments_rec_charact; eauto. lia. Qed. Lemma loc_arguments_main: diff --git a/riscV/NeedOp.v b/riscV/NeedOp.v index 117bbcb4..4070431a 100644 --- a/riscV/NeedOp.v +++ b/riscV/NeedOp.v @@ -164,8 +164,8 @@ Lemma operation_is_redundant_sound: vagree v arg1' nv. Proof. intros. destruct op; simpl in *; try discriminate; inv H1; FuncInv; subst. -- apply sign_ext_redundant_sound; auto. omega. -- apply sign_ext_redundant_sound; auto. omega. +- apply sign_ext_redundant_sound; auto. lia. +- apply sign_ext_redundant_sound; auto. lia. - apply andimm_redundant_sound; auto. - apply orimm_redundant_sound; auto. Qed. diff --git a/riscV/SelectLongproof.v b/riscV/SelectLongproof.v index 78a1935d..3794e050 100644 --- a/riscV/SelectLongproof.v +++ b/riscV/SelectLongproof.v @@ -502,8 +502,8 @@ Proof. assert (LTU2: Int.ltu (Int.sub Int64.iwordsize' n) Int64.iwordsize' = true). { unfold Int.ltu; apply zlt_true. unfold Int.sub. change (Int.unsigned Int64.iwordsize') with 64. - rewrite Int.unsigned_repr. omega. - assert (64 < Int.max_unsigned) by reflexivity. omega. } + rewrite Int.unsigned_repr. lia. + assert (64 < Int.max_unsigned) by reflexivity. lia. } assert (X: eval_expr ge sp e m le (Eop (Oshrlimm (Int.repr (Int64.zwordsize - 1))) (a ::: Enil)) (Vlong (Int64.shr' i (Int.repr (Int64.zwordsize - 1))))). @@ -514,7 +514,7 @@ Proof. TrivialExists. constructor. EvalOp. simpl; eauto. constructor. simpl. unfold Int.ltu; rewrite zlt_true. rewrite Int64.shrx'_shr_2 by auto. reflexivity. - change (Int.unsigned Int64.iwordsize') with 64; omega. + change (Int.unsigned Int64.iwordsize') with 64; lia. *) Qed. diff --git a/riscV/SelectOpproof.v b/riscV/SelectOpproof.v index 593be1ed..b0b4b794 100644 --- a/riscV/SelectOpproof.v +++ b/riscV/SelectOpproof.v @@ -365,20 +365,20 @@ Proof. change (Int.ltu (Int.repr 32) Int64.iwordsize') with true; simpl. apply Val.lessdef_same. f_equal. transitivity (Int.repr (Z.shiftr (Int.signed i * Int.signed i0) 32)). - unfold Int.mulhs; f_equal. rewrite Zshiftr_div_two_p by omega. reflexivity. + unfold Int.mulhs; f_equal. rewrite Zshiftr_div_two_p by lia. reflexivity. apply Int.same_bits_eq; intros n N. change Int.zwordsize with 32 in *. - assert (N1: 0 <= n < 64) by omega. + assert (N1: 0 <= n < 64) by lia. rewrite Int64.bits_loword by auto. rewrite Int64.bits_shr' by auto. change (Int.unsigned (Int.repr 32)) with 32. change Int64.zwordsize with 64. - rewrite zlt_true by omega. + rewrite zlt_true by lia. rewrite Int.testbit_repr by auto. - unfold Int64.mul. rewrite Int64.testbit_repr by (change Int64.zwordsize with 64; omega). + unfold Int64.mul. rewrite Int64.testbit_repr by (change Int64.zwordsize with 64; lia). transitivity (Z.testbit (Int.signed i * Int.signed i0) (n + 32)). - rewrite Z.shiftr_spec by omega. auto. + rewrite Z.shiftr_spec by lia. auto. apply Int64.same_bits_eqm. apply Int64.eqm_mult; apply Int64.eqm_unsigned_repr. - change Int64.zwordsize with 64; omega. + change Int64.zwordsize with 64; lia. - TrivialExists. Qed. @@ -393,20 +393,20 @@ Proof. change (Int.ltu (Int.repr 32) Int64.iwordsize') with true; simpl. apply Val.lessdef_same. f_equal. transitivity (Int.repr (Z.shiftr (Int.unsigned i * Int.unsigned i0) 32)). - unfold Int.mulhu; f_equal. rewrite Zshiftr_div_two_p by omega. reflexivity. + unfold Int.mulhu; f_equal. rewrite Zshiftr_div_two_p by lia. reflexivity. apply Int.same_bits_eq; intros n N. change Int.zwordsize with 32 in *. - assert (N1: 0 <= n < 64) by omega. + assert (N1: 0 <= n < 64) by lia. rewrite Int64.bits_loword by auto. rewrite Int64.bits_shru' by auto. change (Int.unsigned (Int.repr 32)) with 32. change Int64.zwordsize with 64. - rewrite zlt_true by omega. + rewrite zlt_true by lia. rewrite Int.testbit_repr by auto. - unfold Int64.mul. rewrite Int64.testbit_repr by (change Int64.zwordsize with 64; omega). + unfold Int64.mul. rewrite Int64.testbit_repr by (change Int64.zwordsize with 64; lia). transitivity (Z.testbit (Int.unsigned i * Int.unsigned i0) (n + 32)). - rewrite Z.shiftr_spec by omega. auto. + rewrite Z.shiftr_spec by lia. auto. apply Int64.same_bits_eqm. apply Int64.eqm_mult; apply Int64.eqm_unsigned_repr. - change Int64.zwordsize with 64; omega. + change Int64.zwordsize with 64; lia. - TrivialExists. Qed. @@ -563,8 +563,8 @@ Proof. assert (LTU2: Int.ltu (Int.sub Int.iwordsize n) Int.iwordsize = true). { unfold Int.ltu; apply zlt_true. unfold Int.sub. change (Int.unsigned Int.iwordsize) with 32. - rewrite Int.unsigned_repr. omega. - assert (32 < Int.max_unsigned) by reflexivity. omega. } + rewrite Int.unsigned_repr. lia. + assert (32 < Int.max_unsigned) by reflexivity. lia. } assert (X: eval_expr ge sp e m le (Eop (Oshrimm (Int.repr (Int.zwordsize - 1))) (a ::: Enil)) (Vint (Int.shr i (Int.repr (Int.zwordsize - 1))))). @@ -575,7 +575,7 @@ Proof. TrivialExists. constructor. EvalOp. simpl; eauto. constructor. simpl. unfold Int.ltu; rewrite zlt_true. rewrite Int.shrx_shr_2 by auto. reflexivity. - change (Int.unsigned Int.iwordsize) with 32; omega. + change (Int.unsigned Int.iwordsize) with 32; lia. *) Qed. @@ -763,7 +763,7 @@ Qed. Theorem eval_cast8unsigned: unary_constructor_sound cast8unsigned (Val.zero_ext 8). Proof. red; intros until x. unfold cast8unsigned. - rewrite Val.zero_ext_and. apply eval_andimm. omega. + rewrite Val.zero_ext_and. apply eval_andimm. lia. Qed. Theorem eval_cast16signed: unary_constructor_sound cast16signed (Val.sign_ext 16). @@ -776,7 +776,7 @@ Qed. Theorem eval_cast16unsigned: unary_constructor_sound cast16unsigned (Val.zero_ext 16). Proof. red; intros until x. unfold cast8unsigned. - rewrite Val.zero_ext_and. apply eval_andimm. omega. + rewrite Val.zero_ext_and. apply eval_andimm. lia. Qed. Theorem eval_intoffloat: diff --git a/riscV/Stacklayout.v b/riscV/Stacklayout.v index d0c6a526..25f02aab 100644 --- a/riscV/Stacklayout.v +++ b/riscV/Stacklayout.v @@ -68,15 +68,15 @@ Local Opaque Z.add Z.mul sepconj range. set (ol := align (size_callee_save_area b ocs) 8). set (ostkdata := align (ol + 4 * b.(bound_local)) 8). replace (size_chunk Mptr) with w by (rewrite size_chunk_Mptr; auto). - assert (0 < w) by (unfold w; destruct Archi.ptr64; omega). + assert (0 < w) by (unfold w; destruct Archi.ptr64; lia). generalize b.(bound_local_pos) b.(bound_outgoing_pos) b.(bound_stack_data_pos); intros. - assert (0 <= 4 * b.(bound_outgoing)) by omega. - assert (4 * b.(bound_outgoing) <= olink) by (apply align_le; omega). - assert (olink + w <= oretaddr) by (unfold oretaddr; omega). - assert (oretaddr + w <= ocs) by (unfold ocs; omega). + assert (0 <= 4 * b.(bound_outgoing)) by lia. + assert (4 * b.(bound_outgoing) <= olink) by (apply align_le; lia). + assert (olink + w <= oretaddr) by (unfold oretaddr; lia). + assert (oretaddr + w <= ocs) by (unfold ocs; lia). assert (ocs <= size_callee_save_area b ocs) by (apply size_callee_save_area_incr). - assert (size_callee_save_area b ocs <= ol) by (apply align_le; omega). - assert (ol + 4 * b.(bound_local) <= ostkdata) by (apply align_le; omega). + assert (size_callee_save_area b ocs <= ol) by (apply align_le; lia). + assert (ol + 4 * b.(bound_local) <= ostkdata) by (apply align_le; lia). (* Reorder as: outgoing back link @@ -89,11 +89,11 @@ Local Opaque Z.add Z.mul sepconj range. rewrite sep_swap45. (* Apply range_split and range_split2 repeatedly *) unfold fe_ofs_arg. - apply range_split_2. fold olink; omega. omega. - apply range_split. omega. - apply range_split. omega. - apply range_split_2. fold ol. omega. omega. - apply range_drop_right with ostkdata. omega. + apply range_split_2. fold olink; lia. lia. + apply range_split. lia. + apply range_split. lia. + apply range_split_2. fold ol. lia. lia. + apply range_drop_right with ostkdata. lia. eapply sep_drop2. eexact H. Qed. @@ -109,16 +109,16 @@ Proof. set (ocs := oretaddr + w). set (ol := align (size_callee_save_area b ocs) 8). set (ostkdata := align (ol + 4 * b.(bound_local)) 8). - assert (0 < w) by (unfold w; destruct Archi.ptr64; omega). + assert (0 < w) by (unfold w; destruct Archi.ptr64; lia). generalize b.(bound_local_pos) b.(bound_outgoing_pos) b.(bound_stack_data_pos); intros. - assert (0 <= 4 * b.(bound_outgoing)) by omega. - assert (4 * b.(bound_outgoing) <= olink) by (apply align_le; omega). - assert (olink + w <= oretaddr) by (unfold oretaddr; omega). - assert (oretaddr + w <= ocs) by (unfold ocs; omega). + assert (0 <= 4 * b.(bound_outgoing)) by lia. + assert (4 * b.(bound_outgoing) <= olink) by (apply align_le; lia). + assert (olink + w <= oretaddr) by (unfold oretaddr; lia). + assert (oretaddr + w <= ocs) by (unfold ocs; lia). assert (ocs <= size_callee_save_area b ocs) by (apply size_callee_save_area_incr). - assert (size_callee_save_area b ocs <= ol) by (apply align_le; omega). - assert (ol + 4 * b.(bound_local) <= ostkdata) by (apply align_le; omega). - split. omega. apply align_le. omega. + assert (size_callee_save_area b ocs <= ol) by (apply align_le; lia). + assert (ol + 4 * b.(bound_local) <= ostkdata) by (apply align_le; lia). + split. lia. apply align_le. lia. Qed. Lemma frame_env_aligned: @@ -137,11 +137,11 @@ Proof. set (ocs := oretaddr + w). set (ol := align (size_callee_save_area b ocs) 8). set (ostkdata := align (ol + 4 * b.(bound_local)) 8). - assert (0 < w) by (unfold w; destruct Archi.ptr64; omega). + assert (0 < w) by (unfold w; destruct Archi.ptr64; lia). replace (align_chunk Mptr) with w by (rewrite align_chunk_Mptr; auto). split. apply Z.divide_0_r. - split. apply align_divides; omega. - split. apply align_divides; omega. - split. apply align_divides; omega. - apply Z.divide_add_r. apply align_divides; omega. apply Z.divide_refl. + split. apply align_divides; lia. + split. apply align_divides; lia. + split. apply align_divides; lia. + apply Z.divide_add_r. apply align_divides; lia. apply Z.divide_refl. Qed. -- cgit From e81d015e3cc2cb0c352792d0cac12f1594281bc2 Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Sun, 10 Jan 2021 14:34:58 +0100 Subject: RISC-V: wrong fixup code generated for vararg calls with fixed FP args This is a follow-up to 2076a3bb3. Integer registers were wrongly reserved for fixed FP arguments, causing variadic FP arguments to end up in the wrong integer registers. Added regression test in test/regression/varargs2.c --- riscV/Asmexpand.ml | 30 ++++++++++++++++++------------ 1 file changed, 18 insertions(+), 12 deletions(-) (limited to 'riscV') diff --git a/riscV/Asmexpand.ml b/riscV/Asmexpand.ml index f01ecb23..da97c4a8 100644 --- a/riscV/Asmexpand.ml +++ b/riscV/Asmexpand.ml @@ -72,27 +72,33 @@ let rec fixup_variadic_call fixed ri rf tyl = | (Tint | Tany32) :: tyl -> fixup_variadic_call (fixed - 1) (ri + 1) rf tyl | Tsingle :: tyl -> - if fixed <= 0 then begin + if fixed > 0 then + fixup_variadic_call (fixed - 1) ri (rf + 1) tyl + else begin let rs = float_param_regs.(rf) and rd = int_param_regs.(ri) in - emit (Pfmvxs(rd, rs)) - end; - fixup_variadic_call (fixed - 1) (ri + 1) (rf + 1) tyl + emit (Pfmvxs(rd, rs)); + fixup_variadic_call (fixed - 1) (ri + 1) (rf + 1) tyl + end | Tlong :: tyl -> let ri' = if Archi.ptr64 then ri + 1 else align ri 2 + 2 in fixup_variadic_call (fixed - 1) ri' rf tyl | (Tfloat | Tany64) :: tyl -> if Archi.ptr64 then begin - if fixed <= 0 then begin + if fixed > 0 then + fixup_variadic_call (fixed - 1) ri (rf + 1) tyl + else begin let rs = float_param_regs.(rf) and rd = int_param_regs.(ri) in - emit (Pfmvxd(rd, rs)) - end; - fixup_variadic_call (fixed - 1) (ri + 1) (rf + 1) tyl + emit (Pfmvxd(rd, rs)); + fixup_variadic_call (fixed - 1) (ri + 1) (rf + 1) tyl + end end else begin let ri = align ri 2 in if ri < 8 then begin - if fixed <= 0 then begin + if fixed > 0 then + fixup_variadic_call (fixed - 1) ri (rf + 1) tyl + else begin let rs = float_param_regs.(rf) and rd1 = int_param_regs.(ri) and rd2 = int_param_regs.(ri + 1) in @@ -100,9 +106,9 @@ let rec fixup_variadic_call fixed ri rf tyl = emit (Pfsd(rs, X2, Ofsimm _0)); emit (Plw(rd1, X2, Ofsimm _0)); emit (Plw(rd2, X2, Ofsimm _4)); - emit (Paddiw(X2, X X2, _16)) - end; - fixup_variadic_call (fixed - 1) (ri + 2) (rf + 1) tyl + emit (Paddiw(X2, X X2, _16)); + fixup_variadic_call (fixed - 1) (ri + 2) (rf + 1) tyl + end end end -- cgit From 88567ce6d247562a9fa9151eaa32f7ad63ea37c0 Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Mon, 11 Jan 2021 18:04:25 +0100 Subject: RISC-V: fix FP calling conventions This is a follow-up to e81d015e3. In the RISC-V ABI, FP arguments to functions are passed in integer registers (or pairs of integer registers) in two cases: 1- the FP argument is a variadic argument 2- the FP argument is a fixed argument but all 8 FP registers reserved for parameter passing have been used already. The previous implementation handled only case 1, with some problems. This commit implements both 1 and 2. To this end, 8 extra FP caller-save registers are used to hold the values of the FP arguments that must be passed in integer registers. Fixup code moves these FP registers to integer registers / register pairs. Symmetrically, at function entry, the integer registers / register pairs are moved back to the FP registers. 8 extra FP registers is enough because there are only 8 integer registers used for parameter passing, so at most 8 FP arguments may need to be moved to integer registers. --- riscV/Asm.v | 4 + riscV/Asmexpand.ml | 194 +++++++++++++++++++++++++++++-------------------- riscV/Conventions1.v | 86 ++++++++++++---------- riscV/TargetPrinter.ml | 4 + 4 files changed, 171 insertions(+), 117 deletions(-) (limited to 'riscV') diff --git a/riscV/Asm.v b/riscV/Asm.v index 30b128ec..7e1b1fc8 100644 --- a/riscV/Asm.v +++ b/riscV/Asm.v @@ -256,7 +256,9 @@ Inductive instruction : Type := (* floating point register move *) | Pfmv (rd: freg) (rs: freg) (**r move *) | Pfmvxs (rd: ireg) (rs: freg) (**r move FP single to integer register *) + | Pfmvsx (rd: freg) (rs: ireg) (**r move integer register to FP single *) | Pfmvxd (rd: ireg) (rs: freg) (**r move FP double to integer register *) + | Pfmvdx (rd: freg) (rs: ireg) (**r move integer register to FP double *) (* 32-bit (single-precision) floating point *) | Pfls (rd: freg) (ra: ireg) (ofs: offset) (**r load float *) @@ -969,7 +971,9 @@ Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : out | Pfence | Pfmvxs _ _ + | Pfmvsx _ _ | Pfmvxd _ _ + | Pfmvdx _ _ | Pfmins _ _ _ | Pfmaxs _ _ _ diff --git a/riscV/Asmexpand.ml b/riscV/Asmexpand.ml index da97c4a8..e8c142e9 100644 --- a/riscV/Asmexpand.ml +++ b/riscV/Asmexpand.ml @@ -24,6 +24,7 @@ open Asmexpandaux open AST open Camlcoq open! Integers +open Locations exception Error of string @@ -50,6 +51,86 @@ let expand_addptrofs dst src n = let expand_storeind_ptr src base ofs = List.iter emit (Asmgen.storeind_ptr src base ofs []) +(* Fix-up code around function calls and function entry. + Some floating-point arguments residing in FP registers need to be + moved to integer registers or register pairs. + Symmetrically, some floating-point parameter passed in integer + registers or register pairs need to be moved to FP registers. *) + +let int_param_regs = [| X10; X11; X12; X13; X14; X15; X16; X17 |] + +let move_single_arg fr i = + emit (Pfmvxs(int_param_regs.(i), fr)) + +let move_double_arg fr i = + if Archi.ptr64 then begin + emit (Pfmvxd(int_param_regs.(i), fr)) + end else begin + emit (Paddiw(X2, X X2, Integers.Int.neg _16)); + emit (Pfsd(fr, X2, Ofsimm _0)); + emit (Plw(int_param_regs.(i), X2, Ofsimm _0)); + if i < 7 then begin + emit (Plw(int_param_regs.(i + 1), X2, Ofsimm _4)) + end else begin + emit (Plw(X31, X2, Ofsimm _4)); + emit (Psw(X31, X2, Ofsimm _16)) + end; + emit (Paddiw(X2, X X2, _16)) + end + +let move_single_param fr i = + emit (Pfmvsx(fr, int_param_regs.(i))) + +let move_double_param fr i = + if Archi.ptr64 then begin + emit (Pfmvdx(fr, int_param_regs.(i))) + end else begin + emit (Paddiw(X2, X X2, Integers.Int.neg _16)); + emit (Psw(int_param_regs.(i), X2, Ofsimm _0)); + if i < 7 then begin + emit (Psw(int_param_regs.(i + 1), X2, Ofsimm _4)) + end else begin + emit (Plw(X31, X2, Ofsimm _16)); + emit (Psw(X31, X2, Ofsimm _4)) + end; + emit (Pfld(fr, X2, Ofsimm _0)); + emit (Paddiw(X2, X X2, _16)) + end + +let float_extra_index = function + | Machregs.F0 -> Some (F0, 0) + | Machregs.F1 -> Some (F1, 1) + | Machregs.F2 -> Some (F2, 2) + | Machregs.F3 -> Some (F3, 3) + | Machregs.F4 -> Some (F4, 4) + | Machregs.F5 -> Some (F5, 5) + | Machregs.F6 -> Some (F6, 6) + | Machregs.F7 -> Some (F7, 7) + | _ -> None + +let fixup_gen single double sg = + let fixup ty loc = + match ty, loc with + | Tsingle, One (R r) -> + begin match float_extra_index r with + | Some(r, i) -> single r i + | None -> () + end + | (Tfloat | Tany64), One (R r) -> + begin match float_extra_index r with + | Some(r, i) -> double r i + | None -> () + end + | _, _ -> () + in + List.iter2 fixup sg.sig_args (Conventions1.loc_arguments sg) + +let fixup_call sg = + fixup_gen move_single_arg move_double_arg sg + +let fixup_function_entry sg = + fixup_gen move_single_param move_double_param sg + (* Built-ins. They come in two flavors: - annotation statements: take their arguments in registers or stack locations; generate no code; @@ -57,66 +138,6 @@ let expand_storeind_ptr src base ofs = registers. *) -(* Fix-up code around calls to variadic functions. - Floating-point variadic arguments residing in FP registers need to - be moved to integer registers. *) - -let int_param_regs = [| X10; X11; X12; X13; X14; X15; X16; X17 |] -let float_param_regs = [| F10; F11; F12; F13; F14; F15; F16; F17 |] - -let rec fixup_variadic_call fixed ri rf tyl = - if ri < 8 then - match tyl with - | [] -> - () - | (Tint | Tany32) :: tyl -> - fixup_variadic_call (fixed - 1) (ri + 1) rf tyl - | Tsingle :: tyl -> - if fixed > 0 then - fixup_variadic_call (fixed - 1) ri (rf + 1) tyl - else begin - let rs = float_param_regs.(rf) - and rd = int_param_regs.(ri) in - emit (Pfmvxs(rd, rs)); - fixup_variadic_call (fixed - 1) (ri + 1) (rf + 1) tyl - end - | Tlong :: tyl -> - let ri' = if Archi.ptr64 then ri + 1 else align ri 2 + 2 in - fixup_variadic_call (fixed - 1) ri' rf tyl - | (Tfloat | Tany64) :: tyl -> - if Archi.ptr64 then begin - if fixed > 0 then - fixup_variadic_call (fixed - 1) ri (rf + 1) tyl - else begin - let rs = float_param_regs.(rf) - and rd = int_param_regs.(ri) in - emit (Pfmvxd(rd, rs)); - fixup_variadic_call (fixed - 1) (ri + 1) (rf + 1) tyl - end - end else begin - let ri = align ri 2 in - if ri < 8 then begin - if fixed > 0 then - fixup_variadic_call (fixed - 1) ri (rf + 1) tyl - else begin - let rs = float_param_regs.(rf) - and rd1 = int_param_regs.(ri) - and rd2 = int_param_regs.(ri + 1) in - emit (Paddiw(X2, X X2, Integers.Int.neg _16)); - emit (Pfsd(rs, X2, Ofsimm _0)); - emit (Plw(rd1, X2, Ofsimm _0)); - emit (Plw(rd2, X2, Ofsimm _4)); - emit (Paddiw(X2, X X2, _16)); - fixup_variadic_call (fixed - 1) (ri + 2) (rf + 1) tyl - end - end - end - -let fixup_call sg = - match sg.sig_cc.cc_vararg with - | None -> () - | Some fixed -> fixup_variadic_call (Z.to_int fixed) 0 0 sg.sig_args - (* Handling of annotations *) let expand_annot_val kind txt targ args res = @@ -323,37 +344,49 @@ let expand_builtin_vstore chunk args = (* Number of integer registers, FP registers, and stack words used to pass the (fixed) arguments to a function. *) -let rec args_size ri rf ofs = function +let arg_int_size ri rf ofs k = + if ri < 8 + then k (ri + 1) rf ofs + else k ri rf (ofs + 1) + +let arg_single_size ri rf ofs k = + if rf < 8 + then k ri (rf + 1) ofs + else arg_int_size ri rf ofs k + +let arg_long_size ri rf ofs k = + if Archi.ptr64 then + if ri < 8 + then k (ri + 1) rf ofs + else k ri rf (ofs + 1) + else + if ri < 7 then k (ri + 2) rf ofs + else if ri = 7 then k (ri + 1) rf (ofs + 1) + else k ri rf (align ofs 2 + 2) + +let arg_double_size ri rf ofs k = + if rf < 8 + then k ri (rf + 1) ofs + else arg_long_size ri rf ofs k + +let rec args_size l ri rf ofs = + match l with | [] -> (ri, rf, ofs) | (Tint | Tany32) :: l -> - if ri < 8 - then args_size (ri + 1) rf ofs l - else args_size ri rf (ofs + 1) l + arg_int_size ri rf ofs (args_size l) | Tsingle :: l -> - if rf < 8 - then args_size ri (rf + 1) ofs l - else args_size ri rf (ofs + 1) l + arg_single_size ri rf ofs (args_size l) | Tlong :: l -> - if Archi.ptr64 then - if ri < 8 - then args_size (ri + 1) rf ofs l - else args_size ri rf (ofs + 1) l - else - if ri < 7 then args_size (ri + 2) rf ofs l - else if ri = 7 then args_size (ri + 1) rf (ofs + 1) l - else args_size ri rf (align ofs 2 + 2) l + arg_long_size ri rf ofs (args_size l) | (Tfloat | Tany64) :: l -> - if rf < 8 - then args_size ri (rf + 1) ofs l - else let ofs' = if Archi.ptr64 then ofs + 1 else align ofs 2 + 2 in - args_size ri rf ofs' l + arg_double_size ri rf ofs (args_size l) (* Size in words of the arguments to a function. This includes both arguments passed in integer registers and arguments passed on stack, but not arguments passed in FP registers. *) let arguments_size sg = - let (ri, _, ofs) = args_size 0 0 0 sg.sig_args in + let (ri, _, ofs) = args_size sg.sig_args 0 0 0 in ri + ofs let save_arguments first_reg base_ofs = @@ -744,6 +777,7 @@ let preg_to_dwarf = function let expand_function id fn = try set_current_function fn; + fixup_function_entry fn.fn_sig; expand id (* sp= *) 2 preg_to_dwarf expand_instruction fn.fn_code; Errors.OK (get_current_function ()) with Error s -> diff --git a/riscV/Conventions1.v b/riscV/Conventions1.v index 5dd50a4c..6cb06c61 100644 --- a/riscV/Conventions1.v +++ b/riscV/Conventions1.v @@ -176,20 +176,21 @@ Qed. - RV64: pass the first 8 integer arguments in integer registers (a1...a8: int_param_regs), the first 8 FP arguments in FP registers - (fa1...fa8: float_param_regs), and the remaining arguments on the - stack, in 8-byte slots. + (fa1...fa8: float_param_regs) then in integer registers (a1...a8), + and the remaining arguments on the stack, in 8-byte slots. -- RV32: same, but arguments of 64-bit integer type are passed in two - consecutive integer registers (a(i), a(i+1)) or in a(8) and on a - 32-bit word on the stack. Stack-allocated arguments are aligned to - their natural alignment. +- RV32: same, but arguments of size 64 bits that must be passed in + integer registers are passed in two consecutive integer registers + (a(i), a(i+1)), or in a(8) and on a 32-bit word on the stack. + Stack-allocated arguments are aligned to their natural alignment. For variadic functions, the fixed arguments are passed as described above, then the variadic arguments receive special treatment: -- RV64: all variadic arguments, including FP arguments, - are passed in the remaining integer registers (a1...a8), - then on the stack, in 8-byte slots. +- RV64: FP registers are not used for passing variadic arguments. + All variadic arguments, including FP arguments, are passed in the + remaining integer registers (a1...a8), then on the stack, in 8-byte + slots. - RV32: likewise, but arguments of 64-bit types (integers as well as floats) are passed in two consecutive aligned integer registers @@ -207,6 +208,15 @@ Definition int_param_regs := Definition float_param_regs := F10 :: F11 :: F12 :: F13 :: F14 :: F15 :: F16 :: F17 :: nil. +(** To evaluate FP arguments that must be passed in integer registers, + we can use any FP caller-save register that is not already used to pass + a fixed FP argument. Since there are 8 integer registers for argument + passing, we need at most 8 extra more FP registers for these FP + arguments. *) + +Definition float_extra_param_regs := + F0 :: F1 :: F2 :: F3 :: F4 :: F5 :: F6 :: F7 :: nil. + Definition int_arg (ri rf ofs: Z) (ty: typ) (rec: Z -> Z -> Z -> list (rpair loc)) := match list_nth_z int_param_regs ri with @@ -220,26 +230,27 @@ Definition int_arg (ri rf ofs: Z) (ty: typ) Definition float_arg (va: bool) (ri rf ofs: Z) (ty: typ) (rec: Z -> Z -> Z -> list (rpair loc)) := - match list_nth_z float_param_regs rf with + match list_nth_z (if va then nil else float_param_regs) rf with | Some r => - if va then - (let ri' := (* reserve 1 or 2 aligned integer registers *) - if Archi.ptr64 || zeq (typesize ty) 1 then ri + 1 else align ri 2 + 2 in - if zle ri' 8 then - (* we have enough integer registers, put argument in FP reg - and fixup code will put it in one or two integer regs *) - One (R r) :: rec ri' (rf + 1) ofs - else - (* we are out of integer registers, pass argument on stack *) + One (R r) :: rec ri (rf + 1) ofs + | None => + (* We are out of FP registers, or cannot use them because vararg, + so try to put the argument in an extra FP register while + reserving an integer register or register pair into which + fixup code will move the extra FP register. *) + let regpair := negb Archi.ptr64 && zeq (typesize ty) 2 in + let ri' := if va && regpair then align ri 2 else ri in + match list_nth_z float_extra_param_regs ri' with + | Some r => + let ri'' := ri' + (if Archi.ptr64 then 1 else typesize ty) in + let ofs'' := if regpair && zeq ri' 7 then ofs + 1 else ofs in + One (R r) :: rec ri'' rf ofs'' + | None => + (* We are out of integer registers, pass argument on stack *) let ofs := align ofs (typesize ty) in One(S Outgoing ofs ty) - :: rec ri' rf (ofs + (if Archi.ptr64 then 2 else typesize ty))) - else - One (R r) :: rec ri (rf + 1) ofs - | None => - let ofs := align ofs (typesize ty) in - One(S Outgoing ofs ty) - :: rec ri rf (ofs + (if Archi.ptr64 then 2 else typesize ty)) + :: rec ri' rf (ofs + (if Archi.ptr64 then 2 else typesize ty)) + end end. Definition split_long_arg (va: bool) (ri rf ofs: Z) @@ -317,6 +328,8 @@ Proof. { decide_goal. } assert (CSF: forall r, In r float_param_regs -> is_callee_save r = false). { decide_goal. } + assert (CSFX: forall r, In r float_extra_param_regs -> is_callee_save r = false). + { decide_goal. } assert (AL: forall ofs ty, ofs >= 0 -> align ofs (typesize ty) >= 0). { intros. assert (ofs <= align ofs (typesize ty)) by (apply align_le; apply typesize_pos). @@ -341,18 +354,17 @@ Proof. assert (B: forall va ri rf ofs ty f, OKF f -> ofs >= 0 -> OK (float_arg va ri rf ofs ty f)). { intros until f; intros OF OO; red; unfold float_arg; intros. - destruct (list_nth_z float_param_regs rf) as [r|] eqn:NTH. - - set (ri' := if Archi.ptr64 || zeq (typesize ty) 1 then ri + 1 else align ri 2 + 2) in *. - destruct va; [destruct (zle ri' 8)|idtac]; destruct H. - + subst p; simpl. apply CSF. eapply list_nth_z_in; eauto. - + eapply OF; eauto. - + subst p; repeat split; auto. - + eapply OF; [idtac|eauto]. generalize (AL ofs ty OO) (SKK ty); lia. - + subst p; simpl. apply CSF. eapply list_nth_z_in; eauto. - + eapply OF; eauto. + destruct (list_nth_z (if va then nil else float_param_regs) rf) as [r|] eqn:NTH. - destruct H. - + subst p; repeat split; auto. - + eapply OF; [idtac|eauto]. generalize (AL ofs ty OO) (SKK ty); lia. + + subst p; simpl. apply CSF. destruct va. simpl in NTH; discriminate. eapply list_nth_z_in; eauto. + + eapply OF; eauto. + - set (regpair := negb Archi.ptr64 && zeq (typesize ty) 2) in *. + set (ri' := if va && regpair then align ri 2 else ri) in *. + destruct (list_nth_z float_extra_param_regs ri') as [r|] eqn:NTH'; destruct H. + + subst p; simpl. apply CSFX. eapply list_nth_z_in; eauto. + + eapply OF; [|eauto]. destruct (regpair && zeq ri' 7); lia. + + subst p; simpl. auto. + + eapply OF; [|eauto]. generalize (AL ofs ty OO) (SKK ty); lia. } assert (C: forall va ri rf ofs f, OKF f -> ofs >= 0 -> OK (split_long_arg va ri rf ofs f)). diff --git a/riscV/TargetPrinter.ml b/riscV/TargetPrinter.ml index 64bcea4c..5cd47b46 100644 --- a/riscV/TargetPrinter.ml +++ b/riscV/TargetPrinter.ml @@ -392,8 +392,12 @@ module Target : TARGET = fprintf oc " fmv.d %a, %a\n" freg fd freg fs | Pfmvxs (rd,fs) -> fprintf oc " fmv.x.s %a, %a\n" ireg rd freg fs + | Pfmvsx (fd,rs) -> + fprintf oc " fmv.s.x %a, %a\n" freg fd ireg rs | Pfmvxd (rd,fs) -> fprintf oc " fmv.x.d %a, %a\n" ireg rd freg fs + | Pfmvdx (fd,rs) -> + fprintf oc " fmv.d.x %a, %a\n" freg fd ireg rs (* 32-bit (single-precision) floating point *) | Pfls (fd, ra, ofs) -> -- cgit From 478ece46d8323ea182ded96a531309becf7445bb Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Sat, 16 Jan 2021 15:27:02 +0100 Subject: Support re-normalization of function parameters at function entry This is complementary to 28f235806 Some ABIs leave more flexibility concerning function parameters than CompCert expects. For instance, the AArch64/ELF ABI allow the caller of a function to leave unspecified the "padding bits" of function parameters. As an example, a parameter of type "unsigned char" may not have zeros in bits 8 to 63, but may have any bits there. When the caller is compiled by CompCert, it normalizes argument values to the parameter types before the call, so padding bits are always correct w.r.t. the type of the argument. This is no longer guaranteed in interoperability scenarios, when the caller is not compiled by CompCert. This commit adds a general mechanism to insert "re-normalization" conversions on the parameters of a function, at function entry. This is controlled by the platform-dependent function Convention1.return_value_needs_normalization. The semantic preservation proof is still conducted against the CompCert model, where the argument values of functions are already normalized. What the proof shows is that the extra conversions have no effect in this case. In future work we could relax the CompCert model, allowing functions to pass arguments that are not normalized. --- riscV/Conventions1.v | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) (limited to 'riscV') diff --git a/riscV/Conventions1.v b/riscV/Conventions1.v index 6cb06c61..eeaae3c4 100644 --- a/riscV/Conventions1.v +++ b/riscV/Conventions1.v @@ -413,8 +413,9 @@ Proof. reflexivity. Qed. -(** ** Normalization of function results *) +(** ** Normalization of function results and parameters *) (** No normalization needed. *) Definition return_value_needs_normalization (t: rettype) := false. +Definition parameter_needs_normalization (t: rettype) := false. -- cgit From fc82b6c80fd3feeb4ef9478e6faa16b5b1104593 Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Thu, 21 Jan 2021 15:44:09 +0100 Subject: Qualify `Hint` as `Global Hint` where appropriate This avoids a new warning of Coq 8.13. Eventually these `Global Hint` should become `#[export] Hint`, with a cleaner but different meaning than `Global Hint`. --- riscV/Asmgenproof1.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) (limited to 'riscV') diff --git a/riscV/Asmgenproof1.v b/riscV/Asmgenproof1.v index 253e769f..8195ce44 100644 --- a/riscV/Asmgenproof1.v +++ b/riscV/Asmgenproof1.v @@ -88,7 +88,7 @@ Proof. intros. apply ireg_of_not_X31 in H. congruence. Qed. -Hint Resolve ireg_of_not_X31 ireg_of_not_X31': asmgen. +Global Hint Resolve ireg_of_not_X31 ireg_of_not_X31': asmgen. (** Useful simplification tactic *) -- cgit From 30feb31c6d6e9235acad42ec5d09d14f3919cc36 Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Wed, 30 Dec 2020 11:41:10 +0100 Subject: Introduce and use PrintAsmaux.variable_section This is a generalization of the previous PrintAsmaux.common_section function that - handles initialized variables in addition to uninitialized variables; - can be used for Section_const, not just for Section_data. --- riscV/TargetPrinter.ml | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) (limited to 'riscV') diff --git a/riscV/TargetPrinter.ml b/riscV/TargetPrinter.ml index 5cd47b46..d8137f84 100644 --- a/riscV/TargetPrinter.ml +++ b/riscV/TargetPrinter.ml @@ -108,9 +108,9 @@ module Target : TARGET = let name_of_section = function | Section_text -> ".text" | Section_data i | Section_small_data i -> - if i then ".data" else common_section () + variable_section ~sec:".data" ~bss:".bss" i | Section_const i | Section_small_const i -> - if i || (not !Clflags.option_fcommon) then ".section .rodata" else "COMM" + variable_section ~sec:".section .rodata" i | Section_string -> ".section .rodata" | Section_literal -> ".section .rodata" | Section_jumptable -> ".section .rodata" -- cgit From 9a0bf569fab7398abd46bd07d2ee777fe745f591 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Mon, 29 Mar 2021 10:28:23 +0200 Subject: fix riscv merge? --- riscV/Asm.v | 1199 +++++++++++++++++++++++++++++++++ riscV/Asmgenproof1.v | 965 +++++++++++++++++++++++++++ riscV/SelectLongproof.v | 620 +++++++++++++++++ riscV/SelectOpproof.v | 1093 ++++++++++++++++++++++++++++++ riscV/TO_MERGE/Asm.v | 1214 --------------------------------- riscV/TO_MERGE/Asmgenproof1.v | 1367 -------------------------------------- riscV/TO_MERGE/SelectLongproof.v | 650 ------------------ riscV/TO_MERGE/SelectOpproof.v | 1124 ------------------------------- riscV/TO_MERGE/TargetPrinter.ml | 677 ------------------- riscV/TargetPrinter.ml | 667 +++++++++++++++++++ 10 files changed, 4544 insertions(+), 5032 deletions(-) create mode 100644 riscV/Asm.v create mode 100644 riscV/Asmgenproof1.v create mode 100644 riscV/SelectLongproof.v create mode 100644 riscV/SelectOpproof.v delete mode 100644 riscV/TO_MERGE/Asm.v delete mode 100644 riscV/TO_MERGE/Asmgenproof1.v delete mode 100644 riscV/TO_MERGE/SelectLongproof.v delete mode 100644 riscV/TO_MERGE/SelectOpproof.v delete mode 100644 riscV/TO_MERGE/TargetPrinter.ml create mode 100644 riscV/TargetPrinter.ml (limited to 'riscV') diff --git a/riscV/Asm.v b/riscV/Asm.v new file mode 100644 index 00000000..7e8ba760 --- /dev/null +++ b/riscV/Asm.v @@ -0,0 +1,1199 @@ +(* *********************************************************************) +(* *) +(* The Compcert verified compiler *) +(* *) +(* Xavier Leroy, INRIA Paris-Rocquencourt *) +(* Prashanth Mundkur, SRI International *) +(* *) +(* 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. *) +(* *) +(* The contributions by Prashanth Mundkur are reused and adapted *) +(* under the terms of a Contributor License Agreement between *) +(* SRI International and INRIA. *) +(* *) +(* *********************************************************************) + +(** Abstract syntax and semantics for RISC-V assembly language. *) + +Require Import Coqlib. +Require Import Maps. +Require Import AST. +Require Import Integers. +Require Import Floats. +Require Import Values. +Require Import Memory. +Require Import Events. +Require Import Globalenvs. +Require Import Smallstep. +Require Import Locations. +Require Stacklayout. +Require Import Conventions. +Require ExtValues. + +(** * Abstract syntax *) + +(** Integer registers. X0 is treated specially because it always reads + as zero and is never used as a destination of an instruction. *) + +Inductive ireg: Type := + | X1: ireg | X2: ireg | X3: ireg | X4: ireg | X5: ireg + | X6: ireg | X7: ireg | X8: ireg | X9: ireg | X10: ireg + | X11: ireg | X12: ireg | X13: ireg | X14: ireg | X15: ireg + | X16: ireg | X17: ireg | X18: ireg | X19: ireg | X20: ireg + | X21: ireg | X22: ireg | X23: ireg | X24: ireg | X25: ireg + | X26: ireg | X27: ireg | X28: ireg | X29: ireg | X30: ireg + | X31: ireg. + +Inductive ireg0: Type := + | X0: ireg0 | X: ireg -> ireg0. + +Coercion X: ireg >-> ireg0. + +(** Floating-point registers *) + +Inductive freg: Type := + | F0: freg | F1: freg | F2: freg | F3: freg + | F4: freg | F5: freg | F6: freg | F7: freg + | F8: freg | F9: freg | F10: freg | F11: freg + | F12: freg | F13: freg | F14: freg | F15: freg + | F16: freg | F17: freg | F18: freg | F19: freg + | F20: freg | F21: freg | F22: freg | F23: freg + | F24: freg | F25: freg | F26: freg | F27: freg + | F28: freg | F29: freg | F30: freg | F31: freg. + +Definition ireg_eq: forall (x y: ireg), {x=y} + {x<>y}. +Proof. decide equality. Defined. + +Definition ireg0_eq: forall (x y: ireg0), {x=y} + {x<>y}. +Proof. decide equality. apply ireg_eq. Defined. + +Lemma freg_eq: forall (x y: freg), {x=y} + {x<>y}. +Proof. decide equality. Defined. + +(** We model the following registers of the RISC-V architecture. *) + +Inductive preg: Type := + | IR: ireg -> preg (**r integer registers *) + | FR: freg -> preg (**r double-precision float registers *) + | PC: preg. (**r program counter *) + +Coercion IR: ireg >-> preg. +Coercion FR: freg >-> preg. + +Lemma preg_eq: forall (x y: preg), {x=y} + {x<>y}. +Proof. decide equality. apply ireg_eq. apply freg_eq. Defined. + +Module PregEq. + Definition t := preg. + Definition eq := preg_eq. +End PregEq. + +Module Pregmap := EMap(PregEq). + +(** Conventional names for stack pointer ([SP]) and return address ([RA]). *) + +Notation "'SP'" := X2 (only parsing) : asm. +Notation "'RA'" := X1 (only parsing) : asm. + +(** Offsets for load and store instructions. An offset is either an + immediate integer or the low part of a symbol. *) + +Inductive offset : Type := + | Ofsimm (ofs: ptrofs) + | Ofslow (id: ident) (ofs: ptrofs). + +(** The RISC-V instruction set is composed of several subsets. We model + the "32I" (32-bit integers), "64I" (64-bit integers), + "M" (multiplication and division), + "F" (single-precision floating-point) + and "D" (double-precision floating-point) subsets. + + For 32- and 64-bit integer arithmetic, the RISC-V instruction set comprises + generic integer operations such as ADD that operate over the full width + of an integer register (either 32 or 64 bit), plus specific instructions + such as ADDW that normalize their results to signed 32-bit integers. + Other instructions such as AND work equally well over 32- and 64-bit + integers, with the convention that 32-bit integers are represented + sign-extended in 64-bit registers. + + This clever design is challenging to formalize in the CompCert value + model. As a first step, we follow a more traditional approach, + also used in the x86 port, whereas we have two sets of (pseudo-) + instructions, one for 32-bit integer arithmetic, with suffix W, + the other for 64-bit integer arithmetic, with suffix L. The mapping + to actual instructions is done when printing assembly code, as follows: + - In 32-bit mode: + ADDW becomes ADD, ADDL is an error, ANDW becomes AND, ANDL is an error. + - In 64-bit mode: + ADDW becomes ADDW, ADDL becomes ADD, ANDW and ANDL both become AND. +*) + +Definition label := positive. + +(** A note on immediates: there are various constraints on immediate + operands to RISC-V instructions. We do not attempt to capture these + restrictions in the abstract syntax nor in the semantics. The + assembler will emit an error if immediate operands exceed the + representable range. Of course, our RISC-V generator (file + [Asmgen]) is careful to respect this range. *) + +Inductive instruction : Type := + | Pmv (rd: ireg) (rs: ireg) (**r integer move *) + +(** 32-bit integer register-immediate instructions *) + | Paddiw (rd: ireg) (rs: ireg0) (imm: int) (**r add immediate *) + | Psltiw (rd: ireg) (rs: ireg0) (imm: int) (**r set-less-than immediate *) + | Psltiuw (rd: ireg) (rs: ireg0) (imm: int) (**r set-less-than unsigned immediate *) + | Pandiw (rd: ireg) (rs: ireg0) (imm: int) (**r and immediate *) + | Poriw (rd: ireg) (rs: ireg0) (imm: int) (**r or immediate *) + | Pxoriw (rd: ireg) (rs: ireg0) (imm: int) (**r xor immediate *) + | Pslliw (rd: ireg) (rs: ireg0) (imm: int) (**r shift-left-logical immediate *) + | Psrliw (rd: ireg) (rs: ireg0) (imm: int) (**r shift-right-logical immediate *) + | Psraiw (rd: ireg) (rs: ireg0) (imm: int) (**r shift-right-arith immediate *) + | Pluiw (rd: ireg) (imm: int) (**r load upper-immediate *) +(** 32-bit integer register-register instructions *) + | Paddw (rd: ireg) (rs1 rs2: ireg0) (**r integer addition *) + | Psubw (rd: ireg) (rs1 rs2: ireg0) (**r integer subtraction *) + + | Pmulw (rd: ireg) (rs1 rs2: ireg0) (**r integer multiply low *) + | Pmulhw (rd: ireg) (rs1 rs2: ireg0) (**r integer multiply high signed *) + | Pmulhuw (rd: ireg) (rs1 rs2: ireg0) (**r integer multiply high unsigned *) + | Pdivw (rd: ireg) (rs1 rs2: ireg0) (**r integer division *) + | Pdivuw (rd: ireg) (rs1 rs2: ireg0) (**r unsigned integer division *) + | Premw (rd: ireg) (rs1 rs2: ireg0) (**r integer remainder *) + | Premuw (rd: ireg) (rs1 rs2: ireg0) (**r unsigned integer remainder *) + | Psltw (rd: ireg) (rs1 rs2: ireg0) (**r set-less-than *) + | Psltuw (rd: ireg) (rs1 rs2: ireg0) (**r set-less-than unsigned *) + | Pseqw (rd: ireg) (rs1 rs2: ireg0) (**r [rd <- rs1 == rs2] (pseudo) *) + | Psnew (rd: ireg) (rs1 rs2: ireg0) (**r [rd <- rs1 != rs2] (pseudo) *) + | Pandw (rd: ireg) (rs1 rs2: ireg0) (**r bitwise and *) + | Porw (rd: ireg) (rs1 rs2: ireg0) (**r bitwise or *) + | Pxorw (rd: ireg) (rs1 rs2: ireg0) (**r bitwise xor *) + | Psllw (rd: ireg) (rs1 rs2: ireg0) (**r shift-left-logical *) + | Psrlw (rd: ireg) (rs1 rs2: ireg0) (**r shift-right-logical *) + | Psraw (rd: ireg) (rs1 rs2: ireg0) (**r shift-right-arith *) + +(** 64-bit integer register-immediate instructions *) + | Paddil (rd: ireg) (rs: ireg0) (imm: int64) (**r add immediate *) + | Psltil (rd: ireg) (rs: ireg0) (imm: int64) (**r set-less-than immediate *) + | Psltiul (rd: ireg) (rs: ireg0) (imm: int64) (**r set-less-than unsigned immediate *) + | Pandil (rd: ireg) (rs: ireg0) (imm: int64) (**r and immediate *) + | Poril (rd: ireg) (rs: ireg0) (imm: int64) (**r or immediate *) + | Pxoril (rd: ireg) (rs: ireg0) (imm: int64) (**r xor immediate *) + | Psllil (rd: ireg) (rs: ireg0) (imm: int) (**r shift-left-logical immediate *) + | Psrlil (rd: ireg) (rs: ireg0) (imm: int) (**r shift-right-logical immediate *) + | Psrail (rd: ireg) (rs: ireg0) (imm: int) (**r shift-right-arith immediate *) + | Pluil (rd: ireg) (imm: int64) (**r load upper-immediate *) +(** 64-bit integer register-register instructions *) + | Paddl (rd: ireg) (rs1 rs2: ireg0) (**r integer addition *) + | Psubl (rd: ireg) (rs1 rs2: ireg0) (**r integer subtraction *) + + | Pmull (rd: ireg) (rs1 rs2: ireg0) (**r integer multiply low *) + | Pmulhl (rd: ireg) (rs1 rs2: ireg0) (**r integer multiply high signed *) + | Pmulhul (rd: ireg) (rs1 rs2: ireg0) (**r integer multiply high unsigned *) + | Pdivl (rd: ireg) (rs1 rs2: ireg0) (**r integer division *) + | Pdivul (rd: ireg) (rs1 rs2: ireg0) (**r unsigned integer division *) + | Preml (rd: ireg) (rs1 rs2: ireg0) (**r integer remainder *) + | Premul (rd: ireg) (rs1 rs2: ireg0) (**r unsigned integer remainder *) + | Psltl (rd: ireg) (rs1 rs2: ireg0) (**r set-less-than *) + | Psltul (rd: ireg) (rs1 rs2: ireg0) (**r set-less-than unsigned *) + | Pseql (rd: ireg) (rs1 rs2: ireg0) (**r [rd <- rs1 == rs2] (pseudo) *) + | Psnel (rd: ireg) (rs1 rs2: ireg0) (**r [rd <- rs1 != rs2] (pseudo) *) + | Pandl (rd: ireg) (rs1 rs2: ireg0) (**r bitwise and *) + | Porl (rd: ireg) (rs1 rs2: ireg0) (**r bitwise or *) + | Pxorl (rd: ireg) (rs1 rs2: ireg0) (**r bitwise xor *) + | Pslll (rd: ireg) (rs1 rs2: ireg0) (**r shift-left-logical *) + | Psrll (rd: ireg) (rs1 rs2: ireg0) (**r shift-right-logical *) + | Psral (rd: ireg) (rs1 rs2: ireg0) (**r shift-right-arith *) + + | Pcvtl2w (rd: ireg) (rs: ireg0) (**r int64->int32 (pseudo) *) + | Pcvtw2l (r: ireg) (**r int32 signed -> int64 (pseudo) *) + + (* Unconditional jumps. Links are always to X1/RA. *) + | Pj_l (l: label) (**r jump to label *) + | Pj_s (symb: ident) (sg: signature) (**r jump to symbol *) + | Pj_r (r: ireg) (sg: signature) (**r jump register *) + | Pjal_s (symb: ident) (sg: signature) (**r jump-and-link symbol *) + | Pjal_r (r: ireg) (sg: signature) (**r jump-and-link register *) + + (* Conditional branches, 32-bit comparisons *) + | Pbeqw (rs1 rs2: ireg0) (l: label) (**r branch-if-equal *) + | Pbnew (rs1 rs2: ireg0) (l: label) (**r branch-if-not-equal signed *) + | Pbltw (rs1 rs2: ireg0) (l: label) (**r branch-if-less signed *) + | Pbltuw (rs1 rs2: ireg0) (l: label) (**r branch-if-less unsigned *) + | Pbgew (rs1 rs2: ireg0) (l: label) (**r branch-if-greater-or-equal signed *) + | Pbgeuw (rs1 rs2: ireg0) (l: label) (**r branch-if-greater-or-equal unsigned *) + + (* Conditional branches, 64-bit comparisons *) + | Pbeql (rs1 rs2: ireg0) (l: label) (**r branch-if-equal *) + | Pbnel (rs1 rs2: ireg0) (l: label) (**r branch-if-not-equal signed *) + | Pbltl (rs1 rs2: ireg0) (l: label) (**r branch-if-less signed *) + | Pbltul (rs1 rs2: ireg0) (l: label) (**r branch-if-less unsigned *) + | Pbgel (rs1 rs2: ireg0) (l: label) (**r branch-if-greater-or-equal signed *) + | Pbgeul (rs1 rs2: ireg0) (l: label) (**r branch-if-greater-or-equal unsigned *) + + (* Loads and stores *) + | Plb (rd: ireg) (ra: ireg) (ofs: offset) (**r load signed int8 *) + | Plbu (rd: ireg) (ra: ireg) (ofs: offset) (**r load unsigned int8 *) + | Plh (rd: ireg) (ra: ireg) (ofs: offset) (**r load signed int16 *) + | Plhu (rd: ireg) (ra: ireg) (ofs: offset) (**r load unsigned int16 *) + | Plw (rd: ireg) (ra: ireg) (ofs: offset) (**r load int32 *) + | Plw_a (rd: ireg) (ra: ireg) (ofs: offset) (**r load any32 *) + | Pld (rd: ireg) (ra: ireg) (ofs: offset) (**r load int64 *) + | Pld_a (rd: ireg) (ra: ireg) (ofs: offset) (**r load any64 *) + + | Psb (rs: ireg) (ra: ireg) (ofs: offset) (**r store int8 *) + | Psh (rs: ireg) (ra: ireg) (ofs: offset) (**r store int16 *) + | Psw (rs: ireg) (ra: ireg) (ofs: offset) (**r store int32 *) + | Psw_a (rs: ireg) (ra: ireg) (ofs: offset) (**r store any32 *) + | Psd (rs: ireg) (ra: ireg) (ofs: offset) (**r store int64 *) + | Psd_a (rs: ireg) (ra: ireg) (ofs: offset) (**r store any64 *) + + (* Synchronization *) + | Pfence (**r fence *) + + (* floating point register move *) + | Pfmv (rd: freg) (rs: freg) (**r move *) + | Pfmvxs (rd: ireg) (rs: freg) (**r bitwise move FP single to integer register *) + | Pfmvxd (rd: ireg) (rs: freg) (**r bitwise move FP double to integer register *) + | Pfmvsx (rd: freg) (rs: ireg) (**r bitwise move integer register to FP single *) + | Pfmvdx (rd: freg) (rs: ireg) (**r bitwise move integer register to FP double*) + + (* 32-bit (single-precision) floating point *) + | Pfls (rd: freg) (ra: ireg) (ofs: offset) (**r load float *) + | Pfss (rs: freg) (ra: ireg) (ofs: offset) (**r store float *) + + | Pfnegs (rd: freg) (rs: freg) (**r negation *) + | Pfabss (rd: freg) (rs: freg) (**r absolute value *) + + | Pfadds (rd: freg) (rs1 rs2: freg) (**r addition *) + | Pfsubs (rd: freg) (rs1 rs2: freg) (**r subtraction *) + | Pfmuls (rd: freg) (rs1 rs2: freg) (**r multiplication *) + | Pfdivs (rd: freg) (rs1 rs2: freg) (**r division *) + | Pfmins (rd: freg) (rs1 rs2: freg) (**r minimum *) + | Pfmaxs (rd: freg) (rs1 rs2: freg) (**r maximum *) + + | Pfeqs (rd: ireg) (rs1 rs2: freg) (**r compare equal *) + | Pflts (rd: ireg) (rs1 rs2: freg) (**r compare less-than *) + | Pfles (rd: ireg) (rs1 rs2: freg) (**r compare less-than/equal *) + + | Pfsqrts (rd: freg) (rs: freg) (**r square-root *) + + | Pfmadds (rd: freg) (rs1 rs2 rs3: freg) (**r fused multiply-add *) + | Pfmsubs (rd: freg) (rs1 rs2 rs3: freg) (**r fused multiply-sub *) + | Pfnmadds (rd: freg) (rs1 rs2 rs3: freg) (**r fused negated multiply-add *) + | Pfnmsubs (rd: freg) (rs1 rs2 rs3: freg) (**r fused negated multiply-sub *) + + | Pfcvtws (rd: ireg) (rs: freg) (**r float32 -> int32 conversion *) + | Pfcvtwus (rd: ireg) (rs: freg) (**r float32 -> unsigned int32 conversion *) + | Pfcvtsw (rd: freg) (rs: ireg0) (**r int32 -> float32 conversion *) + | Pfcvtswu (rd: freg) (rs: ireg0) (**r unsigned int32 -> float32 conversion *) + + | Pfcvtls (rd: ireg) (rs: freg) (**r float32 -> int64 conversion *) + | Pfcvtlus (rd: ireg) (rs: freg) (**r float32 -> unsigned int64 conversion *) + | Pfcvtsl (rd: freg) (rs: ireg0) (**r int64 -> float32 conversion *) + | Pfcvtslu (rd: freg) (rs: ireg0) (**r unsigned int 64-> float32 conversion *) + + (* 64-bit (double-precision) floating point *) + | Pfld (rd: freg) (ra: ireg) (ofs: offset) (**r load 64-bit float *) + | Pfld_a (rd: freg) (ra: ireg) (ofs: offset) (**r load any64 *) + | Pfsd (rd: freg) (ra: ireg) (ofs: offset) (**r store 64-bit float *) + | Pfsd_a (rd: freg) (ra: ireg) (ofs: offset) (**r store any64 *) + + | Pfnegd (rd: freg) (rs: freg) (**r negation *) + | Pfabsd (rd: freg) (rs: freg) (**r absolute value *) + + | Pfaddd (rd: freg) (rs1 rs2: freg) (**r addition *) + | Pfsubd (rd: freg) (rs1 rs2: freg) (**r subtraction *) + | Pfmuld (rd: freg) (rs1 rs2: freg) (**r multiplication *) + | Pfdivd (rd: freg) (rs1 rs2: freg) (**r division *) + | Pfmind (rd: freg) (rs1 rs2: freg) (**r minimum *) + | Pfmaxd (rd: freg) (rs1 rs2: freg) (**r maximum *) + + | Pfeqd (rd: ireg) (rs1 rs2: freg) (**r compare equal *) + | Pfltd (rd: ireg) (rs1 rs2: freg) (**r compare less-than *) + | Pfled (rd: ireg) (rs1 rs2: freg) (**r compare less-than/equal *) + + | Pfsqrtd (rd: freg) (rs: freg) (**r square-root *) + + | Pfmaddd (rd: freg) (rs1 rs2 rs3: freg) (**r fused multiply-add *) + | Pfmsubd (rd: freg) (rs1 rs2 rs3: freg) (**r fused multiply-sub *) + | Pfnmaddd (rd: freg) (rs1 rs2 rs3: freg) (**r fused negated multiply-add *) + | Pfnmsubd (rd: freg) (rs1 rs2 rs3: freg) (**r fused negated multiply-sub *) + + | Pfcvtwd (rd: ireg) (rs: freg) (**r float -> int32 conversion *) + | Pfcvtwud (rd: ireg) (rs: freg) (**r float -> unsigned int32 conversion *) + | Pfcvtdw (rd: freg) (rs: ireg0) (**r int32 -> float conversion *) + | Pfcvtdwu (rd: freg) (rs: ireg0) (**r unsigned int32 -> float conversion *) + + | Pfcvtld (rd: ireg) (rs: freg) (**r float -> int64 conversion *) + | Pfcvtlud (rd: ireg) (rs: freg) (**r float -> unsigned int64 conversion *) + | Pfcvtdl (rd: freg) (rs: ireg0) (**r int64 -> float conversion *) + | Pfcvtdlu (rd: freg) (rs: ireg0) (**r unsigned int64 -> float conversion *) + + | Pfcvtds (rd: freg) (rs: freg) (**r float32 -> float *) + | Pfcvtsd (rd: freg) (rs: freg) (**r float -> float32 *) + + (* Pseudo-instructions *) + | Pallocframe (sz: Z) (pos: ptrofs) (**r allocate new stack frame *) + | Pfreeframe (sz: Z) (pos: ptrofs) (**r deallocate stack frame and restore previous frame *) + | Plabel (lbl: label) (**r define a code label *) + | Ploadsymbol (rd: ireg) (id: ident) (ofs: ptrofs) (**r load the address of a symbol *) + | Ploadsymbol_high (rd: ireg) (id: ident) (ofs: ptrofs) (**r load the high part of the address of a symbol *) + | Ploadli (rd: ireg) (i: int64) (**r load an immediate int64 *) + | Ploadfi (rd: freg) (f: float) (**r load an immediate float *) + | Ploadsi (rd: freg) (f: float32) (**r load an immediate single *) + | Pbtbl (r: ireg) (tbl: list label) (**r N-way branch through a jump table *) + | Pbuiltin: external_function -> list (builtin_arg preg) + -> builtin_res preg -> instruction (**r built-in function (pseudo) *) + | Pselectl (rd: ireg) (rb: ireg0) (rt: ireg0) (rf: ireg0) + | Pnop : instruction. (**r nop instruction *) + + +(** The pseudo-instructions are the following: + +- [Plabel]: define a code label at the current program point. + +- [Ploadsymbol]: load the address of a symbol in an integer register. + Expands to the [la] assembler pseudo-instruction, which does the right + thing even if we are in PIC mode. + +- [Ploadli rd ival]: load an immediate 64-bit integer into an integer + register rd. Expands to a load from an address in the constant data section, + using the integer register x31 as temporary: +<< + lui x31, %hi(lbl) + ld rd, %lo(lbl)(x31) +lbl: + .quad ival +>> + +- [Ploadfi rd fval]: similar to [Ploadli] but loads a double FP constant fval + into a float register rd. + +- [Ploadsi rd fval]: similar to [Ploadli] but loads a single FP constant fval + into a float register rd. + +- [Pallocframe sz pos]: in the formal semantics, this + pseudo-instruction allocates a memory block with bounds [0] and + [sz], stores the value of the stack pointer at offset [pos] in this + block, and sets the stack pointer to the address of the bottom of + this block. + In the printed ASM assembly code, this allocation is: +<< + mv x30, sp + sub sp, sp, #sz + sw x30, #pos(sp) +>> + This cannot be expressed in our memory model, which does not reflect + the fact that stack frames are adjacent and allocated/freed + following a stack discipline. + +- [Pfreeframe sz pos]: in the formal semantics, this pseudo-instruction + reads the word at [pos] of the block pointed by the stack pointer, + frees this block, and sets the stack pointer to the value read. + In the printed ASM assembly code, this freeing is just an increment of [sp]: +<< + add sp, sp, #sz +>> + Again, our memory model cannot comprehend that this operation + frees (logically) the current stack frame. + +- [Pbtbl reg table]: this is a N-way branch, implemented via a jump table + as follows: +<< + la x31, table + add x31, x31, reg + jr x31 +table: .long table[0], table[1], ... +>> + Note that [reg] contains 4 times the index of the desired table entry. + +- [Pseq rd rs1 rs2]: since unsigned comparisons have particular + semantics for pointers, we cannot encode equality directly using + xor/sub etc, which have only integer semantics. +<< + xor rd, rs1, rs2 + sltiu rd, rd, 1 +>> + The [xor] is omitted if one of [rs1] and [rs2] is [x0]. + +- [Psne rd rs1 rs2]: similarly for unsigned inequality. +<< + xor rd, rs1, rs2 + sltu rd, x0, rd +>> +*) + +Definition code := list instruction. +Record function : Type := mkfunction { fn_sig: signature; fn_code: code }. +Definition fundef := AST.fundef function. +Definition program := AST.program fundef unit. + +(** * Operational semantics *) + +(** The semantics operates over a single mapping from registers + (type [preg]) to values. We maintain + the convention that integer registers are mapped to values of + type [Tint] or [Tlong] (in 64 bit mode), + and float registers to values of type [Tsingle] or [Tfloat]. *) + +Definition regset := Pregmap.t val. +Definition genv := Genv.t fundef unit. + +Definition get0w (rs: regset) (r: ireg0) : val := + match r with + | X0 => Vint Int.zero + | X r => rs r + end. + +Definition get0l (rs: regset) (r: ireg0) : val := + match r with + | X0 => Vlong Int64.zero + | X r => rs r + end. + +Notation "a # b" := (a b) (at level 1, only parsing) : asm. +Notation "a ## b" := (get0w a b) (at level 1) : asm. +Notation "a ### b" := (get0l a b) (at level 1) : asm. +Notation "a # b <- c" := (Pregmap.set b c a) (at level 1, b at next level) : asm. + +Open Scope asm. + +(** Undefining some registers *) + +Fixpoint undef_regs (l: list preg) (rs: regset) : regset := + match l with + | nil => rs + | r :: l' => undef_regs l' (rs#r <- Vundef) + end. + +(** Assigning a register pair *) + +Definition set_pair (p: rpair preg) (v: val) (rs: regset) : regset := + match p with + | One r => rs#r <- v + | Twolong rhi rlo => rs#rhi <- (Val.hiword v) #rlo <- (Val.loword v) + end. + +(** Assigning multiple registers *) + +Fixpoint set_regs (rl: list preg) (vl: list val) (rs: regset) : regset := + match rl, vl with + | r1 :: rl', v1 :: vl' => set_regs rl' vl' (rs#r1 <- v1) + | _, _ => rs + end. + +(** Assigning the result of a builtin *) + +Fixpoint set_res (res: builtin_res preg) (v: val) (rs: regset) : regset := + match res with + | BR r => rs#r <- v + | BR_none => rs + | BR_splitlong hi lo => set_res lo (Val.loword v) (set_res hi (Val.hiword v) rs) + end. + +Section RELSEM. + +(** Looking up instructions in a code sequence by position. *) + +Fixpoint find_instr (pos: Z) (c: code) {struct c} : option instruction := + match c with + | nil => None + | i :: il => if zeq pos 0 then Some i else find_instr (pos - 1) il + end. + +(** Position corresponding to a label *) + +Definition is_label (lbl: label) (instr: instruction) : bool := + match instr with + | Plabel lbl' => if peq lbl lbl' then true else false + | _ => false + end. + +Lemma is_label_correct: + forall lbl instr, + if is_label lbl instr then instr = Plabel lbl else instr <> Plabel lbl. +Proof. + intros. destruct instr; simpl; try discriminate. + case (peq lbl lbl0); intro; congruence. +Qed. + +Fixpoint label_pos (lbl: label) (pos: Z) (c: code) {struct c} : option Z := + match c with + | nil => None + | instr :: c' => + if is_label lbl instr then Some (pos + 1) else label_pos lbl (pos + 1) c' + end. + +Variable ge: genv. + +(** The two functions below axiomatize how the linker processes + symbolic references [symbol + offset)] and splits their + actual values into a 20-bit high part [%hi(symbol + offset)] and + a 12-bit low part [%lo(symbol + offset)]. *) + +Parameter low_half: genv -> ident -> ptrofs -> ptrofs. +Parameter high_half: genv -> ident -> ptrofs -> val. + +(** The fundamental property of these operations is that, when applied + to the address of a symbol, their results can be recombined by + addition, rebuilding the original address. *) + +Axiom low_high_half: + forall id ofs, + Val.offset_ptr (high_half ge id ofs) (low_half ge id ofs) = Genv.symbol_address ge id ofs. + +(** The semantics is purely small-step and defined as a function + from the current state (a register set + a memory state) + to either [Next rs' m'] where [rs'] and [m'] are the updated register + set and memory state after execution of the instruction at [rs#PC], + or [Stuck] if the processor is stuck. *) + +Inductive outcome: Type := + | Next: regset -> mem -> outcome + | Stuck: outcome. + +(** Manipulations over the [PC] register: continuing with the next + instruction ([nextinstr]) or branching to a label ([goto_label]). *) + +Definition nextinstr (rs: regset) := + rs#PC <- (Val.offset_ptr rs#PC Ptrofs.one). + +Definition goto_label (f: function) (lbl: label) (rs: regset) (m: mem) := + match label_pos lbl 0 (fn_code f) with + | None => Stuck + | Some pos => + match rs#PC with + | Vptr b ofs => Next (rs#PC <- (Vptr b (Ptrofs.repr pos))) m + | _ => Stuck + end + end. + +(** Auxiliaries for memory accesses *) + +Definition eval_offset (ofs: offset) : ptrofs := + match ofs with + | Ofsimm n => n + | Ofslow id delta => low_half ge id delta + end. + +Definition exec_load (chunk: memory_chunk) (rs: regset) (m: mem) + (d: preg) (a: ireg) (ofs: offset) := + match Mem.loadv chunk m (Val.offset_ptr (rs a) (eval_offset ofs)) with + | None => Stuck + | Some v => Next (nextinstr (rs#d <- v)) m + end. + +Definition exec_store (chunk: memory_chunk) (rs: regset) (m: mem) + (s: preg) (a: ireg) (ofs: offset) := + match Mem.storev chunk m (Val.offset_ptr (rs a) (eval_offset ofs)) (rs s) with + | None => Stuck + | Some m' => Next (nextinstr rs) m' + end. + +(** Evaluating a branch *) + +Definition eval_branch (f: function) (l: label) (rs: regset) (m: mem) (res: option bool) : outcome := + match res with + | Some true => goto_label f l rs m + | Some false => Next (nextinstr rs) m + | None => Stuck + end. + +(** Execution of a single instruction [i] in initial state [rs] and + [m]. Return updated state. For instructions that correspond to + actual RISC-V instructions, the cases are straightforward + transliterations of the informal descriptions given in the RISC-V + user-mode specification. For pseudo-instructions, refer to the + informal descriptions given above. + + Note that we set to [Vundef] the registers used as temporaries by + the expansions of the pseudo-instructions, so that the RISC-V code + we generate cannot use those registers to hold values that must + survive the execution of the pseudo-instruction. *) + +Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : outcome := + match i with + | Pmv d s => + Next (nextinstr (rs#d <- (rs#s))) m + +(** 32-bit integer register-immediate instructions *) + | Paddiw d s i => + Next (nextinstr (rs#d <- (Val.add rs##s (Vint i)))) m + | Psltiw d s i => + Next (nextinstr (rs#d <- (Val.cmp Clt rs##s (Vint i)))) m + | Psltiuw d s i => + Next (nextinstr (rs#d <- (Val.cmpu (Mem.valid_pointer m) Clt rs##s (Vint i)))) m + | Pandiw d s i => + Next (nextinstr (rs#d <- (Val.and rs##s (Vint i)))) m + | Poriw d s i => + Next (nextinstr (rs#d <- (Val.or rs##s (Vint i)))) m + | Pxoriw d s i => + Next (nextinstr (rs#d <- (Val.xor rs##s (Vint i)))) m + | Pslliw d s i => + Next (nextinstr (rs#d <- (Val.shl rs##s (Vint i)))) m + | Psrliw d s i => + Next (nextinstr (rs#d <- (Val.shru rs##s (Vint i)))) m + | Psraiw d s i => + Next (nextinstr (rs#d <- (Val.shr rs##s (Vint i)))) m + | Pluiw d i => + Next (nextinstr (rs#d <- (Vint (Int.shl i (Int.repr 12))))) m + +(** 32-bit integer register-register instructions *) + | Paddw d s1 s2 => + Next (nextinstr (rs#d <- (Val.add rs##s1 rs##s2))) m + | Psubw d s1 s2 => + Next (nextinstr (rs#d <- (Val.sub rs##s1 rs##s2))) m + | Pmulw d s1 s2 => + Next (nextinstr (rs#d <- (Val.mul rs##s1 rs##s2))) m + | Pmulhw d s1 s2 => + Next (nextinstr (rs#d <- (Val.mulhs rs##s1 rs##s2))) m + | Pmulhuw d s1 s2 => + Next (nextinstr (rs#d <- (Val.mulhu rs##s1 rs##s2))) m + | Pdivw d s1 s2 => + Next (nextinstr (rs#d <- (Val.maketotal (Val.divs rs##s1 rs##s2)))) m + | Pdivuw d s1 s2 => + Next (nextinstr (rs#d <- (Val.maketotal (Val.divu rs##s1 rs##s2)))) m + | Premw d s1 s2 => + Next (nextinstr (rs#d <- (Val.maketotal (Val.mods rs##s1 rs##s2)))) m + | Premuw d s1 s2 => + Next (nextinstr (rs#d <- (Val.maketotal (Val.modu rs##s1 rs##s2)))) m + | Psltw d s1 s2 => + Next (nextinstr (rs#d <- (Val.cmp Clt rs##s1 rs##s2))) m + | Psltuw d s1 s2 => + Next (nextinstr (rs#d <- (Val.cmpu (Mem.valid_pointer m) Clt rs##s1 rs##s2))) m + | Pseqw d s1 s2 => + Next (nextinstr (rs#d <- (Val.cmpu (Mem.valid_pointer m) Ceq rs##s1 rs##s2))) m + | Psnew d s1 s2 => + Next (nextinstr (rs#d <- (Val.cmpu (Mem.valid_pointer m) Cne rs##s1 rs##s2))) m + | Pandw d s1 s2 => + Next (nextinstr (rs#d <- (Val.and rs##s1 rs##s2))) m + | Porw d s1 s2 => + Next (nextinstr (rs#d <- (Val.or rs##s1 rs##s2))) m + | Pxorw d s1 s2 => + Next (nextinstr (rs#d <- (Val.xor rs##s1 rs##s2))) m + | Psllw d s1 s2 => + Next (nextinstr (rs#d <- (Val.shl rs##s1 rs##s2))) m + | Psrlw d s1 s2 => + Next (nextinstr (rs#d <- (Val.shru rs##s1 rs##s2))) m + | Psraw d s1 s2 => + Next (nextinstr (rs#d <- (Val.shr rs##s1 rs##s2))) m + +(** 64-bit integer register-immediate instructions *) + | Paddil d s i => + Next (nextinstr (rs#d <- (Val.addl rs###s (Vlong i)))) m + | Psltil d s i => + Next (nextinstr (rs#d <- (Val.maketotal (Val.cmpl Clt rs###s (Vlong i))))) m + | Psltiul d s i => + Next (nextinstr (rs#d <- (Val.maketotal (Val.cmplu (Mem.valid_pointer m) Clt rs###s (Vlong i))))) m + | Pandil d s i => + Next (nextinstr (rs#d <- (Val.andl rs###s (Vlong i)))) m + | Poril d s i => + Next (nextinstr (rs#d <- (Val.orl rs###s (Vlong i)))) m + | Pxoril d s i => + Next (nextinstr (rs#d <- (Val.xorl rs###s (Vlong i)))) m + | Psllil d s i => + Next (nextinstr (rs#d <- (Val.shll rs###s (Vint i)))) m + | Psrlil d s i => + Next (nextinstr (rs#d <- (Val.shrlu rs###s (Vint i)))) m + | Psrail d s i => + Next (nextinstr (rs#d <- (Val.shrl rs###s (Vint i)))) m + | Pluil d i => + Next (nextinstr (rs#d <- (Vlong (Int64.sign_ext 32 (Int64.shl i (Int64.repr 12)))))) m + +(** 64-bit integer register-register instructions *) + | Paddl d s1 s2 => + Next (nextinstr (rs#d <- (Val.addl rs###s1 rs###s2))) m + | Psubl d s1 s2 => + Next (nextinstr (rs#d <- (Val.subl rs###s1 rs###s2))) m + | Pmull d s1 s2 => + Next (nextinstr (rs#d <- (Val.mull rs###s1 rs###s2))) m + | Pmulhl d s1 s2 => + Next (nextinstr (rs#d <- (Val.mullhs rs###s1 rs###s2))) m + | Pmulhul d s1 s2 => + Next (nextinstr (rs#d <- (Val.mullhu rs###s1 rs###s2))) m + | Pdivl d s1 s2 => + Next (nextinstr (rs#d <- (Val.maketotal (Val.divls rs###s1 rs###s2)))) m + | Pdivul d s1 s2 => + Next (nextinstr (rs#d <- (Val.maketotal (Val.divlu rs###s1 rs###s2)))) m + | Preml d s1 s2 => + Next (nextinstr (rs#d <- (Val.maketotal (Val.modls rs###s1 rs###s2)))) m + | Premul d s1 s2 => + Next (nextinstr (rs#d <- (Val.maketotal (Val.modlu rs###s1 rs###s2)))) m + | Psltl d s1 s2 => + Next (nextinstr (rs#d <- (Val.maketotal (Val.cmpl Clt rs###s1 rs###s2)))) m + | Psltul d s1 s2 => + Next (nextinstr (rs#d <- (Val.maketotal (Val.cmplu (Mem.valid_pointer m) Clt rs###s1 rs###s2)))) m + | Pseql d s1 s2 => + Next (nextinstr (rs#d <- (Val.maketotal (Val.cmplu (Mem.valid_pointer m) Ceq rs###s1 rs###s2)))) m + | Psnel d s1 s2 => + Next (nextinstr (rs#d <- (Val.maketotal (Val.cmplu (Mem.valid_pointer m) Cne rs###s1 rs###s2)))) m + | Pandl d s1 s2 => + Next (nextinstr (rs#d <- (Val.andl rs###s1 rs###s2))) m + | Porl d s1 s2 => + Next (nextinstr (rs#d <- (Val.orl rs###s1 rs###s2))) m + | Pxorl d s1 s2 => + Next (nextinstr (rs#d <- (Val.xorl rs###s1 rs###s2))) m + | Pslll d s1 s2 => + Next (nextinstr (rs#d <- (Val.shll rs###s1 rs###s2))) m + | Psrll d s1 s2 => + Next (nextinstr (rs#d <- (Val.shrlu rs###s1 rs###s2))) m + | Psral d s1 s2 => + Next (nextinstr (rs#d <- (Val.shrl rs###s1 rs###s2))) m + + | Pcvtl2w d s => + Next (nextinstr (rs#d <- (Val.loword rs##s))) m + | Pcvtw2l r => + Next (nextinstr (rs#r <- (Val.longofint rs#r))) m + +(** Unconditional jumps. Links are always to X1/RA. *) + | Pj_l l => + goto_label f l rs m + | Pj_s s sg => + Next (rs#PC <- (Genv.symbol_address ge s Ptrofs.zero)) m + | Pj_r r sg => + Next (rs#PC <- (rs#r)) m + | Pjal_s s sg => + Next (rs#PC <- (Genv.symbol_address ge s Ptrofs.zero) + #RA <- (Val.offset_ptr rs#PC Ptrofs.one) + ) m + | Pjal_r r sg => + Next (rs#PC <- (rs#r) + #RA <- (Val.offset_ptr rs#PC Ptrofs.one) + ) m +(** Conditional branches, 32-bit comparisons *) + | Pbeqw s1 s2 l => + eval_branch f l rs m (Val.cmpu_bool (Mem.valid_pointer m) Ceq rs##s1 rs##s2) + | Pbnew s1 s2 l => + eval_branch f l rs m (Val.cmpu_bool (Mem.valid_pointer m) Cne rs##s1 rs##s2) + | Pbltw s1 s2 l => + eval_branch f l rs m (Val.cmp_bool Clt rs##s1 rs##s2) + | Pbltuw s1 s2 l => + eval_branch f l rs m (Val.cmpu_bool (Mem.valid_pointer m) Clt rs##s1 rs##s2) + | Pbgew s1 s2 l => + eval_branch f l rs m (Val.cmp_bool Cge rs##s1 rs##s2) + | Pbgeuw s1 s2 l => + eval_branch f l rs m (Val.cmpu_bool (Mem.valid_pointer m) Cge rs##s1 rs##s2) + +(** Conditional branches, 64-bit comparisons *) + | Pbeql s1 s2 l => + eval_branch f l rs m (Val.cmplu_bool (Mem.valid_pointer m) Ceq rs###s1 rs###s2) + | Pbnel s1 s2 l => + eval_branch f l rs m (Val.cmplu_bool (Mem.valid_pointer m) Cne rs###s1 rs###s2) + | Pbltl s1 s2 l => + eval_branch f l rs m (Val.cmpl_bool Clt rs###s1 rs###s2) + | Pbltul s1 s2 l => + eval_branch f l rs m (Val.cmplu_bool (Mem.valid_pointer m) Clt rs###s1 rs###s2) + | Pbgel s1 s2 l => + eval_branch f l rs m (Val.cmpl_bool Cge rs###s1 rs###s2) + | Pbgeul s1 s2 l => + eval_branch f l rs m (Val.cmplu_bool (Mem.valid_pointer m) Cge rs###s1 rs###s2) + +(** Loads and stores *) + | Plb d a ofs => + exec_load Mint8signed rs m d a ofs + | Plbu d a ofs => + exec_load Mint8unsigned rs m d a ofs + | Plh d a ofs => + exec_load Mint16signed rs m d a ofs + | Plhu d a ofs => + exec_load Mint16unsigned rs m d a ofs + | Plw d a ofs => + exec_load Mint32 rs m d a ofs + | Plw_a d a ofs => + exec_load Many32 rs m d a ofs + | Pld d a ofs => + exec_load Mint64 rs m d a ofs + | Pld_a d a ofs => + exec_load Many64 rs m d a ofs + | Psb s a ofs => + exec_store Mint8unsigned rs m s a ofs + | Psh s a ofs => + exec_store Mint16unsigned rs m s a ofs + | Psw s a ofs => + exec_store Mint32 rs m s a ofs + | Psw_a s a ofs => + exec_store Many32 rs m s a ofs + | Psd s a ofs => + exec_store Mint64 rs m s a ofs + | Psd_a s a ofs => + exec_store Many64 rs m s a ofs + +(** Floating point register move *) + | Pfmv d s => + Next (nextinstr (rs#d <- (rs#s))) m + +(** 32-bit (single-precision) floating point *) + | Pfls d a ofs => + exec_load Mfloat32 rs m d a ofs + | Pfss s a ofs => + exec_store Mfloat32 rs m s a ofs + + | Pfnegs d s => + Next (nextinstr (rs#d <- (Val.negfs rs#s))) m + | Pfabss d s => + Next (nextinstr (rs#d <- (Val.absfs rs#s))) m + + | Pfadds d s1 s2 => + Next (nextinstr (rs#d <- (Val.addfs rs#s1 rs#s2))) m + | Pfsubs d s1 s2 => + Next (nextinstr (rs#d <- (Val.subfs rs#s1 rs#s2))) m + | Pfmuls d s1 s2 => + Next (nextinstr (rs#d <- (Val.mulfs rs#s1 rs#s2))) m + | Pfdivs d s1 s2 => + Next (nextinstr (rs#d <- (Val.divfs rs#s1 rs#s2))) m + | Pfeqs d s1 s2 => + Next (nextinstr (rs#d <- (Val.cmpfs Ceq rs#s1 rs#s2))) m + | Pflts d s1 s2 => + Next (nextinstr (rs#d <- (Val.cmpfs Clt rs#s1 rs#s2))) m + | Pfles d s1 s2 => + Next (nextinstr (rs#d <- (Val.cmpfs Cle rs#s1 rs#s2))) m + + | Pfcvtws d s => + Next (nextinstr (rs#d <- (Val.maketotal (Val.intofsingle rs#s)))) m + | Pfcvtwus d s => + Next (nextinstr (rs#d <- (Val.maketotal (Val.intuofsingle rs#s)))) m + | Pfcvtsw d s => + Next (nextinstr (rs#d <- (Val.maketotal (Val.singleofint rs##s)))) m + | Pfcvtswu d s => + Next (nextinstr (rs#d <- (Val.maketotal (Val.singleofintu rs##s)))) m + + | Pfcvtls d s => + Next (nextinstr (rs#d <- (Val.maketotal (Val.longofsingle rs#s)))) m + | Pfcvtlus d s => + Next (nextinstr (rs#d <- (Val.maketotal (Val.longuofsingle rs#s)))) m + | Pfcvtsl d s => + Next (nextinstr (rs#d <- (Val.maketotal (Val.singleoflong rs###s)))) m + | Pfcvtslu d s => + Next (nextinstr (rs#d <- (Val.maketotal (Val.singleoflongu rs###s)))) m + +(** 64-bit (double-precision) floating point *) + | Pfld d a ofs => + exec_load Mfloat64 rs m d a ofs + | Pfld_a d a ofs => + exec_load Many64 rs m d a ofs + | Pfsd s a ofs => + exec_store Mfloat64 rs m s a ofs + | Pfsd_a s a ofs => + exec_store Many64 rs m s a ofs + + | Pfnegd d s => + Next (nextinstr (rs#d <- (Val.negf rs#s))) m + | Pfabsd d s => + Next (nextinstr (rs#d <- (Val.absf rs#s))) m + + | Pfaddd d s1 s2 => + Next (nextinstr (rs#d <- (Val.addf rs#s1 rs#s2))) m + | Pfsubd d s1 s2 => + Next (nextinstr (rs#d <- (Val.subf rs#s1 rs#s2))) m + | Pfmuld d s1 s2 => + Next (nextinstr (rs#d <- (Val.mulf rs#s1 rs#s2))) m + | Pfdivd d s1 s2 => + Next (nextinstr (rs#d <- (Val.divf rs#s1 rs#s2))) m + | Pfeqd d s1 s2 => + Next (nextinstr (rs#d <- (Val.cmpf Ceq rs#s1 rs#s2))) m + | Pfltd d s1 s2 => + Next (nextinstr (rs#d <- (Val.cmpf Clt rs#s1 rs#s2))) m + | Pfled d s1 s2 => + Next (nextinstr (rs#d <- (Val.cmpf Cle rs#s1 rs#s2))) m + + | Pfcvtwd d s => + Next (nextinstr (rs#d <- (Val.maketotal (Val.intoffloat rs#s)))) m + | Pfcvtwud d s => + Next (nextinstr (rs#d <- (Val.maketotal (Val.intuoffloat rs#s)))) m + | Pfcvtdw d s => + Next (nextinstr (rs#d <- (Val.maketotal (Val.floatofint rs##s)))) m + | Pfcvtdwu d s => + Next (nextinstr (rs#d <- (Val.maketotal (Val.floatofintu rs##s)))) m + + | Pfcvtld d s => + Next (nextinstr (rs#d <- (Val.maketotal (Val.longoffloat rs#s)))) m + | Pfcvtlud d s => + Next (nextinstr (rs#d <- (Val.maketotal (Val.longuoffloat rs#s)))) m + | Pfcvtdl d s => + Next (nextinstr (rs#d <- (Val.maketotal (Val.floatoflong rs###s)))) m + | Pfcvtdlu d s => + Next (nextinstr (rs#d <- (Val.maketotal (Val.floatoflongu rs###s)))) m + + | Pfcvtds d s => + Next (nextinstr (rs#d <- (Val.floatofsingle rs#s))) m + | Pfcvtsd d s => + Next (nextinstr (rs#d <- (Val.singleoffloat rs#s))) m + + | Pfmvxs d s => + Next (nextinstr (rs#d <- (ExtValues.bits_of_single rs#s))) m + | Pfmvxd d s => + Next (nextinstr (rs#d <- (ExtValues.bits_of_float rs#s))) m + + | Pfmvsx d s => + Next (nextinstr (rs#d <- (ExtValues.single_of_bits rs#s))) m + | Pfmvdx d s => + Next (nextinstr (rs#d <- (ExtValues.float_of_bits rs#s))) m + + +(** Pseudo-instructions *) + | Pallocframe sz pos => + let (m1, stk) := Mem.alloc m 0 sz in + let sp := (Vptr stk Ptrofs.zero) in + match Mem.storev Mptr m1 (Val.offset_ptr sp pos) rs#SP with + | None => Stuck + | Some m2 => Next (nextinstr (rs #X30 <- (rs SP) #SP <- sp #X31 <- Vundef)) m2 + end + | Pfreeframe sz pos => + match Mem.loadv Mptr m (Val.offset_ptr rs#SP pos) with + | None => Stuck + | Some v => + match rs SP with + | Vptr stk ofs => + match Mem.free m stk 0 sz with + | None => Stuck + | Some m' => Next (nextinstr (rs#SP <- v #X31 <- Vundef)) m' + end + | _ => Stuck + end + end + | Pselectl rd rb rt rf => + Next (nextinstr (rs#rd <- (ExtValues.select01_long + (rs###rb) (rs###rt) (rs###rf))) + #X31 <- Vundef) m + | Plabel lbl => + Next (nextinstr rs) m + | Ploadsymbol rd s ofs => + Next (nextinstr (rs#rd <- (Genv.symbol_address ge s ofs))) m + | Ploadsymbol_high rd s ofs => + Next (nextinstr (rs#rd <- (high_half ge s ofs))) m + | Ploadli rd i => + Next (nextinstr (rs#X31 <- Vundef #rd <- (Vlong i))) m + | Ploadfi rd f => + Next (nextinstr (rs#X31 <- Vundef #rd <- (Vfloat f))) m + | Ploadsi rd f => + Next (nextinstr (rs#X31 <- Vundef #rd <- (Vsingle f))) m + | Pbtbl r tbl => + match rs r with + | Vint n => + match list_nth_z tbl (Int.unsigned n) with + | None => Stuck + | Some lbl => goto_label f lbl (rs#X5 <- Vundef #X31 <- Vundef) m + end + | _ => Stuck + end + | Pbuiltin ef args res => + Stuck (**r treated specially below *) + + (** The following instructions and directives are not generated directly by Asmgen, + so we do not model them. *) + | Pfence + + | Pfmins _ _ _ + | Pfmaxs _ _ _ + | Pfsqrts _ _ + | Pfmadds _ _ _ _ + | Pfmsubs _ _ _ _ + | Pfnmadds _ _ _ _ + | Pfnmsubs _ _ _ _ + + | Pfmind _ _ _ + | Pfmaxd _ _ _ + | Pfsqrtd _ _ + | Pfmaddd _ _ _ _ + | Pfmsubd _ _ _ _ + | Pfnmaddd _ _ _ _ + | Pfnmsubd _ _ _ _ + | Pnop + => Stuck + end. + +(** Translation of the LTL/Linear/Mach view of machine registers to + the RISC-V view. Note that no LTL register maps to [X31]. This + register is reserved as temporary, to be used by the generated RV32G + code. *) + +Definition preg_of (r: mreg) : preg := + match r with + | R5 => X5 | R6 => X6 | R7 => X7 + | R8 => X8 | R9 => X9 | R10 => X10 | R11 => X11 + | R12 => X12 | R13 => X13 | R14 => X14 | R15 => X15 + | R16 => X16 | R17 => X17 | R18 => X18 | R19 => X19 + | R20 => X20 | R21 => X21 | R22 => X22 | R23 => X23 + | R24 => X24 | R25 => X25 | R26 => X26 | R27 => X27 + | R28 => X28 | R29 => X29 | R30 => X30 + + | Machregs.F0 => F0 | Machregs.F1 => F1 | Machregs.F2 => F2 | Machregs.F3 => F3 + | Machregs.F4 => F4 | Machregs.F5 => F5 | Machregs.F6 => F6 | Machregs.F7 => F7 + | Machregs.F8 => F8 | Machregs.F9 => F9 | Machregs.F10 => F10 | Machregs.F11 => F11 + | Machregs.F12 => F12 | Machregs.F13 => F13 | Machregs.F14 => F14 | Machregs.F15 => F15 + | Machregs.F16 => F16 | Machregs.F17 => F17 | Machregs.F18 => F18 | Machregs.F19 => F19 + | Machregs.F20 => F20 | Machregs.F21 => F21 | Machregs.F22 => F22 | Machregs.F23 => F23 + | Machregs.F24 => F24 | Machregs.F25 => F25 | Machregs.F26 => F26 | Machregs.F27 => F27 + | Machregs.F28 => F28 | Machregs.F29 => F29 | Machregs.F30 => F30 | Machregs.F31 => F31 + end. + +(** Undefine all registers except SP and callee-save registers *) + +Definition undef_caller_save_regs (rs: regset) : regset := + fun r => + if preg_eq r SP + || In_dec preg_eq r (List.map preg_of (List.filter is_callee_save all_mregs)) + then rs r + else Vundef. + +(** Extract the values of the arguments of an external call. + We exploit the calling conventions from module [Conventions], except that + we use RISC-V registers instead of locations. *) + +Inductive extcall_arg (rs: regset) (m: mem): loc -> val -> Prop := + | extcall_arg_reg: forall r, + extcall_arg rs m (R r) (rs (preg_of r)) + | extcall_arg_stack: forall ofs ty bofs v, + bofs = Stacklayout.fe_ofs_arg + 4 * ofs -> + Mem.loadv (chunk_of_type ty) m + (Val.offset_ptr rs#SP (Ptrofs.repr bofs)) = Some v -> + extcall_arg rs m (S Outgoing ofs ty) v. + +Inductive extcall_arg_pair (rs: regset) (m: mem): rpair loc -> val -> Prop := + | extcall_arg_one: forall l v, + extcall_arg rs m l v -> + extcall_arg_pair rs m (One l) v + | extcall_arg_twolong: forall hi lo vhi vlo, + extcall_arg rs m hi vhi -> + extcall_arg rs m lo vlo -> + extcall_arg_pair rs m (Twolong hi lo) (Val.longofwords vhi vlo). + +Definition extcall_arguments + (rs: regset) (m: mem) (sg: signature) (args: list val) : Prop := + list_forall2 (extcall_arg_pair rs m) (loc_arguments sg) args. + +Definition loc_external_result (sg: signature) : rpair preg := + map_rpair preg_of (loc_result sg). + +(** Execution of the instruction at [rs PC]. *) + +Inductive state: Type := + | State: regset -> mem -> state. + +Inductive step: state -> trace -> state -> Prop := + | exec_step_internal: + forall b ofs f i rs m rs' m', + rs PC = Vptr b ofs -> + Genv.find_funct_ptr ge b = Some (Internal f) -> + find_instr (Ptrofs.unsigned ofs) (fn_code f) = Some i -> + exec_instr f i rs m = Next rs' m' -> + step (State rs m) E0 (State rs' m') + | exec_step_builtin: + forall b ofs f ef args res rs m vargs t vres rs' m', + rs PC = Vptr b ofs -> + Genv.find_funct_ptr ge b = Some (Internal f) -> + find_instr (Ptrofs.unsigned ofs) f.(fn_code) = Some (Pbuiltin ef args res) -> + eval_builtin_args ge rs (rs SP) m args vargs -> + external_call ef ge vargs m t vres m' -> + rs' = nextinstr + (set_res res vres + (undef_regs (map preg_of (destroyed_by_builtin ef)) + (rs#X31 <- Vundef))) -> + step (State rs m) t (State rs' m') + | exec_step_external: + forall b ef args res rs m t rs' m', + rs PC = Vptr b Ptrofs.zero -> + Genv.find_funct_ptr ge b = Some (External ef) -> + external_call ef ge args m t res m' -> + extcall_arguments rs m (ef_sig ef) args -> + rs' = (set_pair (loc_external_result (ef_sig ef) ) res (undef_caller_save_regs rs))#PC <- (rs RA) -> + step (State rs m) t (State rs' m'). + +End RELSEM. + +(** Execution of whole programs. *) + +Inductive initial_state (p: program): state -> Prop := + | initial_state_intro: forall m0, + let ge := Genv.globalenv p in + let rs0 := + (Pregmap.init Vundef) + # PC <- (Genv.symbol_address ge p.(prog_main) Ptrofs.zero) + # SP <- Vnullptr + # RA <- Vnullptr in + Genv.init_mem p = Some m0 -> + initial_state p (State rs0 m0). + +Inductive final_state: state -> int -> Prop := + | final_state_intro: forall rs m r, + rs PC = Vnullptr -> + rs X10 = Vint r -> + final_state (State rs m) r. + +Definition semantics (p: program) := + Semantics step (initial_state p) final_state (Genv.globalenv p). + +(** Determinacy of the [Asm] semantics. *) + +Remark extcall_arguments_determ: + forall rs m sg args1 args2, + extcall_arguments rs m sg args1 -> extcall_arguments rs m sg args2 -> args1 = args2. +Proof. + intros until m. + assert (A: forall l v1 v2, + extcall_arg rs m l v1 -> extcall_arg rs m l v2 -> v1 = v2). + { intros. inv H; inv H0; congruence. } + assert (B: forall p v1 v2, + extcall_arg_pair rs m p v1 -> extcall_arg_pair rs m p v2 -> v1 = v2). + { intros. inv H; inv H0. + eapply A; eauto. + f_equal; eapply A; eauto. } + assert (C: forall ll vl1, list_forall2 (extcall_arg_pair rs m) ll vl1 -> + forall vl2, list_forall2 (extcall_arg_pair rs m) ll vl2 -> vl1 = vl2). + { + induction 1; intros vl2 EA; inv EA. + auto. + f_equal; eauto. } + intros. eapply C; eauto. +Qed. + +Lemma semantics_determinate: forall p, determinate (semantics p). +Proof. +Ltac Equalities := + match goal with + | [ H1: ?a = ?b, H2: ?a = ?c |- _ ] => + rewrite H1 in H2; inv H2; Equalities + | _ => idtac + end. + intros; constructor; simpl; intros. +- (* determ *) + inv H; inv H0; Equalities. + split. constructor. auto. + discriminate. + discriminate. + assert (vargs0 = vargs) by (eapply eval_builtin_args_determ; eauto). subst vargs0. + exploit external_call_determ. eexact H5. eexact H11. intros [A B]. + split. auto. intros. destruct B; auto. subst. auto. + assert (args0 = args) by (eapply extcall_arguments_determ; eauto). subst args0. + exploit external_call_determ. eexact H3. eexact H8. intros [A B]. + split. auto. intros. destruct B; auto. subst. auto. +- (* trace length *) + red; intros. inv H; simpl. + lia. + eapply external_call_trace_length; eauto. + eapply external_call_trace_length; eauto. +- (* initial states *) + inv H; inv H0. f_equal. congruence. +- (* final no step *) + assert (NOTNULL: forall b ofs, Vnullptr <> Vptr b ofs). + { intros; unfold Vnullptr; destruct Archi.ptr64; congruence. } + inv H. unfold Vzero in H0. red; intros; red; intros. + inv H; rewrite H0 in *; eelim NOTNULL; eauto. +- (* final states *) + inv H; inv H0. congruence. +Qed. + +(** Classification functions for processor registers (used in Asmgenproof). *) + +Definition data_preg (r: preg) : bool := + match r with + | IR RA => false + | IR X31 => false + | IR _ => true + | FR _ => true + | PC => false + end. diff --git a/riscV/Asmgenproof1.v b/riscV/Asmgenproof1.v new file mode 100644 index 00000000..379b5789 --- /dev/null +++ b/riscV/Asmgenproof1.v @@ -0,0 +1,965 @@ +(* *********************************************************************) +(* *) +(* The Compcert verified compiler *) +(* *) +(* Xavier Leroy, INRIA Paris-Rocquencourt *) +(* Prashanth Mundkur, SRI International *) +(* *) +(* 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. *) +(* *) +(* The contributions by Prashanth Mundkur are reused and adapted *) +(* under the terms of a Contributor License Agreement between *) +(* SRI International and INRIA. *) +(* *) +(* *********************************************************************) + +Require Import Coqlib Errors Maps. +Require Import AST Zbits Integers Floats Values Memory Globalenvs. +Require Import Op Locations Mach Conventions. +Require Import Asm Asmgen Asmgenproof0. + +(** Decomposition of integer constants. *) + +Lemma make_immed32_sound: + forall n, + match make_immed32 n with + | Imm32_single imm => n = imm + | Imm32_pair hi lo => n = Int.add (Int.shl hi (Int.repr 12)) lo + end. +Proof. + intros; unfold make_immed32. set (lo := Int.sign_ext 12 n). + predSpec Int.eq Int.eq_spec n lo. +- auto. +- set (m := Int.sub n lo). + assert (A: eqmod (two_p 12) (Int.unsigned lo) (Int.unsigned n)) by (apply Int.eqmod_sign_ext'; compute; auto). + assert (B: eqmod (two_p 12) (Int.unsigned n - Int.unsigned lo) 0). + { replace 0 with (Int.unsigned n - Int.unsigned n) by lia. + auto using eqmod_sub, eqmod_refl. } + assert (C: eqmod (two_p 12) (Int.unsigned m) 0). + { apply eqmod_trans with (Int.unsigned n - Int.unsigned lo); auto. + apply eqmod_divides with Int.modulus. apply Int.eqm_sym; apply Int.eqm_unsigned_repr. + exists (two_p (32-12)); auto. } + assert (D: Int.modu m (Int.repr 4096) = Int.zero). + { apply eqmod_mod_eq in C. unfold Int.modu. + change (Int.unsigned (Int.repr 4096)) with (two_p 12). rewrite C. + reflexivity. + apply two_p_gt_ZERO; lia. } + rewrite <- (Int.divu_pow2 m (Int.repr 4096) (Int.repr 12)) by auto. + rewrite Int.shl_mul_two_p. + change (two_p (Int.unsigned (Int.repr 12))) with 4096. + replace (Int.mul (Int.divu m (Int.repr 4096)) (Int.repr 4096)) with m. + unfold m. rewrite Int.sub_add_opp. rewrite Int.add_assoc. rewrite <- (Int.add_commut lo). + rewrite Int.add_neg_zero. rewrite Int.add_zero. auto. + rewrite (Int.modu_divu_Euclid m (Int.repr 4096)) at 1 by (vm_compute; congruence). + rewrite D. apply Int.add_zero. +Qed. + +Lemma make_immed64_sound: + forall n, + match make_immed64 n with + | Imm64_single imm => n = imm + | Imm64_pair hi lo => n = Int64.add (Int64.sign_ext 32 (Int64.shl hi (Int64.repr 12))) lo + | Imm64_large imm => n = imm + end. +Proof. + intros; unfold make_immed64. set (lo := Int64.sign_ext 12 n). + predSpec Int64.eq Int64.eq_spec n lo. +- auto. +- set (m := Int64.sub n lo). + set (p := Int64.zero_ext 20 (Int64.shru m (Int64.repr 12))). + predSpec Int64.eq Int64.eq_spec n (Int64.add (Int64.sign_ext 32 (Int64.shl p (Int64.repr 12))) lo). + auto. + auto. +Qed. + +(** Properties of registers *) + +Lemma ireg_of_not_X31: + forall m r, ireg_of m = OK r -> IR r <> IR X31. +Proof. + intros. erewrite <- ireg_of_eq; eauto with asmgen. +Qed. + +Lemma ireg_of_not_X31': + forall m r, ireg_of m = OK r -> r <> X31. +Proof. + intros. apply ireg_of_not_X31 in H. congruence. +Qed. + +Global Hint Resolve ireg_of_not_X31 ireg_of_not_X31': asmgen. + +(** Useful simplification tactic *) + +Ltac Simplif := + ((rewrite nextinstr_inv by eauto with asmgen) + || (rewrite nextinstr_inv1 by eauto with asmgen) + || (rewrite Pregmap.gss) + || (rewrite nextinstr_pc) + || (rewrite Pregmap.gso by eauto with asmgen)); auto with asmgen. + +Ltac Simpl := repeat Simplif. + +(** * Correctness of RISC-V constructor functions *) + +Section CONSTRUCTORS. + +Variable ge: genv. +Variable fn: function. + +(** 32-bit integer constants and arithmetic *) + +Lemma load_hilo32_correct: + forall rd hi lo k rs m, + exists rs', + exec_straight ge fn (load_hilo32 rd hi lo k) rs m k rs' m + /\ rs'#rd = Vint (Int.add (Int.shl hi (Int.repr 12)) lo) + /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r. +Proof. + unfold load_hilo32; intros. + predSpec Int.eq Int.eq_spec lo Int.zero. +- subst lo. econstructor; split. + apply exec_straight_one. simpl; eauto. auto. + split. rewrite Int.add_zero. Simpl. + intros; Simpl. +- econstructor; split. + eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto. + split. Simpl. + intros; Simpl. +Qed. + +Lemma loadimm32_correct: + forall rd n k rs m, + exists rs', + exec_straight ge fn (loadimm32 rd n k) rs m k rs' m + /\ rs'#rd = Vint n + /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r. +Proof. + unfold loadimm32; intros. generalize (make_immed32_sound n); intros E. + destruct (make_immed32 n). +- subst imm. econstructor; split. + apply exec_straight_one. simpl; eauto. auto. + split. rewrite Int.add_zero_l; Simpl. + intros; Simpl. +- rewrite E. apply load_hilo32_correct. +Qed. + +Lemma opimm32_correct: + forall (op: ireg -> ireg0 -> ireg0 -> instruction) + (opi: ireg -> ireg0 -> int -> instruction) + (sem: val -> val -> val) m, + (forall d s1 s2 rs, + exec_instr ge fn (op d s1 s2) rs m = Next (nextinstr (rs#d <- (sem rs##s1 rs##s2))) m) -> + (forall d s n rs, + exec_instr ge fn (opi d s n) rs m = Next (nextinstr (rs#d <- (sem rs##s (Vint n)))) m) -> + forall rd r1 n k rs, + r1 <> X31 -> + exists rs', + exec_straight ge fn (opimm32 op opi rd r1 n k) rs m k rs' m + /\ rs'#rd = sem rs##r1 (Vint n) + /\ forall r, r <> PC -> r <> rd -> r <> X31 -> rs'#r = rs#r. +Proof. + intros. unfold opimm32. generalize (make_immed32_sound n); intros E. + destruct (make_immed32 n). +- subst imm. econstructor; split. + apply exec_straight_one. rewrite H0. simpl; eauto. auto. + split. Simpl. intros; Simpl. +- destruct (load_hilo32_correct X31 hi lo (op rd r1 X31 :: k) rs m) + as (rs' & A & B & C). + econstructor; split. + eapply exec_straight_trans. eexact A. apply exec_straight_one. + rewrite H; eauto. auto. + split. Simpl. simpl. rewrite B, C, E. auto. congruence. congruence. + intros; Simpl. +Qed. + +(** 64-bit integer constants and arithmetic *) + +Lemma load_hilo64_correct: + forall rd hi lo k rs m, + exists rs', + exec_straight ge fn (load_hilo64 rd hi lo k) rs m k rs' m + /\ rs'#rd = Vlong (Int64.add (Int64.sign_ext 32 (Int64.shl hi (Int64.repr 12))) lo) + /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r. +Proof. + unfold load_hilo64; intros. + predSpec Int64.eq Int64.eq_spec lo Int64.zero. +- subst lo. econstructor; split. + apply exec_straight_one. simpl; eauto. auto. + split. rewrite Int64.add_zero. Simpl. + intros; Simpl. +- econstructor; split. + eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto. + split. Simpl. + intros; Simpl. +Qed. + +Lemma loadimm64_correct: + forall rd n k rs m, + exists rs', + exec_straight ge fn (loadimm64 rd n k) rs m k rs' m + /\ rs'#rd = Vlong n + /\ forall r, r <> PC -> r <> rd -> r <> X31 -> rs'#r = rs#r. +Proof. + unfold loadimm64; intros. generalize (make_immed64_sound n); intros E. + destruct (make_immed64 n). +- subst imm. econstructor; split. + apply exec_straight_one. simpl; eauto. auto. + split. rewrite Int64.add_zero_l; Simpl. + intros; Simpl. +- exploit load_hilo64_correct; eauto. intros (rs' & A & B & C). + rewrite E. exists rs'; eauto. +- subst imm. econstructor; split. + apply exec_straight_one. simpl; eauto. auto. + split. Simpl. + intros; Simpl. +Qed. + +Lemma opimm64_correct: + forall (op: ireg -> ireg0 -> ireg0 -> instruction) + (opi: ireg -> ireg0 -> int64 -> instruction) + (sem: val -> val -> val) m, + (forall d s1 s2 rs, + exec_instr ge fn (op d s1 s2) rs m = Next (nextinstr (rs#d <- (sem rs###s1 rs###s2))) m) -> + (forall d s n rs, + exec_instr ge fn (opi d s n) rs m = Next (nextinstr (rs#d <- (sem rs###s (Vlong n)))) m) -> + forall rd r1 n k rs, + r1 <> X31 -> + exists rs', + exec_straight ge fn (opimm64 op opi rd r1 n k) rs m k rs' m + /\ rs'#rd = sem rs##r1 (Vlong n) + /\ forall r, r <> PC -> r <> rd -> r <> X31 -> rs'#r = rs#r. +Proof. + intros. unfold opimm64. generalize (make_immed64_sound n); intros E. + destruct (make_immed64 n). +- subst imm. econstructor; split. + apply exec_straight_one. rewrite H0. simpl; eauto. auto. + split. Simpl. intros; Simpl. +- destruct (load_hilo64_correct X31 hi lo (op rd r1 X31 :: k) rs m) + as (rs' & A & B & C). + econstructor; split. + eapply exec_straight_trans. eexact A. apply exec_straight_one. + rewrite H; eauto. auto. + split. Simpl. simpl. rewrite B, C, E. auto. congruence. congruence. + intros; Simpl. +- subst imm. econstructor; split. + eapply exec_straight_two. simpl; eauto. rewrite H. simpl; eauto. auto. auto. + split. Simpl. intros; Simpl. +Qed. + +(** Add offset to pointer *) + +Lemma addptrofs_correct: + forall rd r1 n k rs m, + r1 <> X31 -> + exists rs', + exec_straight ge fn (addptrofs rd r1 n k) rs m k rs' m + /\ Val.lessdef (Val.offset_ptr rs#r1 n) rs'#rd + /\ forall r, r <> PC -> r <> rd -> r <> X31 -> rs'#r = rs#r. +Proof. + unfold addptrofs; intros. + destruct (Ptrofs.eq_dec n Ptrofs.zero). +- subst n. econstructor; split. + apply exec_straight_one. simpl; eauto. auto. + split. Simpl. destruct (rs r1); simpl; auto. rewrite Ptrofs.add_zero; auto. + intros; Simpl. +- destruct Archi.ptr64 eqn:SF. ++ unfold addimm64. + exploit (opimm64_correct Paddl Paddil Val.addl); eauto. intros (rs' & A & B & C). + exists rs'; split. eexact A. split; auto. + rewrite B. simpl. destruct (rs r1); simpl; auto. rewrite SF. + rewrite Ptrofs.of_int64_to_int64 by auto. auto. ++ unfold addimm32. + exploit (opimm32_correct Paddw Paddiw Val.add); eauto. intros (rs' & A & B & C). + exists rs'; split. eexact A. split; auto. + rewrite B. simpl. destruct (rs r1); simpl; auto. rewrite SF. + rewrite Ptrofs.of_int_to_int by auto. auto. +Qed. + +Lemma addptrofs_correct_2: + forall rd r1 n k (rs: regset) m b ofs, + r1 <> X31 -> rs#r1 = Vptr b ofs -> + exists rs', + exec_straight ge fn (addptrofs rd r1 n k) rs m k rs' m + /\ rs'#rd = Vptr b (Ptrofs.add ofs n) + /\ forall r, r <> PC -> r <> rd -> r <> X31 -> rs'#r = rs#r. +Proof. + intros. exploit (addptrofs_correct rd r1 n); eauto. intros (rs' & A & B & C). + exists rs'; intuition eauto. + rewrite H0 in B. inv B. auto. +Qed. + +Ltac ArgsInv := + repeat (match goal with + | [ H: Error _ = OK _ |- _ ] => discriminate + | [ H: match ?args with nil => _ | _ :: _ => _ end = OK _ |- _ ] => destruct args + | [ H: bind _ _ = OK _ |- _ ] => monadInv H + | [ H: match _ with left _ => _ | right _ => assertion_failed end = OK _ |- _ ] => monadInv H; ArgsInv + | [ H: match _ with true => _ | false => assertion_failed end = OK _ |- _ ] => monadInv H; ArgsInv + end); + subst; + repeat (match goal with + | [ H: ireg_of _ = OK _ |- _ ] => simpl in *; rewrite (ireg_of_eq _ _ H) in * + | [ H: freg_of _ = OK _ |- _ ] => simpl in *; rewrite (freg_of_eq _ _ H) in * + end). + +Lemma transl_cbranch_correct_1: + forall cond args lbl k c m ms b sp rs m', + transl_cbranch cond args lbl k = OK c -> + eval_condition cond (List.map ms args) m = Some b -> + agree ms sp rs -> + Mem.extends m m' -> + exists rs', exists insn, + exec_straight_opt ge fn c rs m' (insn :: k) rs' m' + /\ exec_instr ge fn insn rs' m' = eval_branch fn lbl rs' m' (Some b) + /\ forall r, r <> PC -> r <> X31 -> rs'#r = rs#r. +Proof. + intros until m'; intros TRANSL EVAL AG MEXT. + set (vl' := map rs (map preg_of args)). + assert (EVAL': eval_condition cond vl' m' = Some b). + { apply eval_condition_lessdef with (map ms args) m; auto. eapply preg_vals; eauto. } + clear EVAL MEXT AG. + destruct cond; simpl in TRANSL; ArgsInv. + (* Pbeqw / Cmp *) + { destruct optR0 as [[]|]; + unfold apply_bin_r0, apply_bin_r0_r0r0lbl in *; + unfold zero32, Op.zero32 in *; + eexists; eexists; eauto; split; constructor; auto; + simpl in *. + + destruct (rs x); simpl in *; try congruence. + assert (HB: (Int.eq Int.zero i) = b) by congruence. + rewrite HB; destruct b; simpl; auto. + + destruct (rs x); simpl in *; try congruence. + assert (HB: (Int.eq i Int.zero) = b) by congruence. + rewrite HB; destruct b; simpl; auto. + + destruct (rs x); simpl in *; try congruence. + destruct (rs x0); try congruence. + assert (HB: (Int.eq i i0) = b) by congruence. + rewrite HB; destruct b; simpl; auto. } + (* Pbnew / Cmp *) + { destruct optR0 as [[]|]; + unfold apply_bin_r0, apply_bin_r0_r0r0lbl in *; + unfold zero32, Op.zero32 in *; + eexists; eexists; eauto; split; constructor; auto; + simpl in *. + + destruct (rs x); simpl in *; try congruence. + assert (HB: negb (Int.eq Int.zero i) = b) by congruence. + rewrite HB; destruct b; simpl; auto. + + destruct (rs x); simpl in *; try congruence. + assert (HB: negb (Int.eq i Int.zero) = b) by congruence. + rewrite HB; destruct b; simpl; auto. + + destruct (rs x); simpl in *; try congruence. + destruct (rs x0); try congruence. + assert (HB: negb (Int.eq i i0) = b) by congruence. + rewrite HB; destruct b; simpl; auto. } + (* Pbeqw, Pbnew, Pbltw, Pbtluw, Pbgew, Pbgeuw / Cmpu *) + 1-6: + destruct optR0 as [[]|]; + unfold apply_bin_r0, apply_bin_r0_r0r0lbl in *; + unfold zero32, Op.zero32 in *; + eexists; eexists; eauto; split; constructor; + simpl in *; try rewrite EVAL'; auto. + (* Pbeql / Cmpl *) + { destruct optR0 as [[]|]; + unfold apply_bin_r0, apply_bin_r0_r0r0lbl in *; + unfold zero64, Op.zero64 in *; + eexists; eexists; eauto; split; constructor; + simpl in *; auto. + + destruct (rs x); simpl in *; try congruence. + assert (HB: (Int64.eq Int64.zero i) = b) by congruence. + rewrite HB; destruct b; simpl; auto. + + destruct (rs x); simpl in *; try congruence. + assert (HB: (Int64.eq i Int64.zero) = b) by congruence. + rewrite HB; destruct b; simpl; auto. + + destruct (rs x); simpl in *; try congruence. + destruct (rs x0); try congruence. + assert (HB: (Int64.eq i i0) = b) by congruence. + rewrite HB; destruct b; simpl; auto. } + (* Pbnel / Cmpl *) + { destruct optR0 as [[]|]; + unfold apply_bin_r0, apply_bin_r0_r0r0lbl in *; + unfold zero64, Op.zero64 in *; + eexists; eexists; eauto; split; constructor; + simpl in *; auto. + + destruct (rs x); simpl in *; try congruence. + assert (HB: negb (Int64.eq Int64.zero i) = b) by congruence. + rewrite HB; destruct b; simpl; auto. + + destruct (rs x); simpl in *; try congruence. + assert (HB: negb (Int64.eq i Int64.zero) = b) by congruence. + rewrite HB; destruct b; simpl; auto. + + destruct (rs x); simpl in *; try congruence. + destruct (rs x0); try congruence. + assert (HB: negb (Int64.eq i i0) = b) by congruence. + rewrite HB; destruct b; simpl; auto. } + (* Pbeql, Pbnel, Pbltl, Pbtlul, Pbgel, Pbgeul / Cmplu *) + 1-6: + destruct optR0 as [[]|]; + unfold apply_bin_r0, apply_bin_r0_r0r0lbl in *; + unfold zero64, Op.zero64 in *; + eexists; eexists; eauto; split; constructor; + simpl in *; try rewrite EVAL'; auto. +Qed. + +Lemma transl_cbranch_correct_true: + forall cond args lbl k c m ms sp rs m', + transl_cbranch cond args lbl k = OK c -> + eval_condition cond (List.map ms args) m = Some true -> + agree ms sp rs -> + Mem.extends m m' -> + exists rs', exists insn, + exec_straight_opt ge fn c rs m' (insn :: k) rs' m' + /\ exec_instr ge fn insn rs' m' = goto_label fn lbl rs' m' + /\ forall r, r <> PC -> r <> X31 -> rs'#r = rs#r. +Proof. + intros. eapply transl_cbranch_correct_1 with (b := true); eauto. +Qed. + +Lemma transl_cbranch_correct_false: + forall cond args lbl k c m ms sp rs m', + transl_cbranch cond args lbl k = OK c -> + eval_condition cond (List.map ms args) m = Some false -> + agree ms sp rs -> + Mem.extends m m' -> + exists rs', + exec_straight ge fn c rs m' k rs' m' + /\ forall r, r <> PC -> r <> X31 -> rs'#r = rs#r. +Proof. + intros. exploit transl_cbranch_correct_1; eauto. simpl. + intros (rs' & insn & A & B & C). + exists (nextinstr rs'). + split. eapply exec_straight_opt_right; eauto. apply exec_straight_one; auto. + intros; Simpl. +Qed. + +(** Some arithmetic properties. *) + +Remark cast32unsigned_from_cast32signed: + forall i, Int64.repr (Int.unsigned i) = Int64.zero_ext 32 (Int64.repr (Int.signed i)). +Proof. + intros. apply Int64.same_bits_eq; intros. + rewrite Int64.bits_zero_ext, !Int64.testbit_repr by tauto. + rewrite Int.bits_signed by tauto. fold (Int.testbit i i0). + change Int.zwordsize with 32. + destruct (zlt i0 32). auto. apply Int.bits_above. auto. +Qed. + +(* Translation of arithmetic operations *) + +Ltac SimplEval H := + match type of H with + | Some _ = None _ => discriminate + | Some _ = Some _ => inv H + | ?a = Some ?b => let A := fresh in assert (A: Val.maketotal a = b) by (rewrite H; reflexivity) +end. + +Ltac TranslOpSimpl := + econstructor; split; + [ apply exec_straight_one; [simpl; eauto | reflexivity] + | split; [ apply Val.lessdef_same; Simpl; fail | intros; Simpl; fail ] ]. + +Lemma transl_op_correct: + forall op args res k (rs: regset) m v c, + transl_op op args res k = OK c -> + eval_operation ge (rs#SP) op (map rs (map preg_of args)) m = Some v -> + exists rs', + exec_straight ge fn c rs m k rs' m + /\ Val.lessdef v rs'#(preg_of res) + /\ forall r, data_preg r = true -> r <> preg_of res -> preg_notin r (destroyed_by_op op) -> rs' r = rs r. +Proof. + assert (SAME: forall v1 v2, v1 = v2 -> Val.lessdef v2 v1). { intros; subst; auto. } +Opaque Int.eq. + intros until c; intros TR EV. + unfold transl_op in TR; destruct op; ArgsInv; simpl in EV; SimplEval EV; try TranslOpSimpl. + (* move *) + { destruct (preg_of res), (preg_of m0); inv TR; TranslOpSimpl. } + (* intconst *) + { exploit loadimm32_correct; eauto. intros (rs' & A & B & C). + exists rs'; split; eauto. rewrite B; auto with asmgen. } + (* longconst *) + { exploit loadimm64_correct; eauto. intros (rs' & A & B & C). + exists rs'; split; eauto. rewrite B; auto with asmgen. } + (* floatconst *) + { destruct (Float.eq_dec n Float.zero). + + subst n. econstructor; split. + apply exec_straight_one. simpl; eauto. auto. + split; intros; Simpl. + + econstructor; split. + apply exec_straight_one. simpl; eauto. auto. + split; intros; Simpl. } + (* singleconst *) + { destruct (Float32.eq_dec n Float32.zero). + + subst n. econstructor; split. + apply exec_straight_one. simpl; eauto. auto. + split; intros; Simpl. + + econstructor; split. + apply exec_straight_one. simpl; eauto. auto. + split; intros; Simpl. } + (* addrsymbol *) + { destruct (Archi.pic_code tt && negb (Ptrofs.eq ofs Ptrofs.zero)). + + set (rs1 := nextinstr (rs#x <- (Genv.symbol_address ge id Ptrofs.zero))). + exploit (addptrofs_correct x x ofs k rs1 m); eauto with asmgen. + intros (rs2 & A & B & C). + exists rs2; split. + apply exec_straight_step with rs1 m; auto. + split. replace ofs with (Ptrofs.add Ptrofs.zero ofs) by (apply Ptrofs.add_zero_l). + rewrite Genv.shift_symbol_address. + replace (rs1 x) with (Genv.symbol_address ge id Ptrofs.zero) in B by (unfold rs1; Simpl). + exact B. + intros. rewrite C by eauto with asmgen. unfold rs1; Simpl. + + TranslOpSimpl. } + (* stackoffset *) + { exploit addptrofs_correct. instantiate (1 := X2); auto with asmgen. intros (rs' & A & B & C). + exists rs'; split; eauto. auto with asmgen. } + (* cast8signed *) + { econstructor; split. + eapply exec_straight_two. simpl;eauto. simpl;eauto. auto. auto. + split; intros; Simpl. + assert (A: Int.ltu (Int.repr 24) Int.iwordsize = true) by auto. + destruct (rs x0); auto; simpl. rewrite A; simpl. rewrite A. + apply Val.lessdef_same. f_equal. apply Int.sign_ext_shr_shl. split; reflexivity. } + (* cast16signed *) + { econstructor; split. + eapply exec_straight_two. simpl;eauto. simpl;eauto. auto. auto. + split; intros; Simpl. + assert (A: Int.ltu (Int.repr 16) Int.iwordsize = true) by auto. + destruct (rs x0); auto; simpl. rewrite A; simpl. rewrite A. + apply Val.lessdef_same. f_equal. apply Int.sign_ext_shr_shl. split; reflexivity. } + (* addimm *) + { exploit (opimm32_correct Paddw Paddiw Val.add); auto. instantiate (1 := x0); eauto with asmgen. + intros (rs' & A & B & C). + exists rs'; split; eauto. rewrite B; auto with asmgen. } + (* andimm *) + { exploit (opimm32_correct Pandw Pandiw Val.and); auto. instantiate (1 := x0); eauto with asmgen. + intros (rs' & A & B & C). + exists rs'; split; eauto. rewrite B; auto with asmgen. } + (* orimm *) + exploit (opimm32_correct Porw Poriw Val.or); auto. instantiate (1 := x0); eauto with asmgen. + { intros (rs' & A & B & C). + exists rs'; split; eauto. rewrite B; auto with asmgen. } + (* xorimm *) + { exploit (opimm32_correct Pxorw Pxoriw Val.xor); auto. instantiate (1 := x0); eauto with asmgen. + intros (rs' & A & B & C). + exists rs'; split; eauto. rewrite B; auto with asmgen. } + (* shrximm *) + { destruct (Val.shrx (rs x0) (Vint n)) eqn:TOTAL; cbn. + { + exploit Val.shrx_shr_3; eauto. intros E; subst v. + destruct (Int.eq n Int.zero). + + econstructor; split. apply exec_straight_one. simpl; eauto. auto. + split; intros; Simpl. + + destruct (Int.eq n Int.one). + * econstructor; split. + eapply exec_straight_step. simpl; reflexivity. auto. + eapply exec_straight_step. simpl; reflexivity. auto. + apply exec_straight_one. simpl; reflexivity. auto. + split; intros; Simpl. + * change (Int.repr 32) with Int.iwordsize. set (n' := Int.sub Int.iwordsize n). + econstructor; split. + eapply exec_straight_step. simpl; reflexivity. auto. + eapply exec_straight_step. simpl; reflexivity. auto. + eapply exec_straight_step. simpl; reflexivity. auto. + apply exec_straight_one. simpl; reflexivity. auto. + split; intros; Simpl. + } + destruct (Int.eq n Int.zero). + + econstructor; split. apply exec_straight_one. simpl; eauto. auto. + split; intros; Simpl. + + destruct (Int.eq n Int.one). + * econstructor; split. + eapply exec_straight_step. simpl; reflexivity. auto. + eapply exec_straight_step. simpl; reflexivity. auto. + apply exec_straight_one. simpl; reflexivity. auto. + split; intros; Simpl. + * change (Int.repr 32) with Int.iwordsize. set (n' := Int.sub Int.iwordsize n). + econstructor; split. + eapply exec_straight_step. simpl; reflexivity. auto. + eapply exec_straight_step. simpl; reflexivity. auto. + eapply exec_straight_step. simpl; reflexivity. auto. + apply exec_straight_one. simpl; reflexivity. auto. + split; intros; Simpl. } + (* longofintu *) + { econstructor; split. + eapply exec_straight_three. simpl; eauto. simpl; eauto. simpl; eauto. auto. auto. auto. + split; intros; Simpl. destruct (rs x0); auto. simpl. + assert (A: Int.ltu (Int.repr 32) Int64.iwordsize' = true) by auto. + rewrite A; simpl. rewrite A. apply Val.lessdef_same. f_equal. + rewrite cast32unsigned_from_cast32signed. apply Int64.zero_ext_shru_shl. compute; auto. } + (* addlimm *) + { exploit (opimm64_correct Paddl Paddil Val.addl); auto. instantiate (1 := x0); eauto with asmgen. + intros (rs' & A & B & C). + exists rs'; split; eauto. rewrite B; auto with asmgen. } + (* andimm *) + { exploit (opimm64_correct Pandl Pandil Val.andl); auto. instantiate (1 := x0); eauto with asmgen. + intros (rs' & A & B & C). + exists rs'; split; eauto. rewrite B; auto with asmgen. } + (* orimm *) + { exploit (opimm64_correct Porl Poril Val.orl); auto. instantiate (1 := x0); eauto with asmgen. + intros (rs' & A & B & C). + exists rs'; split; eauto. rewrite B; auto with asmgen. } + (* xorimm *) + { exploit (opimm64_correct Pxorl Pxoril Val.xorl); auto. instantiate (1 := x0); eauto with asmgen. + intros (rs' & A & B & C). + exists rs'; split; eauto. rewrite B; auto with asmgen. } + (* shrxlimm *) + { destruct (Val.shrxl (rs x0) (Vint n)) eqn:TOTAL. + { + exploit Val.shrxl_shrl_3; eauto. intros E; subst v. + destruct (Int.eq n Int.zero). + + econstructor; split. apply exec_straight_one. simpl; eauto. auto. + split; intros; Simpl. + + destruct (Int.eq n Int.one). + * econstructor; split. + eapply exec_straight_step. simpl; reflexivity. auto. + eapply exec_straight_step. simpl; reflexivity. auto. + apply exec_straight_one. simpl; reflexivity. auto. + split; intros; Simpl. + + * change (Int.repr 64) with Int64.iwordsize'. set (n' := Int.sub Int64.iwordsize' n). + econstructor; split. + eapply exec_straight_step. simpl; reflexivity. auto. + eapply exec_straight_step. simpl; reflexivity. auto. + eapply exec_straight_step. simpl; reflexivity. auto. + apply exec_straight_one. simpl; reflexivity. auto. + split; intros; Simpl. + } + destruct (Int.eq n Int.zero). + + econstructor; split. apply exec_straight_one. simpl; eauto. auto. + split; intros; Simpl. + + destruct (Int.eq n Int.one). + * econstructor; split. + eapply exec_straight_step. simpl; reflexivity. auto. + eapply exec_straight_step. simpl; reflexivity. auto. + apply exec_straight_one. simpl; reflexivity. auto. + split; intros; Simpl. + + * change (Int.repr 64) with Int64.iwordsize'. set (n' := Int.sub Int64.iwordsize' n). + econstructor; split. + eapply exec_straight_step. simpl; reflexivity. auto. + eapply exec_straight_step. simpl; reflexivity. auto. + eapply exec_straight_step. simpl; reflexivity. auto. + apply exec_straight_one. simpl; reflexivity. auto. + split; intros; Simpl. } + (* Expanded instructions from RTL *) + 7,8,15,16: + econstructor; split; try apply exec_straight_one; simpl; eauto; + split; intros; Simpl; unfold may_undef_int; try destruct is_long; simpl; + try rewrite Int.add_commut; try rewrite Int64.add_commut; + destruct (rs (preg_of m0)); try discriminate; eauto. + 1-12: + destruct optR0 as [[]|]; unfold apply_bin_r0_r0r0, apply_bin_r0; + econstructor; split; try apply exec_straight_one; simpl; eauto; + split; intros; Simpl; + destruct (rs x0); auto; + destruct (rs x1); auto. + (* select *) + { econstructor; split. apply exec_straight_one. simpl; eauto. auto. + split; intros; Simpl. + apply Val.lessdef_normalize. } +Qed. + +(** Memory accesses *) + +Lemma indexed_memory_access_correct: + forall mk_instr base ofs k rs m, + base <> X31 -> + exists base' ofs' rs', + exec_straight_opt ge fn (indexed_memory_access mk_instr base ofs k) rs m + (mk_instr base' ofs' :: k) rs' m + /\ Val.offset_ptr rs'#base' (eval_offset ge ofs') = Val.offset_ptr rs#base ofs + /\ forall r, r <> PC -> r <> X31 -> rs'#r = rs#r. +Proof. + unfold indexed_memory_access; intros. + destruct Archi.ptr64 eqn:SF. +- generalize (make_immed64_sound (Ptrofs.to_int64 ofs)); intros EQ. + destruct (make_immed64 (Ptrofs.to_int64 ofs)). ++ econstructor; econstructor; econstructor; split. + apply exec_straight_opt_refl. + split; auto. simpl. subst imm. rewrite Ptrofs.of_int64_to_int64 by auto. auto. ++ econstructor; econstructor; econstructor; split. + constructor. eapply exec_straight_two. + simpl; eauto. simpl; eauto. auto. auto. + split; intros; Simpl. destruct (rs base); auto; simpl. rewrite SF. simpl. + rewrite Ptrofs.add_assoc. f_equal. f_equal. + rewrite <- (Ptrofs.of_int64_to_int64 SF ofs). rewrite EQ. + symmetry; auto with ptrofs. ++ econstructor; econstructor; econstructor; split. + constructor. eapply exec_straight_two. + simpl; eauto. simpl; eauto. auto. auto. + split; intros; Simpl. unfold eval_offset. destruct (rs base); auto; simpl. rewrite SF. simpl. + rewrite Ptrofs.add_zero. subst imm. rewrite Ptrofs.of_int64_to_int64 by auto. auto. +- generalize (make_immed32_sound (Ptrofs.to_int ofs)); intros EQ. + destruct (make_immed32 (Ptrofs.to_int ofs)). ++ econstructor; econstructor; econstructor; split. + apply exec_straight_opt_refl. + split; auto. simpl. subst imm. rewrite Ptrofs.of_int_to_int by auto. auto. ++ econstructor; econstructor; econstructor; split. + constructor. eapply exec_straight_two. + simpl; eauto. simpl; eauto. auto. auto. + split; intros; Simpl. destruct (rs base); auto; simpl. rewrite SF. simpl. + rewrite Ptrofs.add_assoc. f_equal. f_equal. + rewrite <- (Ptrofs.of_int_to_int SF ofs). rewrite EQ. + symmetry; auto with ptrofs. +Qed. + +Lemma indexed_load_access_correct: + forall chunk (mk_instr: ireg -> offset -> instruction) rd m, + (forall base ofs rs, + exec_instr ge fn (mk_instr base ofs) rs m = exec_load ge chunk rs m rd base ofs) -> + forall (base: ireg) ofs k (rs: regset) v, + Mem.loadv chunk m (Val.offset_ptr rs#base ofs) = Some v -> + base <> X31 -> rd <> PC -> + exists rs', + exec_straight ge fn (indexed_memory_access mk_instr base ofs k) rs m k rs' m + /\ rs'#rd = v + /\ forall r, r <> PC -> r <> X31 -> r <> rd -> rs'#r = rs#r. +Proof. + intros until m; intros EXEC; intros until v; intros LOAD NOT31 NOTPC. + exploit indexed_memory_access_correct; eauto. + intros (base' & ofs' & rs' & A & B & C). + econstructor; split. + eapply exec_straight_opt_right. eexact A. apply exec_straight_one. rewrite EXEC. + unfold exec_load. rewrite B, LOAD. eauto. Simpl. + split; intros; Simpl. +Qed. + +Lemma indexed_store_access_correct: + forall chunk (mk_instr: ireg -> offset -> instruction) r1 m, + (forall base ofs rs, + exec_instr ge fn (mk_instr base ofs) rs m = exec_store ge chunk rs m r1 base ofs) -> + forall (base: ireg) ofs k (rs: regset) m', + Mem.storev chunk m (Val.offset_ptr rs#base ofs) (rs#r1) = Some m' -> + base <> X31 -> r1 <> X31 -> r1 <> PC -> + exists rs', + exec_straight ge fn (indexed_memory_access mk_instr base ofs k) rs m k rs' m' + /\ forall r, r <> PC -> r <> X31 -> rs'#r = rs#r. +Proof. + intros until m; intros EXEC; intros until m'; intros STORE NOT31 NOT31' NOTPC. + exploit indexed_memory_access_correct; eauto. + intros (base' & ofs' & rs' & A & B & C). + econstructor; split. + eapply exec_straight_opt_right. eexact A. apply exec_straight_one. rewrite EXEC. + unfold exec_store. rewrite B, C, STORE by auto. eauto. auto. + intros; Simpl. +Qed. + +Lemma loadind_correct: + forall (base: ireg) ofs ty dst k c (rs: regset) m v, + loadind base ofs ty dst k = OK c -> + Mem.loadv (chunk_of_type ty) m (Val.offset_ptr rs#base ofs) = Some v -> + base <> X31 -> + exists rs', + exec_straight ge fn c rs m k rs' m + /\ rs'#(preg_of dst) = v + /\ forall r, r <> PC -> r <> X31 -> r <> preg_of dst -> rs'#r = rs#r. +Proof. + intros until v; intros TR LOAD NOT31. + assert (A: exists mk_instr, + c = indexed_memory_access mk_instr base ofs k + /\ forall base' ofs' rs', + exec_instr ge fn (mk_instr base' ofs') rs' m = + exec_load ge (chunk_of_type ty) rs' m (preg_of dst) base' ofs'). + { unfold loadind in TR. destruct ty, (preg_of dst); inv TR; econstructor; split; eauto. } + destruct A as (mk_instr & B & C). subst c. + eapply indexed_load_access_correct; eauto with asmgen. +Qed. + +Lemma storeind_correct: + forall (base: ireg) ofs ty src k c (rs: regset) m m', + storeind src base ofs ty k = OK c -> + Mem.storev (chunk_of_type ty) m (Val.offset_ptr rs#base ofs) rs#(preg_of src) = Some m' -> + base <> X31 -> + exists rs', + exec_straight ge fn c rs m k rs' m' + /\ forall r, r <> PC -> r <> X31 -> rs'#r = rs#r. +Proof. + intros until m'; intros TR STORE NOT31. + assert (A: exists mk_instr, + c = indexed_memory_access mk_instr base ofs k + /\ forall base' ofs' rs', + exec_instr ge fn (mk_instr base' ofs') rs' m = + exec_store ge (chunk_of_type ty) rs' m (preg_of src) base' ofs'). + { unfold storeind in TR. destruct ty, (preg_of src); inv TR; econstructor; split; eauto. } + destruct A as (mk_instr & B & C). subst c. + eapply indexed_store_access_correct; eauto with asmgen. +Qed. + +Lemma loadind_ptr_correct: + forall (base: ireg) ofs (dst: ireg) k (rs: regset) m v, + Mem.loadv Mptr m (Val.offset_ptr rs#base ofs) = Some v -> + base <> X31 -> + exists rs', + exec_straight ge fn (loadind_ptr base ofs dst k) rs m k rs' m + /\ rs'#dst = v + /\ forall r, r <> PC -> r <> X31 -> r <> dst -> rs'#r = rs#r. +Proof. + intros. eapply indexed_load_access_correct; eauto with asmgen. + intros. unfold Mptr. destruct Archi.ptr64; auto. +Qed. + +Lemma storeind_ptr_correct: + forall (base: ireg) ofs (src: ireg) k (rs: regset) m m', + Mem.storev Mptr m (Val.offset_ptr rs#base ofs) rs#src = Some m' -> + base <> X31 -> src <> X31 -> + exists rs', + exec_straight ge fn (storeind_ptr src base ofs k) rs m k rs' m' + /\ forall r, r <> PC -> r <> X31 -> rs'#r = rs#r. +Proof. + intros. eapply indexed_store_access_correct with (r1 := src); eauto with asmgen. + intros. unfold Mptr. destruct Archi.ptr64; auto. +Qed. + +Lemma transl_memory_access_correct: + forall mk_instr addr args k c (rs: regset) m v, + transl_memory_access mk_instr addr args k = OK c -> + eval_addressing ge rs#SP addr (map rs (map preg_of args)) = Some v -> + exists base ofs rs', + exec_straight_opt ge fn c rs m (mk_instr base ofs :: k) rs' m + /\ Val.offset_ptr rs'#base (eval_offset ge ofs) = v + /\ forall r, r <> PC -> r <> X31 -> rs'#r = rs#r. +Proof. + intros until v; intros TR EV. + unfold transl_memory_access in TR; destruct addr; ArgsInv. +- (* indexed *) + inv EV. apply indexed_memory_access_correct; eauto with asmgen. +- (* global *) + simpl in EV. inv EV. inv TR. econstructor; econstructor; econstructor; split. + constructor. apply exec_straight_one. simpl; eauto. auto. + split; intros; Simpl. unfold eval_offset. apply low_high_half. +- (* stack *) + inv TR. inv EV. apply indexed_memory_access_correct; eauto with asmgen. +Qed. + +Lemma transl_load_access_correct: + forall chunk (mk_instr: ireg -> offset -> instruction) addr args k c rd (rs: regset) m v v', + (forall base ofs rs, + exec_instr ge fn (mk_instr base ofs) rs m = exec_load ge chunk rs m rd base ofs) -> + transl_memory_access mk_instr addr args k = OK c -> + eval_addressing ge rs#SP addr (map rs (map preg_of args)) = Some v -> + Mem.loadv chunk m v = Some v' -> + rd <> PC -> + exists rs', + exec_straight ge fn c rs m k rs' m + /\ rs'#rd = v' + /\ forall r, r <> PC -> r <> X31 -> r <> rd -> rs'#r = rs#r. +Proof. + intros until v'; intros INSTR TR EV LOAD NOTPC. + exploit transl_memory_access_correct; eauto. + intros (base & ofs & rs' & A & B & C). + econstructor; split. + eapply exec_straight_opt_right. eexact A. apply exec_straight_one. + rewrite INSTR. unfold exec_load. rewrite B, LOAD. reflexivity. Simpl. + split; intros; Simpl. +Qed. + +Lemma transl_store_access_correct: + forall chunk (mk_instr: ireg -> offset -> instruction) addr args k c r1 (rs: regset) m v m', + (forall base ofs rs, + exec_instr ge fn (mk_instr base ofs) rs m = exec_store ge chunk rs m r1 base ofs) -> + transl_memory_access mk_instr addr args k = OK c -> + eval_addressing ge rs#SP addr (map rs (map preg_of args)) = Some v -> + Mem.storev chunk m v rs#r1 = Some m' -> + r1 <> PC -> r1 <> X31 -> + exists rs', + exec_straight ge fn c rs m k rs' m' + /\ forall r, r <> PC -> r <> X31 -> rs'#r = rs#r. +Proof. + intros until m'; intros INSTR TR EV STORE NOTPC NOT31. + exploit transl_memory_access_correct; eauto. + intros (base & ofs & rs' & A & B & C). + econstructor; split. + eapply exec_straight_opt_right. eexact A. apply exec_straight_one. + rewrite INSTR. unfold exec_store. rewrite B, C, STORE by auto. reflexivity. auto. + intros; Simpl. +Qed. + +Lemma transl_load_correct: + forall trap chunk addr args dst k c (rs: regset) m a v, + transl_load trap chunk addr args dst k = OK c -> + eval_addressing ge rs#SP addr (map rs (map preg_of args)) = Some a -> + Mem.loadv chunk m a = Some v -> + exists rs', + exec_straight ge fn c rs m k rs' m + /\ rs'#(preg_of dst) = v + /\ forall r, r <> PC -> r <> X31 -> r <> preg_of dst -> rs'#r = rs#r. +Proof. + intros until v; intros TR EV LOAD. + destruct trap; try (simpl in *; discriminate). + assert (A: exists mk_instr, + transl_memory_access mk_instr addr args k = OK c + /\ forall base ofs rs, + exec_instr ge fn (mk_instr base ofs) rs m = exec_load ge chunk rs m (preg_of dst) base ofs). + { unfold transl_load in TR; destruct chunk; ArgsInv; econstructor; (split; [eassumption|auto]). } + destruct A as (mk_instr & B & C). + eapply transl_load_access_correct; eauto with asmgen. +Qed. + +Lemma transl_store_correct: + forall chunk addr args src k c (rs: regset) m a m', + transl_store chunk addr args src k = OK c -> + eval_addressing ge rs#SP addr (map rs (map preg_of args)) = Some a -> + Mem.storev chunk m a rs#(preg_of src) = Some m' -> + exists rs', + exec_straight ge fn c rs m k rs' m' + /\ forall r, r <> PC -> r <> X31 -> rs'#r = rs#r. +Proof. + intros until m'; intros TR EV STORE. + assert (A: exists mk_instr chunk', + transl_memory_access mk_instr addr args k = OK c + /\ (forall base ofs rs, + exec_instr ge fn (mk_instr base ofs) rs m = exec_store ge chunk' rs m (preg_of src) base ofs) + /\ Mem.storev chunk m a rs#(preg_of src) = Mem.storev chunk' m a rs#(preg_of src)). + { unfold transl_store in TR; destruct chunk; ArgsInv; + (econstructor; econstructor; split; [eassumption | split; [ intros; simpl; reflexivity | auto]]). + destruct a; auto. apply Mem.store_signed_unsigned_8. + destruct a; auto. apply Mem.store_signed_unsigned_16. + } + destruct A as (mk_instr & chunk' & B & C & D). + rewrite D in STORE; clear D. + eapply transl_store_access_correct; eauto with asmgen. +Qed. + +(** Function epilogues *) + +Lemma make_epilogue_correct: + forall ge0 f m stk soff cs m' ms rs k tm, + load_stack m (Vptr stk soff) Tptr f.(fn_link_ofs) = Some (parent_sp cs) -> + load_stack m (Vptr stk soff) Tptr f.(fn_retaddr_ofs) = Some (parent_ra cs) -> + Mem.free m stk 0 f.(fn_stacksize) = Some m' -> + agree ms (Vptr stk soff) rs -> + Mem.extends m tm -> + match_stack ge0 cs -> + exists rs', exists tm', + exec_straight ge fn (make_epilogue f k) rs tm k rs' tm' + /\ agree ms (parent_sp cs) rs' + /\ Mem.extends m' tm' + /\ rs'#RA = parent_ra cs + /\ rs'#SP = parent_sp cs + /\ (forall r, r <> PC -> r <> RA -> r <> SP -> r <> X31 -> rs'#r = rs#r). +Proof. + intros until tm; intros LP LRA FREE AG MEXT MCS. + exploit Mem.loadv_extends. eauto. eexact LP. auto. simpl. intros (parent' & LP' & LDP'). + exploit Mem.loadv_extends. eauto. eexact LRA. auto. simpl. intros (ra' & LRA' & LDRA'). + exploit lessdef_parent_sp; eauto. intros EQ; subst parent'; clear LDP'. + exploit lessdef_parent_ra; eauto. intros EQ; subst ra'; clear LDRA'. + exploit Mem.free_parallel_extends; eauto. intros (tm' & FREE' & MEXT'). + unfold make_epilogue. + rewrite chunk_of_Tptr in *. + exploit (loadind_ptr_correct SP (fn_retaddr_ofs f) RA (Pfreeframe (fn_stacksize f) (fn_link_ofs f) :: k) rs tm). + rewrite <- (sp_val _ _ _ AG). simpl. eexact LRA'. congruence. + intros (rs1 & A1 & B1 & C1). + econstructor; econstructor; split. + eapply exec_straight_trans. eexact A1. apply exec_straight_one. simpl. + rewrite (C1 X2) by auto with asmgen. rewrite <- (sp_val _ _ _ AG). simpl; rewrite LP'. + rewrite FREE'. eauto. auto. + split. apply agree_nextinstr. apply agree_set_other; auto with asmgen. + apply agree_change_sp with (Vptr stk soff). + apply agree_exten with rs; auto. intros; apply C1; auto with asmgen. + eapply parent_sp_def; eauto. + split. auto. + split. Simpl. + split. Simpl. + intros. Simpl. +Qed. + +End CONSTRUCTORS. diff --git a/riscV/SelectLongproof.v b/riscV/SelectLongproof.v new file mode 100644 index 00000000..0fc578bf --- /dev/null +++ b/riscV/SelectLongproof.v @@ -0,0 +1,620 @@ +(* *********************************************************************) +(* *) +(* The Compcert verified compiler *) +(* *) +(* Xavier Leroy, INRIA Paris *) +(* Prashanth Mundkur, SRI International *) +(* *) +(* 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. *) +(* *) +(* The contributions by Prashanth Mundkur are reused and adapted *) +(* under the terms of a Contributor License Agreement between *) +(* SRI International and INRIA. *) +(* *) +(* *********************************************************************) + +(** Correctness of instruction selection for 64-bit integer operations *) + +Require Import String Coqlib Maps Integers Floats Errors. +Require Archi. +Require Import AST Values Memory Globalenvs Events. +Require Import Cminor Op CminorSel. +Require Import OpHelpers OpHelpersproof. +Require Import SelectOp SelectOpproof SplitLong SplitLongproof. +Require Import SelectLong. + +Local Open Scope cminorsel_scope. +Local Open Scope string_scope. + +(** * Correctness of the instruction selection functions for 64-bit operators *) + +Section CMCONSTR. + +Variable prog: program. +Variable hf: helper_functions. +Hypothesis HELPERS: helper_functions_declared prog hf. +Let ge := Genv.globalenv prog. +Variable sp: val. +Variable e: env. +Variable m: mem. + +Definition unary_constructor_sound (cstr: expr -> expr) (sem: val -> val) : Prop := + forall le a x, + eval_expr ge sp e m le a x -> + exists v, eval_expr ge sp e m le (cstr a) v /\ Val.lessdef (sem x) v. + +Definition binary_constructor_sound (cstr: expr -> expr -> expr) (sem: val -> val -> val) : Prop := + forall le a x b y, + eval_expr ge sp e m le a x -> + eval_expr ge sp e m le b y -> + exists v, eval_expr ge sp e m le (cstr a b) v /\ Val.lessdef (sem x y) v. + +Definition partial_unary_constructor_sound (cstr: expr -> expr) (sem: val -> option val) : Prop := + forall le a x y, + eval_expr ge sp e m le a x -> + sem x = Some y -> + exists v, eval_expr ge sp e m le (cstr a) v /\ Val.lessdef y v. + +Definition partial_binary_constructor_sound (cstr: expr -> expr -> expr) (sem: val -> val -> option val) : Prop := + forall le a x b y z, + eval_expr ge sp e m le a x -> + eval_expr ge sp e m le b y -> + sem x y = Some z -> + exists v, eval_expr ge sp e m le (cstr a b) v /\ Val.lessdef z v. + +Theorem eval_longconst: + forall le n, eval_expr ge sp e m le (longconst n) (Vlong n). +Proof. + unfold longconst; intros; destruct Archi.splitlong. + apply SplitLongproof.eval_longconst. + EvalOp. +Qed. + +Lemma is_longconst_sound: + forall v a n le, + is_longconst a = Some n -> eval_expr ge sp e m le a v -> v = Vlong n. +Proof with (try discriminate). + intros. unfold is_longconst in *. destruct Archi.splitlong. + eapply SplitLongproof.is_longconst_sound; eauto. + assert (a = Eop (Olongconst n) Enil). + { destruct a... destruct o... destruct e0... congruence. } + subst a. InvEval. auto. +Qed. + +Theorem eval_intoflong: unary_constructor_sound intoflong Val.loword. +Proof. + unfold intoflong; destruct Archi.splitlong. apply SplitLongproof.eval_intoflong. + red; intros. destruct (is_longconst a) as [n|] eqn:C. +- TrivialExists. simpl. erewrite (is_longconst_sound x) by eauto. auto. +- TrivialExists. +Qed. + +Theorem eval_longofintu: unary_constructor_sound longofintu Val.longofintu. +Proof. + unfold longofintu; destruct Archi.splitlong. apply SplitLongproof.eval_longofintu. + red; intros. destruct (is_intconst a) as [n|] eqn:C. +- econstructor; split. apply eval_longconst. + exploit is_intconst_sound; eauto. intros; subst x. auto. +- TrivialExists. +Qed. + +Theorem eval_longofint: unary_constructor_sound longofint Val.longofint. +Proof. + unfold longofint; destruct Archi.splitlong. apply SplitLongproof.eval_longofint. + red; intros. destruct (is_intconst a) as [n|] eqn:C. +- econstructor; split. apply eval_longconst. + exploit is_intconst_sound; eauto. intros; subst x. auto. +- TrivialExists. +Qed. + +Theorem eval_negl: unary_constructor_sound negl Val.negl. +Proof. + unfold negl. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_negl; auto. + red; intros. destruct (is_longconst a) as [n|] eqn:C. +- exploit is_longconst_sound; eauto. intros EQ; subst x. + econstructor; split. apply eval_longconst. auto. +- TrivialExists. +Qed. + +Theorem eval_addlimm: forall n, unary_constructor_sound (addlimm n) (fun v => Val.addl v (Vlong n)). +Proof. + unfold addlimm; intros; red; intros. + predSpec Int64.eq Int64.eq_spec n Int64.zero. + subst. exists x; split; auto. + destruct x; simpl; rewrite ?Int64.add_zero, ?Ptrofs.add_zero; auto. + destruct Archi.ptr64; auto. + destruct (addlimm_match a); InvEval. +- econstructor; split. apply eval_longconst. rewrite Int64.add_commut; auto. +- econstructor; split. EvalOp. simpl; eauto. + unfold Genv.symbol_address. destruct (Genv.find_symbol ge s); simpl; auto. + destruct Archi.ptr64; auto. rewrite Ptrofs.add_commut; auto. +- econstructor; split. EvalOp. simpl; eauto. + destruct sp; simpl; auto. destruct Archi.ptr64; auto. + rewrite Ptrofs.add_assoc, (Ptrofs.add_commut m0). auto. +- subst x. rewrite Val.addl_assoc. rewrite Int64.add_commut. TrivialExists. +- TrivialExists. +Qed. + +Theorem eval_addl: binary_constructor_sound addl Val.addl. +Proof. + unfold addl. destruct Archi.splitlong eqn:SL. + apply SplitLongproof.eval_addl. apply Archi.splitlong_ptr32; auto. +(* + assert (SF: Archi.ptr64 = true). + { Local Transparent Archi.splitlong. unfold Archi.splitlong in SL. + destruct Archi.ptr64; simpl in *; congruence. } +*) +(* + assert (B: forall id ofs n, + Genv.symbol_address ge id (Ptrofs.add ofs (Ptrofs.repr n)) = + Val.addl (Genv.symbol_address ge id ofs) (Vlong (Int64.repr n))). + { intros. replace (Ptrofs.repr n) with (Ptrofs.of_int64 (Int64.repr n)) by auto with ptrofs. + apply Genv.shift_symbol_address_64; auto. } + +*) + red; intros until y. + case (addl_match a b); intros; InvEval. + - rewrite Val.addl_commut. apply eval_addlimm; auto. + - apply eval_addlimm; auto. + - subst. + replace (Val.addl (Val.addl v1 (Vlong n1)) (Val.addl v0 (Vlong n2))) + with (Val.addl (Val.addl v1 v0) (Val.addl (Vlong n1) (Vlong n2))). + apply eval_addlimm. EvalOp. + repeat rewrite Val.addl_assoc. decEq. apply Val.addl_permut. + - subst. econstructor; split. + EvalOp. constructor. EvalOp. simpl; eauto. constructor. eauto. constructor. simpl; eauto. + rewrite Val.addl_commut. destruct sp; simpl; auto. + destruct v1; simpl; auto. + destruct Archi.ptr64 eqn:SF; auto. + apply Val.lessdef_same. f_equal. rewrite ! Ptrofs.add_assoc. f_equal. + rewrite (Ptrofs.add_commut (Ptrofs.of_int64 n1)), Ptrofs.add_assoc. f_equal. auto with ptrofs. + destruct Archi.ptr64 eqn:SF; auto. + - subst. econstructor; split. + EvalOp. constructor. EvalOp. simpl; eauto. constructor. eauto. constructor. simpl; eauto. + destruct sp; simpl; auto. + destruct v1; simpl; auto. + destruct Archi.ptr64 eqn:SF; auto. + apply Val.lessdef_same. f_equal. rewrite ! Ptrofs.add_assoc. f_equal. f_equal. + rewrite Ptrofs.add_commut. auto with ptrofs. + destruct Archi.ptr64 eqn:SF; auto. + - subst. + replace (Val.addl (Val.addl v1 (Vlong n1)) y) + with (Val.addl (Val.addl v1 y) (Vlong n1)). + apply eval_addlimm. EvalOp. + repeat rewrite Val.addl_assoc. decEq. apply Val.addl_commut. + - subst. + replace (Val.addl x (Val.addl v1 (Vlong n2))) + with (Val.addl (Val.addl x v1) (Vlong n2)). + apply eval_addlimm. EvalOp. + repeat rewrite Val.addl_assoc. reflexivity. + - TrivialExists. +Qed. + +Theorem eval_subl: binary_constructor_sound subl Val.subl. +Proof. + unfold subl. destruct Archi.splitlong eqn:SL. + apply SplitLongproof.eval_subl. apply Archi.splitlong_ptr32; auto. + red; intros; destruct (subl_match a b); InvEval. +- rewrite Val.subl_addl_opp. apply eval_addlimm; auto. +- subst. rewrite Val.subl_addl_l. rewrite Val.subl_addl_r. + rewrite Val.addl_assoc. simpl. rewrite Int64.add_commut. rewrite <- Int64.sub_add_opp. + apply eval_addlimm; EvalOp. +- subst. rewrite Val.subl_addl_l. apply eval_addlimm; EvalOp. +- subst. rewrite Val.subl_addl_r. + apply eval_addlimm; EvalOp. +- TrivialExists. +Qed. + +Theorem eval_shllimm: forall n, unary_constructor_sound (fun e => shllimm e n) (fun v => Val.shll v (Vint n)). +Proof. + intros; unfold shllimm. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_shllimm; auto. + red; intros. + predSpec Int.eq Int.eq_spec n Int.zero. + exists x; split; auto. subst n; destruct x; simpl; auto. + destruct (Int.ltu Int.zero Int64.iwordsize'); auto. + change (Int64.shl' i Int.zero) with (Int64.shl i Int64.zero). rewrite Int64.shl_zero; auto. + destruct (Int.ltu n Int64.iwordsize') eqn:LT; simpl. + assert (DEFAULT: exists v, eval_expr ge sp e m le (Eop (Oshllimm n) (a:::Enil)) v + /\ Val.lessdef (Val.shll x (Vint n)) v) by TrivialExists. + destruct (shllimm_match a); InvEval. +- econstructor; split. apply eval_longconst. simpl; rewrite LT; auto. +- destruct (Int.ltu (Int.add n n1) Int64.iwordsize') eqn:LT'; auto. + subst. econstructor; split. EvalOp. simpl; eauto. + destruct v1; simpl; auto. rewrite LT'. + destruct (Int.ltu n1 Int64.iwordsize') eqn:LT1; auto. + simpl; rewrite LT. rewrite Int.add_commut, Int64.shl'_shl'; auto. rewrite Int.add_commut; auto. +- apply DEFAULT. +- TrivialExists. constructor; eauto. constructor. EvalOp. simpl; eauto. constructor. auto. +Qed. + +Theorem eval_shrluimm: forall n, unary_constructor_sound (fun e => shrluimm e n) (fun v => Val.shrlu v (Vint n)). +Proof. + intros; unfold shrluimm. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_shrluimm; auto. + red; intros. + predSpec Int.eq Int.eq_spec n Int.zero. + exists x; split; auto. subst n; destruct x; simpl; auto. + destruct (Int.ltu Int.zero Int64.iwordsize'); auto. + change (Int64.shru' i Int.zero) with (Int64.shru i Int64.zero). rewrite Int64.shru_zero; auto. + destruct (Int.ltu n Int64.iwordsize') eqn:LT; simpl. + assert (DEFAULT: exists v, eval_expr ge sp e m le (Eop (Oshrluimm n) (a:::Enil)) v + /\ Val.lessdef (Val.shrlu x (Vint n)) v) by TrivialExists. + destruct (shrluimm_match a); InvEval. +- econstructor; split. apply eval_longconst. simpl; rewrite LT; auto. +- destruct (Int.ltu (Int.add n n1) Int64.iwordsize') eqn:LT'; auto. + subst. econstructor; split. EvalOp. simpl; eauto. + destruct v1; simpl; auto. rewrite LT'. + destruct (Int.ltu n1 Int64.iwordsize') eqn:LT1; auto. + simpl; rewrite LT. rewrite Int.add_commut, Int64.shru'_shru'; auto. rewrite Int.add_commut; auto. +- apply DEFAULT. +- TrivialExists. constructor; eauto. constructor. EvalOp. simpl; eauto. constructor. auto. +Qed. + +Theorem eval_shrlimm: forall n, unary_constructor_sound (fun e => shrlimm e n) (fun v => Val.shrl v (Vint n)). +Proof. + intros; unfold shrlimm. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_shrlimm; auto. + red; intros. + predSpec Int.eq Int.eq_spec n Int.zero. + exists x; split; auto. subst n; destruct x; simpl; auto. + destruct (Int.ltu Int.zero Int64.iwordsize'); auto. + change (Int64.shr' i Int.zero) with (Int64.shr i Int64.zero). rewrite Int64.shr_zero; auto. + destruct (Int.ltu n Int64.iwordsize') eqn:LT; simpl. + assert (DEFAULT: exists v, eval_expr ge sp e m le (Eop (Oshrlimm n) (a:::Enil)) v + /\ Val.lessdef (Val.shrl x (Vint n)) v) by TrivialExists. + destruct (shrlimm_match a); InvEval. +- econstructor; split. apply eval_longconst. simpl; rewrite LT; auto. +- destruct (Int.ltu (Int.add n n1) Int64.iwordsize') eqn:LT'; auto. + subst. econstructor; split. EvalOp. simpl; eauto. + destruct v1; simpl; auto. rewrite LT'. + destruct (Int.ltu n1 Int64.iwordsize') eqn:LT1; auto. + simpl; rewrite LT. rewrite Int.add_commut, Int64.shr'_shr'; auto. rewrite Int.add_commut; auto. +- apply DEFAULT. +- TrivialExists. constructor; eauto. constructor. EvalOp. simpl; eauto. constructor. auto. +Qed. + +Theorem eval_shll: binary_constructor_sound shll Val.shll. +Proof. + unfold shll. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_shll; auto. + red; intros. destruct (is_intconst b) as [n2|] eqn:C. +- exploit is_intconst_sound; eauto. intros EQ; subst y. apply eval_shllimm; auto. +- TrivialExists. +Qed. + +Theorem eval_shrlu: binary_constructor_sound shrlu Val.shrlu. +Proof. + unfold shrlu. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_shrlu; auto. + red; intros. destruct (is_intconst b) as [n2|] eqn:C. +- exploit is_intconst_sound; eauto. intros EQ; subst y. apply eval_shrluimm; auto. +- TrivialExists. +Qed. + +Theorem eval_shrl: binary_constructor_sound shrl Val.shrl. +Proof. + unfold shrl. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_shrl; auto. + red; intros. destruct (is_intconst b) as [n2|] eqn:C. +- exploit is_intconst_sound; eauto. intros EQ; subst y. apply eval_shrlimm; auto. +- TrivialExists. +Qed. + +Theorem eval_mullimm_base: forall n, unary_constructor_sound (mullimm_base n) (fun v => Val.mull v (Vlong n)). +Proof. + intros; unfold mullimm_base. red; intros. + assert (DEFAULT: exists v, + eval_expr ge sp e m le (Eop Omull (a ::: longconst n ::: Enil)) v + /\ Val.lessdef (Val.mull x (Vlong n)) v). + { econstructor; split. EvalOp. constructor. eauto. constructor. apply eval_longconst. constructor. simpl; eauto. + auto. } + generalize (Int64.one_bits'_decomp n); intros D. + destruct (Int64.one_bits' n) as [ | i [ | j [ | ? ? ]]] eqn:B. +- apply DEFAULT. +- replace (Val.mull x (Vlong n)) with (Val.shll x (Vint i)). + apply eval_shllimm; auto. + simpl in D. rewrite D, Int64.add_zero. destruct x; simpl; auto. + rewrite (Int64.one_bits'_range n) by (rewrite B; auto with coqlib). + rewrite Int64.shl'_mul; auto. +- set (le' := x :: le). + assert (A0: eval_expr ge sp e m le' (Eletvar O) x) by (constructor; reflexivity). + exploit (eval_shllimm i). eexact A0. intros (v1 & A1 & B1). + exploit (eval_shllimm j). eexact A0. intros (v2 & A2 & B2). + exploit (eval_addl). eexact A1. eexact A2. intros (v3 & A3 & B3). + exists v3; split. econstructor; eauto. + rewrite D. simpl. rewrite Int64.add_zero. destruct x; auto. + simpl in *. + rewrite (Int64.one_bits'_range n) in B1 by (rewrite B; auto with coqlib). + rewrite (Int64.one_bits'_range n) in B2 by (rewrite B; auto with coqlib). + inv B1; inv B2. simpl in B3; inv B3. + rewrite Int64.mul_add_distr_r. rewrite <- ! Int64.shl'_mul. auto. +- apply DEFAULT. +Qed. + +Theorem eval_mullimm: forall n, unary_constructor_sound (mullimm n) (fun v => Val.mull v (Vlong n)). +Proof. + unfold mullimm. intros; red; intros. + destruct Archi.splitlong eqn:SL. + eapply SplitLongproof.eval_mullimm; eauto. + predSpec Int64.eq Int64.eq_spec n Int64.zero. + exists (Vlong Int64.zero); split. apply eval_longconst. + destruct x; simpl; auto. subst n; rewrite Int64.mul_zero; auto. + predSpec Int64.eq Int64.eq_spec n Int64.one. + exists x; split; auto. + destruct x; simpl; auto. subst n; rewrite Int64.mul_one; auto. + destruct (mullimm_match a); InvEval. +- econstructor; split. apply eval_longconst. rewrite Int64.mul_commut; auto. +- exploit (eval_mullimm_base n); eauto. intros (v2 & A2 & B2). + exploit (eval_addlimm (Int64.mul n n2)). eexact A2. intros (v3 & A3 & B3). + exists v3; split; auto. + subst x. destruct v1; simpl; auto. + simpl in B2; inv B2. simpl in B3; inv B3. rewrite Int64.mul_add_distr_l. + rewrite (Int64.mul_commut n). auto. + destruct Archi.ptr64; simpl; auto. +- apply eval_mullimm_base; auto. +Qed. + +Theorem eval_mull: binary_constructor_sound mull Val.mull. +Proof. + unfold mull. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_mull; auto. + red; intros; destruct (mull_match a b); InvEval. +- rewrite Val.mull_commut. apply eval_mullimm; auto. +- apply eval_mullimm; auto. +- TrivialExists. +Qed. + +Theorem eval_mullhu: + forall n, unary_constructor_sound (fun a => mullhu a n) (fun v => Val.mullhu v (Vlong n)). +Proof. + unfold mullhu; intros. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_mullhu; auto. + red; intros. TrivialExists. constructor. eauto. constructor. apply eval_longconst. constructor. auto. +Qed. + +Theorem eval_mullhs: + forall n, unary_constructor_sound (fun a => mullhs a n) (fun v => Val.mullhs v (Vlong n)). +Proof. + unfold mullhs; intros. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_mullhs; auto. + red; intros. TrivialExists. constructor. eauto. constructor. apply eval_longconst. constructor. auto. +Qed. + +Theorem eval_andlimm: forall n, unary_constructor_sound (andlimm n) (fun v => Val.andl v (Vlong n)). +Proof. + unfold andlimm; intros; red; intros. + predSpec Int64.eq Int64.eq_spec n Int64.zero. + exists (Vlong Int64.zero); split. apply eval_longconst. + subst. destruct x; simpl; auto. rewrite Int64.and_zero; auto. + predSpec Int64.eq Int64.eq_spec n Int64.mone. + exists x; split. assumption. + subst. destruct x; simpl; auto. rewrite Int64.and_mone; auto. + destruct (andlimm_match a); InvEval; subst. +- econstructor; split. apply eval_longconst. simpl. rewrite Int64.and_commut; auto. +- TrivialExists. simpl. rewrite Val.andl_assoc. rewrite Int64.and_commut; auto. +- TrivialExists. +Qed. + +Theorem eval_andl: binary_constructor_sound andl Val.andl. +Proof. + unfold andl; destruct Archi.splitlong. apply SplitLongproof.eval_andl. + red; intros. destruct (andl_match a b). +- InvEval. rewrite Val.andl_commut. apply eval_andlimm; auto. +- InvEval. apply eval_andlimm; auto. +- TrivialExists. +Qed. + +Theorem eval_orlimm: forall n, unary_constructor_sound (orlimm n) (fun v => Val.orl v (Vlong n)). +Proof. + unfold orlimm; intros; red; intros. + predSpec Int64.eq Int64.eq_spec n Int64.zero. + exists x; split; auto. subst. destruct x; simpl; auto. rewrite Int64.or_zero; auto. + predSpec Int64.eq Int64.eq_spec n Int64.mone. + econstructor; split. apply eval_longconst. subst. destruct x; simpl; auto. rewrite Int64.or_mone; auto. + destruct (orlimm_match a); InvEval; subst. +- econstructor; split. apply eval_longconst. simpl. rewrite Int64.or_commut; auto. +- TrivialExists. simpl. rewrite Val.orl_assoc. rewrite Int64.or_commut; auto. +- TrivialExists. +Qed. + +Theorem eval_orl: binary_constructor_sound orl Val.orl. +Proof. + unfold orl; destruct Archi.splitlong. apply SplitLongproof.eval_orl. + red; intros. + destruct (orl_match a b). +- InvEval. rewrite Val.orl_commut. apply eval_orlimm; auto. +- InvEval. apply eval_orlimm; auto. +- TrivialExists. +Qed. + +Theorem eval_xorlimm: forall n, unary_constructor_sound (xorlimm n) (fun v => Val.xorl v (Vlong n)). +Proof. + unfold xorlimm; intros; red; intros. + predSpec Int64.eq Int64.eq_spec n Int64.zero. + exists x; split; auto. subst. destruct x; simpl; auto. rewrite Int64.xor_zero; auto. + destruct (xorlimm_match a); InvEval; subst. +- econstructor; split. apply eval_longconst. simpl. rewrite Int64.xor_commut; auto. +- rewrite Val.xorl_assoc. simpl. rewrite (Int64.xor_commut n2). + predSpec Int64.eq Int64.eq_spec (Int64.xor n n2) Int64.zero. ++ rewrite H. exists v1; split; auto. destruct v1; simpl; auto. rewrite Int64.xor_zero; auto. ++ TrivialExists. +- TrivialExists. +Qed. + +Theorem eval_xorl: binary_constructor_sound xorl Val.xorl. +Proof. + unfold xorl; destruct Archi.splitlong. apply SplitLongproof.eval_xorl. + red; intros. destruct (xorl_match a b). +- InvEval. rewrite Val.xorl_commut. apply eval_xorlimm; auto. +- InvEval. apply eval_xorlimm; auto. +- TrivialExists. +Qed. + +Theorem eval_notl: unary_constructor_sound notl Val.notl. +Proof. + unfold notl; destruct Archi.splitlong. apply SplitLongproof.eval_notl. + red; intros. rewrite Val.notl_xorl. apply eval_xorlimm; auto. +Qed. + +Theorem eval_divls_base: partial_binary_constructor_sound divls_base Val.divls. +Proof. + unfold divls_base; red; intros. destruct Archi.splitlong eqn:SL. + eapply SplitLongproof.eval_divls_base; eauto. + TrivialExists. + cbn. + rewrite H1. + cbn. + trivial. +Qed. + +Theorem eval_modls_base: partial_binary_constructor_sound modls_base Val.modls. +Proof. + unfold modls_base; red; intros. destruct Archi.splitlong eqn:SL. + eapply SplitLongproof.eval_modls_base; eauto. + TrivialExists. + cbn. + rewrite H1. + cbn. + trivial. +Qed. + +Theorem eval_divlu_base: partial_binary_constructor_sound divlu_base Val.divlu. +Proof. + unfold divlu_base; red; intros. destruct Archi.splitlong eqn:SL. + eapply SplitLongproof.eval_divlu_base; eauto. + TrivialExists. + cbn. + rewrite H1. + cbn. + trivial. +Qed. + +Theorem eval_modlu_base: partial_binary_constructor_sound modlu_base Val.modlu. +Proof. + unfold modlu_base; red; intros. destruct Archi.splitlong eqn:SL. + eapply SplitLongproof.eval_modlu_base; eauto. + TrivialExists. + cbn. + rewrite H1. + cbn. + trivial. +Qed. + +Theorem eval_shrxlimm: + forall le a n x z, + eval_expr ge sp e m le a x -> + Val.shrxl x (Vint n) = Some z -> + exists v, eval_expr ge sp e m le (shrxlimm a n) v /\ Val.lessdef z v. +Proof. + unfold shrxlimm; intros. destruct Archi.splitlong eqn:SL. ++ eapply SplitLongproof.eval_shrxlimm; eauto using Archi.splitlong_ptr32. ++ predSpec Int.eq Int.eq_spec n Int.zero. +- subst n. destruct x; simpl in H0; inv H0. econstructor; split; eauto. + change (Int.ltu Int.zero (Int.repr 63)) with true. simpl. rewrite Int64.shrx'_zero; auto. +- TrivialExists. + cbn. + rewrite H0. + reflexivity. +Qed. + +Theorem eval_cmplu: + forall c le a x b y v, + eval_expr ge sp e m le a x -> + eval_expr ge sp e m le b y -> + Val.cmplu (Mem.valid_pointer m) c x y = Some v -> + eval_expr ge sp e m le (cmplu c a b) v. +Proof. + unfold cmplu; intros. destruct Archi.splitlong eqn:SL. + eapply SplitLongproof.eval_cmplu; eauto using Archi.splitlong_ptr32. + unfold Val.cmplu in H1. + destruct (Val.cmplu_bool (Mem.valid_pointer m) c x y) as [vb|] eqn:C; simpl in H1; inv H1. + destruct (is_longconst a) as [n1|] eqn:LC1; destruct (is_longconst b) as [n2|] eqn:LC2; + try (assert (x = Vlong n1) by (eapply is_longconst_sound; eauto)); + try (assert (y = Vlong n2) by (eapply is_longconst_sound; eauto)); + subst. +- simpl in C; inv C. EvalOp. destruct (Int64.cmpu c n1 n2); reflexivity. +- EvalOp. simpl. rewrite Val.swap_cmplu_bool. rewrite C; auto. +- EvalOp. simpl; rewrite C; auto. +- EvalOp. simpl; rewrite C; auto. +Qed. + +Theorem eval_cmpl: + forall c le a x b y v, + eval_expr ge sp e m le a x -> + eval_expr ge sp e m le b y -> + Val.cmpl c x y = Some v -> + eval_expr ge sp e m le (cmpl c a b) v. +Proof. + unfold cmpl; intros. destruct Archi.splitlong eqn:SL. + eapply SplitLongproof.eval_cmpl; eauto. + unfold Val.cmpl in H1. + destruct (Val.cmpl_bool c x y) as [vb|] eqn:C; simpl in H1; inv H1. + destruct (is_longconst a) as [n1|] eqn:LC1; destruct (is_longconst b) as [n2|] eqn:LC2; + try (assert (x = Vlong n1) by (eapply is_longconst_sound; eauto)); + try (assert (y = Vlong n2) by (eapply is_longconst_sound; eauto)); + subst. +- simpl in C; inv C. EvalOp. destruct (Int64.cmp c n1 n2); reflexivity. +- EvalOp. simpl. rewrite Val.swap_cmpl_bool. rewrite C; auto. +- EvalOp. simpl; rewrite C; auto. +- EvalOp. simpl; rewrite C; auto. +Qed. + +Theorem eval_longoffloat: partial_unary_constructor_sound longoffloat Val.longoffloat. +Proof. + unfold longoffloat; red; intros. destruct Archi.splitlong eqn:SL. + eapply SplitLongproof.eval_longoffloat; eauto. + TrivialExists. + cbn; rewrite H0; reflexivity. +Qed. + +Theorem eval_longuoffloat: partial_unary_constructor_sound longuoffloat Val.longuoffloat. +Proof. + unfold longuoffloat; red; intros. destruct Archi.splitlong eqn:SL. + eapply SplitLongproof.eval_longuoffloat; eauto. + TrivialExists. + cbn; rewrite H0; reflexivity. +Qed. + +Theorem eval_floatoflong: partial_unary_constructor_sound floatoflong Val.floatoflong. +Proof. + unfold floatoflong; red; intros. destruct Archi.splitlong eqn:SL. + eapply SplitLongproof.eval_floatoflong; eauto. + TrivialExists. + cbn; rewrite H0; reflexivity. +Qed. + +Theorem eval_floatoflongu: partial_unary_constructor_sound floatoflongu Val.floatoflongu. +Proof. + unfold floatoflongu; red; intros. destruct Archi.splitlong eqn:SL. + eapply SplitLongproof.eval_floatoflongu; eauto. + TrivialExists. + cbn; rewrite H0; reflexivity. +Qed. + +Theorem eval_longofsingle: partial_unary_constructor_sound longofsingle Val.longofsingle. +Proof. + unfold longofsingle; red; intros. destruct Archi.splitlong eqn:SL. + eapply SplitLongproof.eval_longofsingle; eauto. + TrivialExists. + cbn; rewrite H0; reflexivity. +Qed. + +Theorem eval_longuofsingle: partial_unary_constructor_sound longuofsingle Val.longuofsingle. +Proof. + unfold longuofsingle; red; intros. destruct Archi.splitlong eqn:SL. + eapply SplitLongproof.eval_longuofsingle; eauto. + TrivialExists. + cbn; rewrite H0; reflexivity. +Qed. + +Theorem eval_singleoflong: partial_unary_constructor_sound singleoflong Val.singleoflong. +Proof. + unfold singleoflong; red; intros. destruct Archi.splitlong eqn:SL. + eapply SplitLongproof.eval_singleoflong; eauto. + TrivialExists. + cbn; rewrite H0; reflexivity. +Qed. + +Theorem eval_singleoflongu: partial_unary_constructor_sound singleoflongu Val.singleoflongu. +Proof. + unfold singleoflongu; red; intros. destruct Archi.splitlong eqn:SL. + eapply SplitLongproof.eval_singleoflongu; eauto. + TrivialExists. + cbn; rewrite H0; reflexivity. +Qed. + +End CMCONSTR. diff --git a/riscV/SelectOpproof.v b/riscV/SelectOpproof.v new file mode 100644 index 00000000..f450fe6c --- /dev/null +++ b/riscV/SelectOpproof.v @@ -0,0 +1,1093 @@ +(* *********************************************************************) +(* *) +(* The Compcert verified compiler *) +(* *) +(* Xavier Leroy, INRIA Paris-Rocquencourt *) +(* Prashanth Mundkur, SRI International *) +(* *) +(* 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. *) +(* *) +(* The contributions by Prashanth Mundkur are reused and adapted *) +(* under the terms of a Contributor License Agreement between *) +(* SRI International and INRIA. *) +(* *) +(* *********************************************************************) + +(** Correctness of instruction selection for operators *) + +Require Import Coqlib Zbits. +Require Import AST Integers Floats. +Require Import Values Memory Builtins Globalenvs. +Require Import Cminor Op CminorSel. +Require Import SelectOp. +Require Import OpHelpers. +Require Import OpHelpersproof. +Require Import Lia. + +Local Open Scope cminorsel_scope. + +(** * Useful lemmas and tactics *) + +(** The following are trivial lemmas and custom tactics that help + perform backward (inversion) and forward reasoning over the evaluation + of operator applications. *) + +Ltac EvalOp := eapply eval_Eop; eauto with evalexpr. + +Ltac InvEval1 := + match goal with + | [ H: (eval_expr _ _ _ _ _ (Eop _ Enil) _) |- _ ] => + inv H; InvEval1 + | [ H: (eval_expr _ _ _ _ _ (Eop _ (_ ::: Enil)) _) |- _ ] => + inv H; InvEval1 + | [ H: (eval_expr _ _ _ _ _ (Eop _ (_ ::: _ ::: Enil)) _) |- _ ] => + inv H; InvEval1 + | [ H: (eval_exprlist _ _ _ _ _ Enil _) |- _ ] => + inv H; InvEval1 + | [ H: (eval_exprlist _ _ _ _ _ (_ ::: _) _) |- _ ] => + inv H; InvEval1 + | _ => + idtac + end. + +Ltac InvEval2 := + match goal with + | [ H: (eval_operation _ _ _ nil _ = Some _) |- _ ] => + simpl in H; inv H + | [ H: (eval_operation _ _ _ (_ :: nil) _ = Some _) |- _ ] => + simpl in H; FuncInv + | [ H: (eval_operation _ _ _ (_ :: _ :: nil) _ = Some _) |- _ ] => + simpl in H; FuncInv + | [ H: (eval_operation _ _ _ (_ :: _ :: _ :: nil) _ = Some _) |- _ ] => + simpl in H; FuncInv + | _ => + idtac + end. + +Ltac InvEval := InvEval1; InvEval2; InvEval2. + +Ltac TrivialExists := + match goal with + | [ |- exists v, _ /\ Val.lessdef ?a v ] => exists a; split; [EvalOp | auto] + end. + +(** * Correctness of the smart constructors *) + +Section CMCONSTR. +Variable prog: program. +Variable hf: helper_functions. +Hypothesis HELPERS: helper_functions_declared prog hf. +Let ge := Genv.globalenv prog. +Variable sp: val. +Variable e: env. +Variable m: mem. + +(** We now show that the code generated by "smart constructor" functions + such as [Selection.notint] behaves as expected. Continuing the + [notint] example, we show that if the expression [e] + evaluates to some integer value [Vint n], then [Selection.notint e] + evaluates to a value [Vint (Int.not n)] which is indeed the integer + negation of the value of [e]. + + All proofs follow a common pattern: +- Reasoning by case over the result of the classification functions + (such as [add_match] for integer addition), gathering additional + information on the shape of the argument expressions in the non-default + cases. +- Inversion of the evaluations of the arguments, exploiting the additional + information thus gathered. +- Equational reasoning over the arithmetic operations performed, + using the lemmas from the [Int] and [Float] modules. +- Construction of an evaluation derivation for the expression returned + by the smart constructor. +*) + +Definition unary_constructor_sound (cstr: expr -> expr) (sem: val -> val) : Prop := + forall le a x, + eval_expr ge sp e m le a x -> + exists v, eval_expr ge sp e m le (cstr a) v /\ Val.lessdef (sem x) v. + +Definition binary_constructor_sound (cstr: expr -> expr -> expr) (sem: val -> val -> val) : Prop := + forall le a x b y, + eval_expr ge sp e m le a x -> + eval_expr ge sp e m le b y -> + exists v, eval_expr ge sp e m le (cstr a b) v /\ Val.lessdef (sem x y) v. + +Theorem eval_addrsymbol: + forall le id ofs, + exists v, eval_expr ge sp e m le (addrsymbol id ofs) v /\ Val.lessdef (Genv.symbol_address ge id ofs) v. +Proof. + intros. unfold addrsymbol. econstructor; split. + EvalOp. simpl; eauto. + auto. +Qed. + +Theorem eval_addrstack: + forall le ofs, + exists v, eval_expr ge sp e m le (addrstack ofs) v /\ Val.lessdef (Val.offset_ptr sp ofs) v. +Proof. + intros. unfold addrstack. econstructor; split. + EvalOp. simpl; eauto. + auto. +Qed. + +Theorem eval_addimm: + forall n, unary_constructor_sound (addimm n) (fun x => Val.add x (Vint n)). +Proof. + red; unfold addimm; intros until x. + predSpec Int.eq Int.eq_spec n Int.zero. + - subst n. intros. exists x; split; auto. + destruct x; simpl; auto. + rewrite Int.add_zero; auto. + destruct Archi.ptr64; auto. rewrite Ptrofs.add_zero; auto. + - case (addimm_match a); intros; InvEval; simpl. + + TrivialExists; simpl. rewrite Int.add_commut. auto. + + econstructor; split. EvalOp. simpl; eauto. + unfold Genv.symbol_address. destruct (Genv.find_symbol ge s); simpl; auto. + destruct Archi.ptr64; auto. rewrite Ptrofs.add_commut; auto. + + econstructor; split. EvalOp. simpl; eauto. + destruct sp; simpl; auto. destruct Archi.ptr64; auto. + rewrite Ptrofs.add_assoc. rewrite (Ptrofs.add_commut m0). auto. + + TrivialExists; simpl. subst x. rewrite Val.add_assoc. rewrite Int.add_commut. auto. + + TrivialExists. +Qed. + +Theorem eval_add: binary_constructor_sound add Val.add. +Proof. + red; intros until y. + unfold add; case (add_match a b); intros; InvEval. + - rewrite Val.add_commut. apply eval_addimm; auto. + - apply eval_addimm; auto. + - subst. + replace (Val.add (Val.add v1 (Vint n1)) (Val.add v0 (Vint n2))) + with (Val.add (Val.add v1 v0) (Val.add (Vint n1) (Vint n2))). + apply eval_addimm. EvalOp. + repeat rewrite Val.add_assoc. decEq. apply Val.add_permut. + - subst. econstructor; split. + EvalOp. constructor. EvalOp. simpl; eauto. constructor. eauto. constructor. simpl; eauto. + rewrite Val.add_commut. destruct sp; simpl; auto. + destruct v1; simpl; auto. + destruct Archi.ptr64 eqn:SF; auto. + apply Val.lessdef_same. f_equal. rewrite ! Ptrofs.add_assoc. f_equal. + rewrite (Ptrofs.add_commut (Ptrofs.of_int n1)), Ptrofs.add_assoc. f_equal. auto with ptrofs. + destruct Archi.ptr64 eqn:SF; auto. + - subst. econstructor; split. + EvalOp. constructor. EvalOp. simpl; eauto. constructor. eauto. constructor. simpl; eauto. + destruct sp; simpl; auto. + destruct v1; simpl; auto. + destruct Archi.ptr64 eqn:SF; auto. + apply Val.lessdef_same. f_equal. rewrite ! Ptrofs.add_assoc. f_equal. f_equal. + rewrite Ptrofs.add_commut. auto with ptrofs. + destruct Archi.ptr64 eqn:SF; auto. + - subst. + replace (Val.add (Val.add v1 (Vint n1)) y) + with (Val.add (Val.add v1 y) (Vint n1)). + apply eval_addimm. EvalOp. + repeat rewrite Val.add_assoc. decEq. apply Val.add_commut. + - subst. + replace (Val.add x (Val.add v1 (Vint n2))) + with (Val.add (Val.add x v1) (Vint n2)). + apply eval_addimm. EvalOp. + repeat rewrite Val.add_assoc. reflexivity. + - TrivialExists. +Qed. + +Theorem eval_sub: binary_constructor_sound sub Val.sub. +Proof. + red; intros until y. + unfold sub; case (sub_match a b); intros; InvEval. + - rewrite Val.sub_add_opp. apply eval_addimm; auto. + - subst. rewrite Val.sub_add_l. rewrite Val.sub_add_r. + rewrite Val.add_assoc. simpl. rewrite Int.add_commut. rewrite <- Int.sub_add_opp. + apply eval_addimm; EvalOp. + - subst. rewrite Val.sub_add_l. apply eval_addimm; EvalOp. + - subst. rewrite Val.sub_add_r. apply eval_addimm; EvalOp. + - TrivialExists. +Qed. + +Theorem eval_negint: unary_constructor_sound negint (fun v => Val.sub Vzero v). +Proof. + red; intros until x. unfold negint. case (negint_match a); intros; InvEval. + TrivialExists. + TrivialExists. +Qed. + +Theorem eval_shlimm: + forall n, unary_constructor_sound (fun a => shlimm a n) + (fun x => Val.shl x (Vint n)). +Proof. + red; intros until x. unfold shlimm. + + predSpec Int.eq Int.eq_spec n Int.zero. + intros; subst. exists x; split; auto. destruct x; simpl; auto. rewrite Int.shl_zero; auto. + + destruct (Int.ltu n Int.iwordsize) eqn:LT; simpl. + destruct (shlimm_match a); intros; InvEval. + - exists (Vint (Int.shl n1 n)); split. EvalOp. + simpl. rewrite LT. auto. + - destruct (Int.ltu (Int.add n n1) Int.iwordsize) eqn:?. + + exists (Val.shl v1 (Vint (Int.add n n1))); split. EvalOp. + subst. destruct v1; simpl; auto. + rewrite Heqb. + destruct (Int.ltu n1 Int.iwordsize) eqn:?; simpl; auto. + destruct (Int.ltu n Int.iwordsize) eqn:?; simpl; auto. + rewrite Int.add_commut. rewrite Int.shl_shl; auto. rewrite Int.add_commut; auto. + + subst. TrivialExists. econstructor. EvalOp. simpl; eauto. constructor. + simpl. auto. + - TrivialExists. + - intros; TrivialExists. constructor. eauto. constructor. EvalOp. simpl; eauto. constructor. + auto. +Qed. + +Theorem eval_shruimm: + forall n, unary_constructor_sound (fun a => shruimm a n) + (fun x => Val.shru x (Vint n)). +Proof. + red; intros until x. unfold shruimm. + + predSpec Int.eq Int.eq_spec n Int.zero. + intros; subst. exists x; split; auto. destruct x; simpl; auto. rewrite Int.shru_zero; auto. + + destruct (Int.ltu n Int.iwordsize) eqn:LT; simpl. + destruct (shruimm_match a); intros; InvEval. + - exists (Vint (Int.shru n1 n)); split. EvalOp. + simpl. rewrite LT; auto. + - destruct (Int.ltu (Int.add n n1) Int.iwordsize) eqn:?. + exists (Val.shru v1 (Vint (Int.add n n1))); split. EvalOp. + subst. destruct v1; simpl; auto. + rewrite Heqb. + destruct (Int.ltu n1 Int.iwordsize) eqn:?; simpl; auto. + rewrite LT. rewrite Int.add_commut. rewrite Int.shru_shru; auto. rewrite Int.add_commut; auto. + subst. TrivialExists. econstructor. EvalOp. simpl; eauto. constructor. + simpl. auto. + - TrivialExists. + - intros; TrivialExists. constructor. eauto. constructor. EvalOp. simpl; eauto. constructor. + auto. +Qed. + +Theorem eval_shrimm: + forall n, unary_constructor_sound (fun a => shrimm a n) + (fun x => Val.shr x (Vint n)). +Proof. + red; intros until x. unfold shrimm. + + predSpec Int.eq Int.eq_spec n Int.zero. + intros; subst. exists x; split; auto. destruct x; simpl; auto. rewrite Int.shr_zero; auto. + + destruct (Int.ltu n Int.iwordsize) eqn:LT; simpl. + destruct (shrimm_match a); intros; InvEval. + - exists (Vint (Int.shr n1 n)); split. EvalOp. + simpl. rewrite LT; auto. + - destruct (Int.ltu (Int.add n n1) Int.iwordsize) eqn:?. + exists (Val.shr v1 (Vint (Int.add n n1))); split. EvalOp. + subst. destruct v1; simpl; auto. + rewrite Heqb. + destruct (Int.ltu n1 Int.iwordsize) eqn:?; simpl; auto. + rewrite LT. + rewrite Int.add_commut. rewrite Int.shr_shr; auto. rewrite Int.add_commut; auto. + subst. TrivialExists. econstructor. EvalOp. simpl; eauto. constructor. + simpl. auto. + - TrivialExists. + - intros; TrivialExists. constructor. eauto. constructor. EvalOp. simpl; eauto. constructor. + auto. +Qed. + +Lemma eval_mulimm_base: + forall n, unary_constructor_sound (mulimm_base n) (fun x => Val.mul x (Vint n)). +Proof. + intros; red; intros; unfold mulimm_base. + + assert (DFL: exists v, eval_expr ge sp e m le (Eop Omul (Eop (Ointconst n) Enil ::: a ::: Enil)) v /\ Val.lessdef (Val.mul x (Vint n)) v). + TrivialExists. econstructor. EvalOp. simpl; eauto. econstructor. eauto. constructor. + rewrite Val.mul_commut. auto. + + generalize (Int.one_bits_decomp n). + generalize (Int.one_bits_range n). + destruct (Int.one_bits n). + - intros. auto. + - destruct l. + + intros. rewrite H1. simpl. + rewrite Int.add_zero. + replace (Vint (Int.shl Int.one i)) with (Val.shl Vone (Vint i)). rewrite Val.shl_mul. + apply eval_shlimm. auto. simpl. rewrite H0; auto with coqlib. + + destruct l. + intros. rewrite H1. simpl. + exploit (eval_shlimm i (x :: le) (Eletvar 0) x). constructor; auto. intros [v1 [A1 B1]]. + exploit (eval_shlimm i0 (x :: le) (Eletvar 0) x). constructor; auto. intros [v2 [A2 B2]]. + exploit (eval_add (x :: le)). eexact A1. eexact A2. intros [v [A B]]. + exists v; split. econstructor; eauto. + rewrite Int.add_zero. + replace (Vint (Int.add (Int.shl Int.one i) (Int.shl Int.one i0))) + with (Val.add (Val.shl Vone (Vint i)) (Val.shl Vone (Vint i0))). + rewrite Val.mul_add_distr_r. + repeat rewrite Val.shl_mul. eapply Val.lessdef_trans. 2: eauto. apply Val.add_lessdef; auto. + simpl. repeat rewrite H0; auto with coqlib. + intros. auto. +Qed. + +Theorem eval_mulimm: + forall n, unary_constructor_sound (mulimm n) (fun x => Val.mul x (Vint n)). +Proof. + intros; red; intros until x; unfold mulimm. + + predSpec Int.eq Int.eq_spec n Int.zero. + intros. exists (Vint Int.zero); split. EvalOp. + destruct x; simpl; auto. subst n. rewrite Int.mul_zero. auto. + + predSpec Int.eq Int.eq_spec n Int.one. + intros. exists x; split; auto. + destruct x; simpl; auto. subst n. rewrite Int.mul_one. auto. + + case (mulimm_match a); intros; InvEval. + - TrivialExists. simpl. rewrite Int.mul_commut; auto. + - subst. rewrite Val.mul_add_distr_l. + exploit eval_mulimm_base; eauto. instantiate (1 := n). intros [v' [A1 B1]]. + exploit (eval_addimm (Int.mul n n2) le (mulimm_base n t2) v'). auto. intros [v'' [A2 B2]]. + exists v''; split; auto. eapply Val.lessdef_trans. eapply Val.add_lessdef; eauto. + rewrite Val.mul_commut; auto. + - apply eval_mulimm_base; auto. +Qed. + +Theorem eval_mul: binary_constructor_sound mul Val.mul. +Proof. + red; intros until y. + unfold mul; case (mul_match a b); intros; InvEval. + rewrite Val.mul_commut. apply eval_mulimm. auto. + apply eval_mulimm. auto. + TrivialExists. +Qed. + +Theorem eval_mulhs: binary_constructor_sound mulhs Val.mulhs. +Proof. + red; intros. unfold mulhs; destruct Archi.ptr64 eqn:SF. +- econstructor; split. + EvalOp. constructor. EvalOp. constructor. EvalOp. constructor. EvalOp. simpl; eauto. + constructor. EvalOp. simpl; eauto. constructor. + simpl; eauto. constructor. simpl; eauto. constructor. simpl; eauto. + destruct x; simpl; auto. destruct y; simpl; auto. + change (Int.ltu (Int.repr 32) Int64.iwordsize') with true; simpl. + apply Val.lessdef_same. f_equal. + transitivity (Int.repr (Z.shiftr (Int.signed i * Int.signed i0) 32)). + unfold Int.mulhs; f_equal. rewrite Zshiftr_div_two_p by lia. reflexivity. + apply Int.same_bits_eq; intros n N. + change Int.zwordsize with 32 in *. + assert (N1: 0 <= n < 64) by lia. + rewrite Int64.bits_loword by auto. + rewrite Int64.bits_shr' by auto. + change (Int.unsigned (Int.repr 32)) with 32. change Int64.zwordsize with 64. + rewrite zlt_true by lia. + rewrite Int.testbit_repr by auto. + unfold Int64.mul. rewrite Int64.testbit_repr by (change Int64.zwordsize with 64; lia). + transitivity (Z.testbit (Int.signed i * Int.signed i0) (n + 32)). + rewrite Z.shiftr_spec by lia. auto. + apply Int64.same_bits_eqm. apply Int64.eqm_mult; apply Int64.eqm_unsigned_repr. + change Int64.zwordsize with 64; lia. +- TrivialExists. +Qed. + +Theorem eval_mulhu: binary_constructor_sound mulhu Val.mulhu. +Proof. + red; intros. unfold mulhu; destruct Archi.ptr64 eqn:SF. +- econstructor; split. + EvalOp. constructor. EvalOp. constructor. EvalOp. constructor. EvalOp. simpl; eauto. + constructor. EvalOp. simpl; eauto. constructor. + simpl; eauto. constructor. simpl; eauto. constructor. simpl; eauto. + destruct x; simpl; auto. destruct y; simpl; auto. + change (Int.ltu (Int.repr 32) Int64.iwordsize') with true; simpl. + apply Val.lessdef_same. f_equal. + transitivity (Int.repr (Z.shiftr (Int.unsigned i * Int.unsigned i0) 32)). + unfold Int.mulhu; f_equal. rewrite Zshiftr_div_two_p by lia. reflexivity. + apply Int.same_bits_eq; intros n N. + change Int.zwordsize with 32 in *. + assert (N1: 0 <= n < 64) by lia. + rewrite Int64.bits_loword by auto. + rewrite Int64.bits_shru' by auto. + change (Int.unsigned (Int.repr 32)) with 32. change Int64.zwordsize with 64. + rewrite zlt_true by lia. + rewrite Int.testbit_repr by auto. + unfold Int64.mul. rewrite Int64.testbit_repr by (change Int64.zwordsize with 64; lia). + transitivity (Z.testbit (Int.unsigned i * Int.unsigned i0) (n + 32)). + rewrite Z.shiftr_spec by lia. auto. + apply Int64.same_bits_eqm. apply Int64.eqm_mult; apply Int64.eqm_unsigned_repr. + change Int64.zwordsize with 64; lia. +- TrivialExists. +Qed. + +Theorem eval_andimm: + forall n, unary_constructor_sound (andimm n) (fun x => Val.and x (Vint n)). +Proof. + intros; red; intros until x. unfold andimm. + + predSpec Int.eq Int.eq_spec n Int.zero. + intros. exists (Vint Int.zero); split. EvalOp. + destruct x; simpl; auto. subst n. rewrite Int.and_zero. auto. + + predSpec Int.eq Int.eq_spec n Int.mone. + intros. exists x; split; auto. + subst. destruct x; simpl; auto. rewrite Int.and_mone; auto. + + case (andimm_match a); intros. + - InvEval. TrivialExists. simpl. rewrite Int.and_commut; auto. + - InvEval. subst. rewrite Val.and_assoc. simpl. rewrite Int.and_commut. TrivialExists. + - TrivialExists. +Qed. + +Theorem eval_and: binary_constructor_sound and Val.and. +Proof. + red; intros until y; unfold and; case (and_match a b); intros; InvEval. + - rewrite Val.and_commut. apply eval_andimm; auto. + - apply eval_andimm; auto. + - TrivialExists. +Qed. + +Theorem eval_orimm: + forall n, unary_constructor_sound (orimm n) (fun x => Val.or x (Vint n)). +Proof. + intros; red; intros until x. unfold orimm. + + predSpec Int.eq Int.eq_spec n Int.zero. + intros. subst. exists x; split; auto. + destruct x; simpl; auto. rewrite Int.or_zero; auto. + + predSpec Int.eq Int.eq_spec n Int.mone. + intros. exists (Vint Int.mone); split. EvalOp. + destruct x; simpl; auto. subst n. rewrite Int.or_mone. auto. + + destruct (orimm_match a); intros; InvEval. + - TrivialExists. simpl. rewrite Int.or_commut; auto. + - subst. rewrite Val.or_assoc. simpl. rewrite Int.or_commut. TrivialExists. + - TrivialExists. +Qed. + +Theorem eval_or: binary_constructor_sound or Val.or. +Proof. + red; intros until y; unfold or; case (or_match a b); intros; InvEval. + - rewrite Val.or_commut. apply eval_orimm; auto. + - apply eval_orimm; auto. + - TrivialExists. +Qed. + +Theorem eval_xorimm: + forall n, unary_constructor_sound (xorimm n) (fun x => Val.xor x (Vint n)). +Proof. + intros; red; intros until x. unfold xorimm. + + predSpec Int.eq Int.eq_spec n Int.zero. + intros. exists x; split. auto. + destruct x; simpl; auto. subst n. rewrite Int.xor_zero. auto. + + intros. destruct (xorimm_match a); intros; InvEval. + - TrivialExists. simpl. rewrite Int.xor_commut; auto. + - subst. rewrite Val.xor_assoc. simpl. rewrite Int.xor_commut. + predSpec Int.eq Int.eq_spec (Int.xor n2 n) Int.zero. + + exists v1; split; auto. destruct v1; simpl; auto. rewrite H0, Int.xor_zero; auto. + + TrivialExists. + - TrivialExists. +Qed. + +Theorem eval_xor: binary_constructor_sound xor Val.xor. +Proof. + red; intros until y; unfold xor; case (xor_match a b); intros; InvEval. + - rewrite Val.xor_commut. apply eval_xorimm; auto. + - apply eval_xorimm; auto. + - TrivialExists. +Qed. + +Theorem eval_notint: unary_constructor_sound notint Val.notint. +Proof. + unfold notint; red; intros. rewrite Val.not_xor. apply eval_xorimm; auto. +Qed. + +Theorem eval_divs_base: + forall le a b x y z, + eval_expr ge sp e m le a x -> + eval_expr ge sp e m le b y -> + Val.divs x y = Some z -> + exists v, eval_expr ge sp e m le (divs_base a b) v /\ Val.lessdef z v. +Proof. + intros. unfold divs_base. exists z; split. EvalOp. + 2: apply Val.lessdef_refl. + cbn. + rewrite H1. + cbn. + trivial. +Qed. + +Theorem eval_mods_base: + forall le a b x y z, + eval_expr ge sp e m le a x -> + eval_expr ge sp e m le b y -> + Val.mods x y = Some z -> + exists v, eval_expr ge sp e m le (mods_base a b) v /\ Val.lessdef z v. +Proof. + intros. unfold mods_base. exists z; split. EvalOp. + 2: apply Val.lessdef_refl. + cbn. + rewrite H1. + cbn. + trivial. +Qed. + +Theorem eval_divu_base: + forall le a b x y z, + eval_expr ge sp e m le a x -> + eval_expr ge sp e m le b y -> + Val.divu x y = Some z -> + exists v, eval_expr ge sp e m le (divu_base a b) v /\ Val.lessdef z v. +Proof. + intros. unfold divu_base. exists z; split. EvalOp. + 2: apply Val.lessdef_refl. + cbn. + rewrite H1. + cbn. + trivial. +Qed. + +Theorem eval_modu_base: + forall le a b x y z, + eval_expr ge sp e m le a x -> + eval_expr ge sp e m le b y -> + Val.modu x y = Some z -> + exists v, eval_expr ge sp e m le (modu_base a b) v /\ Val.lessdef z v. +Proof. + intros. unfold modu_base. exists z; split. EvalOp. + 2: apply Val.lessdef_refl. + cbn. + rewrite H1. + cbn. + trivial. +Qed. + +Theorem eval_shrximm: + forall le a n x z, + eval_expr ge sp e m le a x -> + Val.shrx x (Vint n) = Some z -> + exists v, eval_expr ge sp e m le (shrximm a n) v /\ Val.lessdef z v. +Proof. + intros. unfold shrximm. + predSpec Int.eq Int.eq_spec n Int.zero. + subst n. exists x; split; auto. + destruct x; simpl in H0; try discriminate. + destruct (Int.ltu Int.zero (Int.repr 31)); inv H0. + replace (Int.shrx i Int.zero) with i. auto. + unfold Int.shrx, Int.divs. rewrite Int.shl_zero. + change (Int.signed Int.one) with 1. rewrite Z.quot_1_r. rewrite Int.repr_signed; auto. + econstructor; split. EvalOp. + cbn. + rewrite H0. + cbn. + reflexivity. + apply Val.lessdef_refl. +Qed. + +Theorem eval_shl: binary_constructor_sound shl Val.shl. +Proof. + red; intros until y; unfold shl; case (shl_match b); intros. + InvEval. apply eval_shlimm; auto. + TrivialExists. +Qed. + +Theorem eval_shr: binary_constructor_sound shr Val.shr. +Proof. + red; intros until y; unfold shr; case (shr_match b); intros. + InvEval. apply eval_shrimm; auto. + TrivialExists. +Qed. + +Theorem eval_shru: binary_constructor_sound shru Val.shru. +Proof. + red; intros until y; unfold shru; case (shru_match b); intros. + InvEval. apply eval_shruimm; auto. + TrivialExists. +Qed. + +Theorem eval_negf: unary_constructor_sound negf Val.negf. +Proof. + red; intros. TrivialExists. +Qed. + +Theorem eval_absf: unary_constructor_sound absf Val.absf. +Proof. + red; intros. TrivialExists. +Qed. + +Theorem eval_addf: binary_constructor_sound addf Val.addf. +Proof. + red; intros; TrivialExists. +Qed. + +Theorem eval_subf: binary_constructor_sound subf Val.subf. +Proof. + red; intros; TrivialExists. +Qed. + +Theorem eval_mulf: binary_constructor_sound mulf Val.mulf. +Proof. + red; intros; TrivialExists. +Qed. + +Theorem eval_negfs: unary_constructor_sound negfs Val.negfs. +Proof. + red; intros. TrivialExists. +Qed. + +Theorem eval_absfs: unary_constructor_sound absfs Val.absfs. +Proof. + red; intros. TrivialExists. +Qed. + +Theorem eval_addfs: binary_constructor_sound addfs Val.addfs. +Proof. + red; intros; TrivialExists. +Qed. + +Theorem eval_subfs: binary_constructor_sound subfs Val.subfs. +Proof. + red; intros; TrivialExists. +Qed. + +Theorem eval_mulfs: binary_constructor_sound mulfs Val.mulfs. +Proof. + red; intros; TrivialExists. +Qed. + +Section COMP_IMM. + +Variable default: comparison -> int -> condition. +Variable intsem: comparison -> int -> int -> bool. +Variable sem: comparison -> val -> val -> val. + +Hypothesis sem_int: forall c x y, sem c (Vint x) (Vint y) = Val.of_bool (intsem c x y). +Hypothesis sem_undef: forall c v, sem c Vundef v = Vundef. +Hypothesis sem_eq: forall x y, sem Ceq (Vint x) (Vint y) = Val.of_bool (Int.eq x y). +Hypothesis sem_ne: forall x y, sem Cne (Vint x) (Vint y) = Val.of_bool (negb (Int.eq x y)). +Hypothesis sem_default: forall c v n, sem c v (Vint n) = Val.of_optbool (eval_condition (default c n) (v :: nil) m). + +Lemma eval_compimm: + forall le c a n2 x, + eval_expr ge sp e m le a x -> + exists v, eval_expr ge sp e m le (compimm default intsem c a n2) v + /\ Val.lessdef (sem c x (Vint n2)) v. +Proof. + intros until x. + unfold compimm; case (compimm_match c a); intros. +(* constant *) + - InvEval. rewrite sem_int. TrivialExists. simpl. destruct (intsem c0 n1 n2); auto. +(* eq cmp *) + - InvEval. inv H. simpl in H5. inv H5. + destruct (Int.eq_dec n2 Int.zero). + + subst n2. TrivialExists. + simpl. rewrite eval_negate_condition. + destruct (eval_condition c0 vl m); simpl. + unfold Vtrue, Vfalse. destruct b; simpl; rewrite sem_eq; auto. + rewrite sem_undef; auto. + + destruct (Int.eq_dec n2 Int.one). subst n2. TrivialExists. + simpl. destruct (eval_condition c0 vl m); simpl. + unfold Vtrue, Vfalse. destruct b; simpl; rewrite sem_eq; auto. + rewrite sem_undef; auto. + exists (Vint Int.zero); split. EvalOp. + destruct (eval_condition c0 vl m); simpl. + unfold Vtrue, Vfalse. destruct b; rewrite sem_eq; rewrite Int.eq_false; auto. + rewrite sem_undef; auto. +(* ne cmp *) + - InvEval. inv H. simpl in H5. inv H5. + destruct (Int.eq_dec n2 Int.zero). + + subst n2. TrivialExists. + simpl. destruct (eval_condition c0 vl m); simpl. + unfold Vtrue, Vfalse. destruct b; simpl; rewrite sem_ne; auto. + rewrite sem_undef; auto. + + destruct (Int.eq_dec n2 Int.one). subst n2. TrivialExists. + simpl. rewrite eval_negate_condition. destruct (eval_condition c0 vl m); simpl. + unfold Vtrue, Vfalse. destruct b; simpl; rewrite sem_ne; auto. + rewrite sem_undef; auto. + exists (Vint Int.one); split. EvalOp. + destruct (eval_condition c0 vl m); simpl. + unfold Vtrue, Vfalse. destruct b; rewrite sem_ne; rewrite Int.eq_false; auto. + rewrite sem_undef; auto. +(* default *) + - TrivialExists. simpl. rewrite sem_default. auto. +Qed. + +Hypothesis sem_swap: + forall c x y, sem (swap_comparison c) x y = sem c y x. + +Lemma eval_compimm_swap: + forall le c a n2 x, + eval_expr ge sp e m le a x -> + exists v, eval_expr ge sp e m le (compimm default intsem (swap_comparison c) a n2) v + /\ Val.lessdef (sem c (Vint n2) x) v. +Proof. + intros. rewrite <- sem_swap. eapply eval_compimm; eauto. +Qed. + +End COMP_IMM. + +Theorem eval_comp: + forall c, binary_constructor_sound (comp c) (Val.cmp c). +Proof. + intros; red; intros until y. unfold comp; case (comp_match a b); intros; InvEval. + eapply eval_compimm_swap; eauto. + intros. unfold Val.cmp. rewrite Val.swap_cmp_bool; auto. + eapply eval_compimm; eauto. + TrivialExists. +Qed. + +Theorem eval_compu: + forall c, binary_constructor_sound (compu c) (Val.cmpu (Mem.valid_pointer m) c). +Proof. + intros; red; intros until y. unfold compu; case (compu_match a b); intros; InvEval. + eapply eval_compimm_swap; eauto. + intros. unfold Val.cmpu. rewrite Val.swap_cmpu_bool; auto. + eapply eval_compimm; eauto. + TrivialExists. +Qed. + +Theorem eval_compf: + forall c, binary_constructor_sound (compf c) (Val.cmpf c). +Proof. + intros; red; intros. unfold compf. TrivialExists. +Qed. + +Theorem eval_compfs: + forall c, binary_constructor_sound (compfs c) (Val.cmpfs c). +Proof. + intros; red; intros. unfold compfs. TrivialExists. +Qed. + +Theorem eval_cast8signed: unary_constructor_sound cast8signed (Val.sign_ext 8). +Proof. + red; intros until x. unfold cast8signed. case (cast8signed_match a); intros; InvEval. + TrivialExists. + TrivialExists. +Qed. + +Theorem eval_cast8unsigned: unary_constructor_sound cast8unsigned (Val.zero_ext 8). +Proof. + red; intros until x. unfold cast8unsigned. + rewrite Val.zero_ext_and. apply eval_andimm. lia. +Qed. + +Theorem eval_cast16signed: unary_constructor_sound cast16signed (Val.sign_ext 16). +Proof. + red; intros until x. unfold cast16signed. case (cast16signed_match a); intros; InvEval. + TrivialExists. + TrivialExists. +Qed. + +Theorem eval_cast16unsigned: unary_constructor_sound cast16unsigned (Val.zero_ext 16). +Proof. + red; intros until x. unfold cast8unsigned. + rewrite Val.zero_ext_and. apply eval_andimm. lia. +Qed. + +Theorem eval_intoffloat: + forall le a x y, + eval_expr ge sp e m le a x -> + Val.intoffloat x = Some y -> + exists v, eval_expr ge sp e m le (intoffloat a) v /\ Val.lessdef y v. +Proof. + intros; unfold intoffloat. TrivialExists. + cbn. rewrite H0. reflexivity. +Qed. + +Theorem eval_intuoffloat: + forall le a x y, + eval_expr ge sp e m le a x -> + Val.intuoffloat x = Some y -> + exists v, eval_expr ge sp e m le (intuoffloat a) v /\ Val.lessdef y v. +Proof. + intros; unfold intuoffloat. TrivialExists. + cbn. rewrite H0. reflexivity. +Qed. + +Theorem eval_floatofintu: + forall le a x y, + eval_expr ge sp e m le a x -> + Val.floatofintu x = Some y -> + exists v, eval_expr ge sp e m le (floatofintu a) v /\ Val.lessdef y v. +Proof. + intros until y; unfold floatofintu. case (floatofintu_match a); intros. + InvEval. simpl in H0. TrivialExists. + TrivialExists. + cbn. rewrite H0. reflexivity. +Qed. + +Theorem eval_floatofint: + forall le a x y, + eval_expr ge sp e m le a x -> + Val.floatofint x = Some y -> + exists v, eval_expr ge sp e m le (floatofint a) v /\ Val.lessdef y v. +Proof. + intros until y; unfold floatofint. case (floatofint_match a); intros. + InvEval. simpl in H0. TrivialExists. + TrivialExists. + cbn. rewrite H0. reflexivity. +Qed. + +Theorem eval_intofsingle: + forall le a x y, + eval_expr ge sp e m le a x -> + Val.intofsingle x = Some y -> + exists v, eval_expr ge sp e m le (intofsingle a) v /\ Val.lessdef y v. +Proof. + intros; unfold intofsingle. TrivialExists. + cbn. rewrite H0. reflexivity. +Qed. + +Theorem eval_singleofint: + forall le a x y, + eval_expr ge sp e m le a x -> + Val.singleofint x = Some y -> + exists v, eval_expr ge sp e m le (singleofint a) v /\ Val.lessdef y v. +Proof. + intros; unfold singleofint; TrivialExists. + cbn. rewrite H0. reflexivity. +Qed. + +Theorem eval_intuofsingle: + forall le a x y, + eval_expr ge sp e m le a x -> + Val.intuofsingle x = Some y -> + exists v, eval_expr ge sp e m le (intuofsingle a) v /\ Val.lessdef y v. +Proof. + intros; unfold intuofsingle. TrivialExists. + cbn. rewrite H0. reflexivity. +Qed. + +Theorem eval_singleofintu: + forall le a x y, + eval_expr ge sp e m le a x -> + Val.singleofintu x = Some y -> + exists v, eval_expr ge sp e m le (singleofintu a) v /\ Val.lessdef y v. +Proof. + intros; unfold intuofsingle. TrivialExists. + cbn. rewrite H0. reflexivity. +Qed. + +Theorem eval_singleoffloat: unary_constructor_sound singleoffloat Val.singleoffloat. +Proof. + red; intros. unfold singleoffloat. TrivialExists. +Qed. + +Theorem eval_floatofsingle: unary_constructor_sound floatofsingle Val.floatofsingle. +Proof. + red; intros. unfold floatofsingle. TrivialExists. +Qed. + +Lemma mod_small_negative: + forall a modulus, + modulus > 0 -> -modulus < a < 0 -> a mod modulus = a + modulus. +Proof. + intros. + replace (a mod modulus) with ((a + modulus) mod modulus). + apply Z.mod_small. + lia. + rewrite <- Zplus_mod_idemp_r. + rewrite Z.mod_same by lia. + rewrite Z.add_0_r. + reflexivity. +Qed. + +Remark normalize_low_long: forall + (PTR64 : Archi.ptr64 = true) v1, + Val.loword (Val.normalize (Val.longofint v1) Tlong) = Val.normalize v1 Tint. +Proof. + intros. + destruct v1; cbn; try rewrite PTR64; trivial. + f_equal. + unfold Int64.loword. + unfold Int.signed. + destruct zlt. + { rewrite Int64.int_unsigned_repr. + apply Int.repr_unsigned. + } + pose proof (Int.unsigned_range i). + rewrite Int64.unsigned_repr_eq. + replace ((Int.unsigned i - Int.modulus) mod Int64.modulus) + with (Int64.modulus + Int.unsigned i - Int.modulus). + { + rewrite <- (Int.repr_unsigned i) at 2. + apply Int.eqm_samerepr. + unfold Int.eqm, eqmod. + change Int.modulus with 4294967296 in *. + change Int64.modulus with 18446744073709551616 in *. + exists 4294967295. + lia. + } + { rewrite mod_small_negative. + lia. + constructor. + constructor. + change Int.modulus with 4294967296 in *. + change Int.half_modulus with 2147483648 in *. + change Int64.modulus with 18446744073709551616 in *. + lia. + lia. + } +Qed. + +Lemma same_expr_pure_correct: + forall le a1 a2 v1 v2 + (PURE : same_expr_pure a1 a2 = true) + (EVAL1 : eval_expr ge sp e m le a1 v1) + (EVAL2 : eval_expr ge sp e m le a2 v2), + v1 = v2. +Proof. + intros. + destruct a1; destruct a2; cbn in *; try discriminate. + inv EVAL1. inv EVAL2. + destruct (ident_eq i i0); congruence. +Qed. + +Theorem eval_select: + forall le ty cond al vl a1 v1 a2 v2 a b, + select ty cond al a1 a2 = Some a -> + eval_exprlist ge sp e m le al vl -> + eval_expr ge sp e m le a1 v1 -> + eval_expr ge sp e m le a2 v2 -> + eval_condition cond vl m = Some b -> + exists v, + eval_expr ge sp e m le a v + /\ Val.lessdef (Val.select (Some b) v1 v2 ty) v. +Proof. + unfold select; intros. + pose proof (same_expr_pure_correct le a1 a2 v1 v2) as PURE. + destruct (same_expr_pure a1 a2). + { rewrite <- PURE by auto. + inv H. + exists v1. split. assumption. + unfold Val.select. + destruct b; apply Val.lessdef_normalize. + } + clear PURE. + destruct Archi.ptr64 eqn:PTR64. + 2: discriminate. + destruct ty; cbn in *; try discriminate. + - (* Tint *) + inv H. TrivialExists. + + cbn. repeat econstructor; eassumption. + + cbn. f_equal. rewrite ExtValues.normalize_select01. + rewrite H3. destruct b. + * rewrite ExtValues.select01_long_true. apply normalize_low_long; assumption. + * rewrite ExtValues.select01_long_false. apply normalize_low_long; assumption. + + - (* Tfloat *) + inv H. TrivialExists. + + cbn. repeat econstructor; eassumption. + + cbn. f_equal. rewrite ExtValues.normalize_select01. + rewrite H3. destruct b. + * rewrite ExtValues.select01_long_true. + apply ExtValues.float_bits_normalize. + * rewrite ExtValues.select01_long_false. + apply ExtValues.float_bits_normalize. + + - (* Tlong *) + inv H. TrivialExists. + + cbn. repeat econstructor; eassumption. + + cbn. f_equal. rewrite ExtValues.normalize_select01. + rewrite H3. destruct b. + * rewrite ExtValues.select01_long_true. reflexivity. + * rewrite ExtValues.select01_long_false. reflexivity. + + - (* Tsingle *) + inv H. TrivialExists. + + cbn. repeat econstructor; eassumption. + + cbn. f_equal. rewrite ExtValues.normalize_select01. + rewrite H3. destruct b. + * rewrite ExtValues.select01_long_true. + rewrite normalize_low_long by assumption. + apply ExtValues.single_bits_normalize. + * rewrite ExtValues.select01_long_false. + rewrite normalize_low_long by assumption. + apply ExtValues.single_bits_normalize. +Qed. + +Theorem eval_addressing: + forall le chunk a v b ofs, + eval_expr ge sp e m le a v -> + v = Vptr b ofs -> + match addressing chunk a with (mode, args) => + exists vl, + eval_exprlist ge sp e m le args vl /\ + eval_addressing ge sp mode vl = Some v + end. +Proof. + intros until v. unfold addressing; case (addressing_match a); intros; InvEval. + - exists (@nil val); split. eauto with evalexpr. simpl. auto. + - destruct (Archi.pic_code tt). + + exists (Vptr b ofs0 :: nil); split. + constructor. EvalOp. simpl. congruence. constructor. simpl. rewrite Ptrofs.add_zero. congruence. + + exists (@nil val); split. constructor. simpl; auto. + - exists (v1 :: nil); split. eauto with evalexpr. simpl. + destruct v1; simpl in H; try discriminate. destruct Archi.ptr64 eqn:SF; inv H. + simpl. auto. + - exists (v1 :: nil); split. eauto with evalexpr. simpl. + destruct v1; simpl in H; try discriminate. destruct Archi.ptr64 eqn:SF; inv H. + simpl. auto. + - exists (v :: nil); split. eauto with evalexpr. subst. simpl. rewrite Ptrofs.add_zero; auto. +Qed. + +Theorem eval_builtin_arg: + forall a v, + eval_expr ge sp e m nil a v -> + CminorSel.eval_builtin_arg ge sp e m (builtin_arg a) v. +Proof. + intros until v. unfold builtin_arg; case (builtin_arg_match a); intros. +- InvEval. constructor. +- InvEval. constructor. +- InvEval. constructor. +- InvEval. simpl in H5. inv H5. constructor. +- InvEval. subst v. constructor; auto. +- inv H. InvEval. simpl in H6; inv H6. constructor; auto. +- destruct Archi.ptr64 eqn:SF. ++ constructor; auto. ++ InvEval. replace v with (if Archi.ptr64 then Val.addl v1 (Vint n) else Val.add v1 (Vint n)). + repeat constructor; auto. + rewrite SF; auto. +- destruct Archi.ptr64 eqn:SF. ++ InvEval. replace v with (if Archi.ptr64 then Val.addl v1 (Vlong n) else Val.add v1 (Vlong n)). + repeat constructor; auto. + rewrite SF; auto. ++ constructor; auto. +- constructor; auto. +Qed. + +(* floating-point division without HELPERS *) +Theorem eval_divf_base: + forall le a b x y, + eval_expr ge sp e m le a x -> + eval_expr ge sp e m le b y -> + exists v, eval_expr ge sp e m le (divf_base a b) v /\ Val.lessdef (Val.divf x y) v. +Proof. + intros; unfold divf_base. + TrivialExists. +Qed. + +Theorem eval_divfs_base: + forall le a b x y, + eval_expr ge sp e m le a x -> + eval_expr ge sp e m le b y -> + exists v, eval_expr ge sp e m le (divfs_base a b) v /\ Val.lessdef (Val.divfs x y) v. +Proof. + intros; unfold divfs_base. + TrivialExists. +Qed. + +(** Platform-specific known builtins *) + +Theorem eval_platform_builtin: + forall bf al a vl v le, + platform_builtin bf al = Some a -> + eval_exprlist ge sp e m le al vl -> + platform_builtin_sem bf vl = Some v -> + exists v', eval_expr ge sp e m le a v' /\ Val.lessdef v v'. +Proof. + destruct bf; intros until le; intro Heval. + all: try (inversion Heval; subst a; clear Heval; + exists v; split; trivial; + repeat (try econstructor; try eassumption)). +Qed. + +End CMCONSTR. diff --git a/riscV/TO_MERGE/Asm.v b/riscV/TO_MERGE/Asm.v deleted file mode 100644 index f75825a1..00000000 --- a/riscV/TO_MERGE/Asm.v +++ /dev/null @@ -1,1214 +0,0 @@ -(* *********************************************************************) -(* *) -(* The Compcert verified compiler *) -(* *) -(* Xavier Leroy, INRIA Paris-Rocquencourt *) -(* Prashanth Mundkur, SRI International *) -(* *) -(* 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. *) -(* *) -(* The contributions by Prashanth Mundkur are reused and adapted *) -(* under the terms of a Contributor License Agreement between *) -(* SRI International and INRIA. *) -(* *) -(* *********************************************************************) - -(** Abstract syntax and semantics for RISC-V assembly language. *) - -Require Import Coqlib. -Require Import Maps. -Require Import AST. -Require Import Integers. -Require Import Floats. -Require Import Values. -Require Import Memory. -Require Import Events. -Require Import Globalenvs. -Require Import Smallstep. -Require Import Locations. -Require Stacklayout. -Require Import Conventions. -Require ExtValues. - -(** * Abstract syntax *) - -(** Integer registers. X0 is treated specially because it always reads - as zero and is never used as a destination of an instruction. *) - -Inductive ireg: Type := - | X1: ireg | X2: ireg | X3: ireg | X4: ireg | X5: ireg - | X6: ireg | X7: ireg | X8: ireg | X9: ireg | X10: ireg - | X11: ireg | X12: ireg | X13: ireg | X14: ireg | X15: ireg - | X16: ireg | X17: ireg | X18: ireg | X19: ireg | X20: ireg - | X21: ireg | X22: ireg | X23: ireg | X24: ireg | X25: ireg - | X26: ireg | X27: ireg | X28: ireg | X29: ireg | X30: ireg - | X31: ireg. - -Inductive ireg0: Type := - | X0: ireg0 | X: ireg -> ireg0. - -Coercion X: ireg >-> ireg0. - -(** Floating-point registers *) - -Inductive freg: Type := - | F0: freg | F1: freg | F2: freg | F3: freg - | F4: freg | F5: freg | F6: freg | F7: freg - | F8: freg | F9: freg | F10: freg | F11: freg - | F12: freg | F13: freg | F14: freg | F15: freg - | F16: freg | F17: freg | F18: freg | F19: freg - | F20: freg | F21: freg | F22: freg | F23: freg - | F24: freg | F25: freg | F26: freg | F27: freg - | F28: freg | F29: freg | F30: freg | F31: freg. - -Definition ireg_eq: forall (x y: ireg), {x=y} + {x<>y}. -Proof. decide equality. Defined. - -Definition ireg0_eq: forall (x y: ireg0), {x=y} + {x<>y}. -Proof. decide equality. apply ireg_eq. Defined. - -Lemma freg_eq: forall (x y: freg), {x=y} + {x<>y}. -Proof. decide equality. Defined. - -(** We model the following registers of the RISC-V architecture. *) - -Inductive preg: Type := - | IR: ireg -> preg (**r integer registers *) - | FR: freg -> preg (**r double-precision float registers *) - | PC: preg. (**r program counter *) - -Coercion IR: ireg >-> preg. -Coercion FR: freg >-> preg. - -Lemma preg_eq: forall (x y: preg), {x=y} + {x<>y}. -Proof. decide equality. apply ireg_eq. apply freg_eq. Defined. - -Module PregEq. - Definition t := preg. - Definition eq := preg_eq. -End PregEq. - -Module Pregmap := EMap(PregEq). - -(** Conventional names for stack pointer ([SP]) and return address ([RA]). *) - -Notation "'SP'" := X2 (only parsing) : asm. -Notation "'RA'" := X1 (only parsing) : asm. - -(** Offsets for load and store instructions. An offset is either an - immediate integer or the low part of a symbol. *) - -Inductive offset : Type := - | Ofsimm (ofs: ptrofs) - | Ofslow (id: ident) (ofs: ptrofs). - -(** The RISC-V instruction set is composed of several subsets. We model - the "32I" (32-bit integers), "64I" (64-bit integers), - "M" (multiplication and division), - "F" (single-precision floating-point) - and "D" (double-precision floating-point) subsets. - - For 32- and 64-bit integer arithmetic, the RISC-V instruction set comprises - generic integer operations such as ADD that operate over the full width - of an integer register (either 32 or 64 bit), plus specific instructions - such as ADDW that normalize their results to signed 32-bit integers. - Other instructions such as AND work equally well over 32- and 64-bit - integers, with the convention that 32-bit integers are represented - sign-extended in 64-bit registers. - - This clever design is challenging to formalize in the CompCert value - model. As a first step, we follow a more traditional approach, - also used in the x86 port, whereas we have two sets of (pseudo-) - instructions, one for 32-bit integer arithmetic, with suffix W, - the other for 64-bit integer arithmetic, with suffix L. The mapping - to actual instructions is done when printing assembly code, as follows: - - In 32-bit mode: - ADDW becomes ADD, ADDL is an error, ANDW becomes AND, ANDL is an error. - - In 64-bit mode: - ADDW becomes ADDW, ADDL becomes ADD, ANDW and ANDL both become AND. -*) - -Definition label := positive. - -(** A note on immediates: there are various constraints on immediate - operands to RISC-V instructions. We do not attempt to capture these - restrictions in the abstract syntax nor in the semantics. The - assembler will emit an error if immediate operands exceed the - representable range. Of course, our RISC-V generator (file - [Asmgen]) is careful to respect this range. *) - -Inductive instruction : Type := - | Pmv (rd: ireg) (rs: ireg) (**r integer move *) - -(** 32-bit integer register-immediate instructions *) - | Paddiw (rd: ireg) (rs: ireg0) (imm: int) (**r add immediate *) - | Psltiw (rd: ireg) (rs: ireg0) (imm: int) (**r set-less-than immediate *) - | Psltiuw (rd: ireg) (rs: ireg0) (imm: int) (**r set-less-than unsigned immediate *) - | Pandiw (rd: ireg) (rs: ireg0) (imm: int) (**r and immediate *) - | Poriw (rd: ireg) (rs: ireg0) (imm: int) (**r or immediate *) - | Pxoriw (rd: ireg) (rs: ireg0) (imm: int) (**r xor immediate *) - | Pslliw (rd: ireg) (rs: ireg0) (imm: int) (**r shift-left-logical immediate *) - | Psrliw (rd: ireg) (rs: ireg0) (imm: int) (**r shift-right-logical immediate *) - | Psraiw (rd: ireg) (rs: ireg0) (imm: int) (**r shift-right-arith immediate *) - | Pluiw (rd: ireg) (imm: int) (**r load upper-immediate *) -(** 32-bit integer register-register instructions *) - | Paddw (rd: ireg) (rs1 rs2: ireg0) (**r integer addition *) - | Psubw (rd: ireg) (rs1 rs2: ireg0) (**r integer subtraction *) - - | Pmulw (rd: ireg) (rs1 rs2: ireg0) (**r integer multiply low *) - | Pmulhw (rd: ireg) (rs1 rs2: ireg0) (**r integer multiply high signed *) - | Pmulhuw (rd: ireg) (rs1 rs2: ireg0) (**r integer multiply high unsigned *) - | Pdivw (rd: ireg) (rs1 rs2: ireg0) (**r integer division *) - | Pdivuw (rd: ireg) (rs1 rs2: ireg0) (**r unsigned integer division *) - | Premw (rd: ireg) (rs1 rs2: ireg0) (**r integer remainder *) - | Premuw (rd: ireg) (rs1 rs2: ireg0) (**r unsigned integer remainder *) - | Psltw (rd: ireg) (rs1 rs2: ireg0) (**r set-less-than *) - | Psltuw (rd: ireg) (rs1 rs2: ireg0) (**r set-less-than unsigned *) - | Pseqw (rd: ireg) (rs1 rs2: ireg0) (**r [rd <- rs1 == rs2] (pseudo) *) - | Psnew (rd: ireg) (rs1 rs2: ireg0) (**r [rd <- rs1 != rs2] (pseudo) *) - | Pandw (rd: ireg) (rs1 rs2: ireg0) (**r bitwise and *) - | Porw (rd: ireg) (rs1 rs2: ireg0) (**r bitwise or *) - | Pxorw (rd: ireg) (rs1 rs2: ireg0) (**r bitwise xor *) - | Psllw (rd: ireg) (rs1 rs2: ireg0) (**r shift-left-logical *) - | Psrlw (rd: ireg) (rs1 rs2: ireg0) (**r shift-right-logical *) - | Psraw (rd: ireg) (rs1 rs2: ireg0) (**r shift-right-arith *) - -(** 64-bit integer register-immediate instructions *) - | Paddil (rd: ireg) (rs: ireg0) (imm: int64) (**r add immediate *) - | Psltil (rd: ireg) (rs: ireg0) (imm: int64) (**r set-less-than immediate *) - | Psltiul (rd: ireg) (rs: ireg0) (imm: int64) (**r set-less-than unsigned immediate *) - | Pandil (rd: ireg) (rs: ireg0) (imm: int64) (**r and immediate *) - | Poril (rd: ireg) (rs: ireg0) (imm: int64) (**r or immediate *) - | Pxoril (rd: ireg) (rs: ireg0) (imm: int64) (**r xor immediate *) - | Psllil (rd: ireg) (rs: ireg0) (imm: int) (**r shift-left-logical immediate *) - | Psrlil (rd: ireg) (rs: ireg0) (imm: int) (**r shift-right-logical immediate *) - | Psrail (rd: ireg) (rs: ireg0) (imm: int) (**r shift-right-arith immediate *) - | Pluil (rd: ireg) (imm: int64) (**r load upper-immediate *) -(** 64-bit integer register-register instructions *) - | Paddl (rd: ireg) (rs1 rs2: ireg0) (**r integer addition *) - | Psubl (rd: ireg) (rs1 rs2: ireg0) (**r integer subtraction *) - - | Pmull (rd: ireg) (rs1 rs2: ireg0) (**r integer multiply low *) - | Pmulhl (rd: ireg) (rs1 rs2: ireg0) (**r integer multiply high signed *) - | Pmulhul (rd: ireg) (rs1 rs2: ireg0) (**r integer multiply high unsigned *) - | Pdivl (rd: ireg) (rs1 rs2: ireg0) (**r integer division *) - | Pdivul (rd: ireg) (rs1 rs2: ireg0) (**r unsigned integer division *) - | Preml (rd: ireg) (rs1 rs2: ireg0) (**r integer remainder *) - | Premul (rd: ireg) (rs1 rs2: ireg0) (**r unsigned integer remainder *) - | Psltl (rd: ireg) (rs1 rs2: ireg0) (**r set-less-than *) - | Psltul (rd: ireg) (rs1 rs2: ireg0) (**r set-less-than unsigned *) - | Pseql (rd: ireg) (rs1 rs2: ireg0) (**r [rd <- rs1 == rs2] (pseudo) *) - | Psnel (rd: ireg) (rs1 rs2: ireg0) (**r [rd <- rs1 != rs2] (pseudo) *) - | Pandl (rd: ireg) (rs1 rs2: ireg0) (**r bitwise and *) - | Porl (rd: ireg) (rs1 rs2: ireg0) (**r bitwise or *) - | Pxorl (rd: ireg) (rs1 rs2: ireg0) (**r bitwise xor *) - | Pslll (rd: ireg) (rs1 rs2: ireg0) (**r shift-left-logical *) - | Psrll (rd: ireg) (rs1 rs2: ireg0) (**r shift-right-logical *) - | Psral (rd: ireg) (rs1 rs2: ireg0) (**r shift-right-arith *) - - | Pcvtl2w (rd: ireg) (rs: ireg0) (**r int64->int32 (pseudo) *) - | Pcvtw2l (r: ireg) (**r int32 signed -> int64 (pseudo) *) - - (* Unconditional jumps. Links are always to X1/RA. *) - | Pj_l (l: label) (**r jump to label *) - | Pj_s (symb: ident) (sg: signature) (**r jump to symbol *) - | Pj_r (r: ireg) (sg: signature) (**r jump register *) - | Pjal_s (symb: ident) (sg: signature) (**r jump-and-link symbol *) - | Pjal_r (r: ireg) (sg: signature) (**r jump-and-link register *) - - (* Conditional branches, 32-bit comparisons *) - | Pbeqw (rs1 rs2: ireg0) (l: label) (**r branch-if-equal *) - | Pbnew (rs1 rs2: ireg0) (l: label) (**r branch-if-not-equal signed *) - | Pbltw (rs1 rs2: ireg0) (l: label) (**r branch-if-less signed *) - | Pbltuw (rs1 rs2: ireg0) (l: label) (**r branch-if-less unsigned *) - | Pbgew (rs1 rs2: ireg0) (l: label) (**r branch-if-greater-or-equal signed *) - | Pbgeuw (rs1 rs2: ireg0) (l: label) (**r branch-if-greater-or-equal unsigned *) - - (* Conditional branches, 64-bit comparisons *) - | Pbeql (rs1 rs2: ireg0) (l: label) (**r branch-if-equal *) - | Pbnel (rs1 rs2: ireg0) (l: label) (**r branch-if-not-equal signed *) - | Pbltl (rs1 rs2: ireg0) (l: label) (**r branch-if-less signed *) - | Pbltul (rs1 rs2: ireg0) (l: label) (**r branch-if-less unsigned *) - | Pbgel (rs1 rs2: ireg0) (l: label) (**r branch-if-greater-or-equal signed *) - | Pbgeul (rs1 rs2: ireg0) (l: label) (**r branch-if-greater-or-equal unsigned *) - - (* Loads and stores *) - | Plb (rd: ireg) (ra: ireg) (ofs: offset) (**r load signed int8 *) - | Plbu (rd: ireg) (ra: ireg) (ofs: offset) (**r load unsigned int8 *) - | Plh (rd: ireg) (ra: ireg) (ofs: offset) (**r load signed int16 *) - | Plhu (rd: ireg) (ra: ireg) (ofs: offset) (**r load unsigned int16 *) - | Plw (rd: ireg) (ra: ireg) (ofs: offset) (**r load int32 *) - | Plw_a (rd: ireg) (ra: ireg) (ofs: offset) (**r load any32 *) - | Pld (rd: ireg) (ra: ireg) (ofs: offset) (**r load int64 *) - | Pld_a (rd: ireg) (ra: ireg) (ofs: offset) (**r load any64 *) - - | Psb (rs: ireg) (ra: ireg) (ofs: offset) (**r store int8 *) - | Psh (rs: ireg) (ra: ireg) (ofs: offset) (**r store int16 *) - | Psw (rs: ireg) (ra: ireg) (ofs: offset) (**r store int32 *) - | Psw_a (rs: ireg) (ra: ireg) (ofs: offset) (**r store any32 *) - | Psd (rs: ireg) (ra: ireg) (ofs: offset) (**r store int64 *) - | Psd_a (rs: ireg) (ra: ireg) (ofs: offset) (**r store any64 *) - - (* Synchronization *) - | Pfence (**r fence *) - - (* floating point register move *) - | Pfmv (rd: freg) (rs: freg) (**r move *) -<<<<<<< HEAD - | Pfmvxs (rd: ireg) (rs: freg) (**r bitwise move FP single to integer register *) - | Pfmvxd (rd: ireg) (rs: freg) (**r bitwise move FP double to integer register *) - | Pfmvsx (rd: freg) (rs: ireg) (**r bitwise move integer register to FP single *) - | Pfmvdx (rd: freg) (rs: ireg) (**r bitwise move integer register to FP double*) -======= - | Pfmvxs (rd: ireg) (rs: freg) (**r move FP single to integer register *) - | Pfmvsx (rd: freg) (rs: ireg) (**r move integer register to FP single *) - | Pfmvxd (rd: ireg) (rs: freg) (**r move FP double to integer register *) - | Pfmvdx (rd: freg) (rs: ireg) (**r move integer register to FP double *) ->>>>>>> master - - (* 32-bit (single-precision) floating point *) - | Pfls (rd: freg) (ra: ireg) (ofs: offset) (**r load float *) - | Pfss (rs: freg) (ra: ireg) (ofs: offset) (**r store float *) - - | Pfnegs (rd: freg) (rs: freg) (**r negation *) - | Pfabss (rd: freg) (rs: freg) (**r absolute value *) - - | Pfadds (rd: freg) (rs1 rs2: freg) (**r addition *) - | Pfsubs (rd: freg) (rs1 rs2: freg) (**r subtraction *) - | Pfmuls (rd: freg) (rs1 rs2: freg) (**r multiplication *) - | Pfdivs (rd: freg) (rs1 rs2: freg) (**r division *) - | Pfmins (rd: freg) (rs1 rs2: freg) (**r minimum *) - | Pfmaxs (rd: freg) (rs1 rs2: freg) (**r maximum *) - - | Pfeqs (rd: ireg) (rs1 rs2: freg) (**r compare equal *) - | Pflts (rd: ireg) (rs1 rs2: freg) (**r compare less-than *) - | Pfles (rd: ireg) (rs1 rs2: freg) (**r compare less-than/equal *) - - | Pfsqrts (rd: freg) (rs: freg) (**r square-root *) - - | Pfmadds (rd: freg) (rs1 rs2 rs3: freg) (**r fused multiply-add *) - | Pfmsubs (rd: freg) (rs1 rs2 rs3: freg) (**r fused multiply-sub *) - | Pfnmadds (rd: freg) (rs1 rs2 rs3: freg) (**r fused negated multiply-add *) - | Pfnmsubs (rd: freg) (rs1 rs2 rs3: freg) (**r fused negated multiply-sub *) - - | Pfcvtws (rd: ireg) (rs: freg) (**r float32 -> int32 conversion *) - | Pfcvtwus (rd: ireg) (rs: freg) (**r float32 -> unsigned int32 conversion *) - | Pfcvtsw (rd: freg) (rs: ireg0) (**r int32 -> float32 conversion *) - | Pfcvtswu (rd: freg) (rs: ireg0) (**r unsigned int32 -> float32 conversion *) - - | Pfcvtls (rd: ireg) (rs: freg) (**r float32 -> int64 conversion *) - | Pfcvtlus (rd: ireg) (rs: freg) (**r float32 -> unsigned int64 conversion *) - | Pfcvtsl (rd: freg) (rs: ireg0) (**r int64 -> float32 conversion *) - | Pfcvtslu (rd: freg) (rs: ireg0) (**r unsigned int 64-> float32 conversion *) - - (* 64-bit (double-precision) floating point *) - | Pfld (rd: freg) (ra: ireg) (ofs: offset) (**r load 64-bit float *) - | Pfld_a (rd: freg) (ra: ireg) (ofs: offset) (**r load any64 *) - | Pfsd (rd: freg) (ra: ireg) (ofs: offset) (**r store 64-bit float *) - | Pfsd_a (rd: freg) (ra: ireg) (ofs: offset) (**r store any64 *) - - | Pfnegd (rd: freg) (rs: freg) (**r negation *) - | Pfabsd (rd: freg) (rs: freg) (**r absolute value *) - - | Pfaddd (rd: freg) (rs1 rs2: freg) (**r addition *) - | Pfsubd (rd: freg) (rs1 rs2: freg) (**r subtraction *) - | Pfmuld (rd: freg) (rs1 rs2: freg) (**r multiplication *) - | Pfdivd (rd: freg) (rs1 rs2: freg) (**r division *) - | Pfmind (rd: freg) (rs1 rs2: freg) (**r minimum *) - | Pfmaxd (rd: freg) (rs1 rs2: freg) (**r maximum *) - - | Pfeqd (rd: ireg) (rs1 rs2: freg) (**r compare equal *) - | Pfltd (rd: ireg) (rs1 rs2: freg) (**r compare less-than *) - | Pfled (rd: ireg) (rs1 rs2: freg) (**r compare less-than/equal *) - - | Pfsqrtd (rd: freg) (rs: freg) (**r square-root *) - - | Pfmaddd (rd: freg) (rs1 rs2 rs3: freg) (**r fused multiply-add *) - | Pfmsubd (rd: freg) (rs1 rs2 rs3: freg) (**r fused multiply-sub *) - | Pfnmaddd (rd: freg) (rs1 rs2 rs3: freg) (**r fused negated multiply-add *) - | Pfnmsubd (rd: freg) (rs1 rs2 rs3: freg) (**r fused negated multiply-sub *) - - | Pfcvtwd (rd: ireg) (rs: freg) (**r float -> int32 conversion *) - | Pfcvtwud (rd: ireg) (rs: freg) (**r float -> unsigned int32 conversion *) - | Pfcvtdw (rd: freg) (rs: ireg0) (**r int32 -> float conversion *) - | Pfcvtdwu (rd: freg) (rs: ireg0) (**r unsigned int32 -> float conversion *) - - | Pfcvtld (rd: ireg) (rs: freg) (**r float -> int64 conversion *) - | Pfcvtlud (rd: ireg) (rs: freg) (**r float -> unsigned int64 conversion *) - | Pfcvtdl (rd: freg) (rs: ireg0) (**r int64 -> float conversion *) - | Pfcvtdlu (rd: freg) (rs: ireg0) (**r unsigned int64 -> float conversion *) - - | Pfcvtds (rd: freg) (rs: freg) (**r float32 -> float *) - | Pfcvtsd (rd: freg) (rs: freg) (**r float -> float32 *) - - (* Pseudo-instructions *) - | Pallocframe (sz: Z) (pos: ptrofs) (**r allocate new stack frame *) - | Pfreeframe (sz: Z) (pos: ptrofs) (**r deallocate stack frame and restore previous frame *) - | Plabel (lbl: label) (**r define a code label *) - | Ploadsymbol (rd: ireg) (id: ident) (ofs: ptrofs) (**r load the address of a symbol *) - | Ploadsymbol_high (rd: ireg) (id: ident) (ofs: ptrofs) (**r load the high part of the address of a symbol *) - | Ploadli (rd: ireg) (i: int64) (**r load an immediate int64 *) - | Ploadfi (rd: freg) (f: float) (**r load an immediate float *) - | Ploadsi (rd: freg) (f: float32) (**r load an immediate single *) - | Pbtbl (r: ireg) (tbl: list label) (**r N-way branch through a jump table *) - | Pbuiltin: external_function -> list (builtin_arg preg) - -> builtin_res preg -> instruction (**r built-in function (pseudo) *) - | Pselectl (rd: ireg) (rb: ireg0) (rt: ireg0) (rf: ireg0) - | Pnop : instruction. (**r nop instruction *) - - -(** The pseudo-instructions are the following: - -- [Plabel]: define a code label at the current program point. - -- [Ploadsymbol]: load the address of a symbol in an integer register. - Expands to the [la] assembler pseudo-instruction, which does the right - thing even if we are in PIC mode. - -- [Ploadli rd ival]: load an immediate 64-bit integer into an integer - register rd. Expands to a load from an address in the constant data section, - using the integer register x31 as temporary: -<< - lui x31, %hi(lbl) - ld rd, %lo(lbl)(x31) -lbl: - .quad ival ->> - -- [Ploadfi rd fval]: similar to [Ploadli] but loads a double FP constant fval - into a float register rd. - -- [Ploadsi rd fval]: similar to [Ploadli] but loads a single FP constant fval - into a float register rd. - -- [Pallocframe sz pos]: in the formal semantics, this - pseudo-instruction allocates a memory block with bounds [0] and - [sz], stores the value of the stack pointer at offset [pos] in this - block, and sets the stack pointer to the address of the bottom of - this block. - In the printed ASM assembly code, this allocation is: -<< - mv x30, sp - sub sp, sp, #sz - sw x30, #pos(sp) ->> - This cannot be expressed in our memory model, which does not reflect - the fact that stack frames are adjacent and allocated/freed - following a stack discipline. - -- [Pfreeframe sz pos]: in the formal semantics, this pseudo-instruction - reads the word at [pos] of the block pointed by the stack pointer, - frees this block, and sets the stack pointer to the value read. - In the printed ASM assembly code, this freeing is just an increment of [sp]: -<< - add sp, sp, #sz ->> - Again, our memory model cannot comprehend that this operation - frees (logically) the current stack frame. - -- [Pbtbl reg table]: this is a N-way branch, implemented via a jump table - as follows: -<< - la x31, table - add x31, x31, reg - jr x31 -table: .long table[0], table[1], ... ->> - Note that [reg] contains 4 times the index of the desired table entry. - -- [Pseq rd rs1 rs2]: since unsigned comparisons have particular - semantics for pointers, we cannot encode equality directly using - xor/sub etc, which have only integer semantics. -<< - xor rd, rs1, rs2 - sltiu rd, rd, 1 ->> - The [xor] is omitted if one of [rs1] and [rs2] is [x0]. - -- [Psne rd rs1 rs2]: similarly for unsigned inequality. -<< - xor rd, rs1, rs2 - sltu rd, x0, rd ->> -*) - -Definition code := list instruction. -Record function : Type := mkfunction { fn_sig: signature; fn_code: code }. -Definition fundef := AST.fundef function. -Definition program := AST.program fundef unit. - -(** * Operational semantics *) - -(** The semantics operates over a single mapping from registers - (type [preg]) to values. We maintain - the convention that integer registers are mapped to values of - type [Tint] or [Tlong] (in 64 bit mode), - and float registers to values of type [Tsingle] or [Tfloat]. *) - -Definition regset := Pregmap.t val. -Definition genv := Genv.t fundef unit. - -Definition get0w (rs: regset) (r: ireg0) : val := - match r with - | X0 => Vint Int.zero - | X r => rs r - end. - -Definition get0l (rs: regset) (r: ireg0) : val := - match r with - | X0 => Vlong Int64.zero - | X r => rs r - end. - -Notation "a # b" := (a b) (at level 1, only parsing) : asm. -Notation "a ## b" := (get0w a b) (at level 1) : asm. -Notation "a ### b" := (get0l a b) (at level 1) : asm. -Notation "a # b <- c" := (Pregmap.set b c a) (at level 1, b at next level) : asm. - -Open Scope asm. - -(** Undefining some registers *) - -Fixpoint undef_regs (l: list preg) (rs: regset) : regset := - match l with - | nil => rs - | r :: l' => undef_regs l' (rs#r <- Vundef) - end. - -(** Assigning a register pair *) - -Definition set_pair (p: rpair preg) (v: val) (rs: regset) : regset := - match p with - | One r => rs#r <- v - | Twolong rhi rlo => rs#rhi <- (Val.hiword v) #rlo <- (Val.loword v) - end. - -(** Assigning multiple registers *) - -Fixpoint set_regs (rl: list preg) (vl: list val) (rs: regset) : regset := - match rl, vl with - | r1 :: rl', v1 :: vl' => set_regs rl' vl' (rs#r1 <- v1) - | _, _ => rs - end. - -(** Assigning the result of a builtin *) - -Fixpoint set_res (res: builtin_res preg) (v: val) (rs: regset) : regset := - match res with - | BR r => rs#r <- v - | BR_none => rs - | BR_splitlong hi lo => set_res lo (Val.loword v) (set_res hi (Val.hiword v) rs) - end. - -Section RELSEM. - -(** Looking up instructions in a code sequence by position. *) - -Fixpoint find_instr (pos: Z) (c: code) {struct c} : option instruction := - match c with - | nil => None - | i :: il => if zeq pos 0 then Some i else find_instr (pos - 1) il - end. - -(** Position corresponding to a label *) - -Definition is_label (lbl: label) (instr: instruction) : bool := - match instr with - | Plabel lbl' => if peq lbl lbl' then true else false - | _ => false - end. - -Lemma is_label_correct: - forall lbl instr, - if is_label lbl instr then instr = Plabel lbl else instr <> Plabel lbl. -Proof. - intros. destruct instr; simpl; try discriminate. - case (peq lbl lbl0); intro; congruence. -Qed. - -Fixpoint label_pos (lbl: label) (pos: Z) (c: code) {struct c} : option Z := - match c with - | nil => None - | instr :: c' => - if is_label lbl instr then Some (pos + 1) else label_pos lbl (pos + 1) c' - end. - -Variable ge: genv. - -(** The two functions below axiomatize how the linker processes - symbolic references [symbol + offset)] and splits their - actual values into a 20-bit high part [%hi(symbol + offset)] and - a 12-bit low part [%lo(symbol + offset)]. *) - -Parameter low_half: genv -> ident -> ptrofs -> ptrofs. -Parameter high_half: genv -> ident -> ptrofs -> val. - -(** The fundamental property of these operations is that, when applied - to the address of a symbol, their results can be recombined by - addition, rebuilding the original address. *) - -Axiom low_high_half: - forall id ofs, - Val.offset_ptr (high_half ge id ofs) (low_half ge id ofs) = Genv.symbol_address ge id ofs. - -(** The semantics is purely small-step and defined as a function - from the current state (a register set + a memory state) - to either [Next rs' m'] where [rs'] and [m'] are the updated register - set and memory state after execution of the instruction at [rs#PC], - or [Stuck] if the processor is stuck. *) - -Inductive outcome: Type := - | Next: regset -> mem -> outcome - | Stuck: outcome. - -(** Manipulations over the [PC] register: continuing with the next - instruction ([nextinstr]) or branching to a label ([goto_label]). *) - -Definition nextinstr (rs: regset) := - rs#PC <- (Val.offset_ptr rs#PC Ptrofs.one). - -Definition goto_label (f: function) (lbl: label) (rs: regset) (m: mem) := - match label_pos lbl 0 (fn_code f) with - | None => Stuck - | Some pos => - match rs#PC with - | Vptr b ofs => Next (rs#PC <- (Vptr b (Ptrofs.repr pos))) m - | _ => Stuck - end - end. - -(** Auxiliaries for memory accesses *) - -Definition eval_offset (ofs: offset) : ptrofs := - match ofs with - | Ofsimm n => n - | Ofslow id delta => low_half ge id delta - end. - -Definition exec_load (chunk: memory_chunk) (rs: regset) (m: mem) - (d: preg) (a: ireg) (ofs: offset) := - match Mem.loadv chunk m (Val.offset_ptr (rs a) (eval_offset ofs)) with - | None => Stuck - | Some v => Next (nextinstr (rs#d <- v)) m - end. - -Definition exec_store (chunk: memory_chunk) (rs: regset) (m: mem) - (s: preg) (a: ireg) (ofs: offset) := - match Mem.storev chunk m (Val.offset_ptr (rs a) (eval_offset ofs)) (rs s) with - | None => Stuck - | Some m' => Next (nextinstr rs) m' - end. - -(** Evaluating a branch *) - -Definition eval_branch (f: function) (l: label) (rs: regset) (m: mem) (res: option bool) : outcome := - match res with - | Some true => goto_label f l rs m - | Some false => Next (nextinstr rs) m - | None => Stuck - end. - -(** Execution of a single instruction [i] in initial state [rs] and - [m]. Return updated state. For instructions that correspond to - actual RISC-V instructions, the cases are straightforward - transliterations of the informal descriptions given in the RISC-V - user-mode specification. For pseudo-instructions, refer to the - informal descriptions given above. - - Note that we set to [Vundef] the registers used as temporaries by - the expansions of the pseudo-instructions, so that the RISC-V code - we generate cannot use those registers to hold values that must - survive the execution of the pseudo-instruction. *) - -Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : outcome := - match i with - | Pmv d s => - Next (nextinstr (rs#d <- (rs#s))) m - -(** 32-bit integer register-immediate instructions *) - | Paddiw d s i => - Next (nextinstr (rs#d <- (Val.add rs##s (Vint i)))) m - | Psltiw d s i => - Next (nextinstr (rs#d <- (Val.cmp Clt rs##s (Vint i)))) m - | Psltiuw d s i => - Next (nextinstr (rs#d <- (Val.cmpu (Mem.valid_pointer m) Clt rs##s (Vint i)))) m - | Pandiw d s i => - Next (nextinstr (rs#d <- (Val.and rs##s (Vint i)))) m - | Poriw d s i => - Next (nextinstr (rs#d <- (Val.or rs##s (Vint i)))) m - | Pxoriw d s i => - Next (nextinstr (rs#d <- (Val.xor rs##s (Vint i)))) m - | Pslliw d s i => - Next (nextinstr (rs#d <- (Val.shl rs##s (Vint i)))) m - | Psrliw d s i => - Next (nextinstr (rs#d <- (Val.shru rs##s (Vint i)))) m - | Psraiw d s i => - Next (nextinstr (rs#d <- (Val.shr rs##s (Vint i)))) m - | Pluiw d i => - Next (nextinstr (rs#d <- (Vint (Int.shl i (Int.repr 12))))) m - -(** 32-bit integer register-register instructions *) - | Paddw d s1 s2 => - Next (nextinstr (rs#d <- (Val.add rs##s1 rs##s2))) m - | Psubw d s1 s2 => - Next (nextinstr (rs#d <- (Val.sub rs##s1 rs##s2))) m - | Pmulw d s1 s2 => - Next (nextinstr (rs#d <- (Val.mul rs##s1 rs##s2))) m - | Pmulhw d s1 s2 => - Next (nextinstr (rs#d <- (Val.mulhs rs##s1 rs##s2))) m - | Pmulhuw d s1 s2 => - Next (nextinstr (rs#d <- (Val.mulhu rs##s1 rs##s2))) m - | Pdivw d s1 s2 => - Next (nextinstr (rs#d <- (Val.maketotal (Val.divs rs##s1 rs##s2)))) m - | Pdivuw d s1 s2 => - Next (nextinstr (rs#d <- (Val.maketotal (Val.divu rs##s1 rs##s2)))) m - | Premw d s1 s2 => - Next (nextinstr (rs#d <- (Val.maketotal (Val.mods rs##s1 rs##s2)))) m - | Premuw d s1 s2 => - Next (nextinstr (rs#d <- (Val.maketotal (Val.modu rs##s1 rs##s2)))) m - | Psltw d s1 s2 => - Next (nextinstr (rs#d <- (Val.cmp Clt rs##s1 rs##s2))) m - | Psltuw d s1 s2 => - Next (nextinstr (rs#d <- (Val.cmpu (Mem.valid_pointer m) Clt rs##s1 rs##s2))) m - | Pseqw d s1 s2 => - Next (nextinstr (rs#d <- (Val.cmpu (Mem.valid_pointer m) Ceq rs##s1 rs##s2))) m - | Psnew d s1 s2 => - Next (nextinstr (rs#d <- (Val.cmpu (Mem.valid_pointer m) Cne rs##s1 rs##s2))) m - | Pandw d s1 s2 => - Next (nextinstr (rs#d <- (Val.and rs##s1 rs##s2))) m - | Porw d s1 s2 => - Next (nextinstr (rs#d <- (Val.or rs##s1 rs##s2))) m - | Pxorw d s1 s2 => - Next (nextinstr (rs#d <- (Val.xor rs##s1 rs##s2))) m - | Psllw d s1 s2 => - Next (nextinstr (rs#d <- (Val.shl rs##s1 rs##s2))) m - | Psrlw d s1 s2 => - Next (nextinstr (rs#d <- (Val.shru rs##s1 rs##s2))) m - | Psraw d s1 s2 => - Next (nextinstr (rs#d <- (Val.shr rs##s1 rs##s2))) m - -(** 64-bit integer register-immediate instructions *) - | Paddil d s i => - Next (nextinstr (rs#d <- (Val.addl rs###s (Vlong i)))) m - | Psltil d s i => - Next (nextinstr (rs#d <- (Val.maketotal (Val.cmpl Clt rs###s (Vlong i))))) m - | Psltiul d s i => - Next (nextinstr (rs#d <- (Val.maketotal (Val.cmplu (Mem.valid_pointer m) Clt rs###s (Vlong i))))) m - | Pandil d s i => - Next (nextinstr (rs#d <- (Val.andl rs###s (Vlong i)))) m - | Poril d s i => - Next (nextinstr (rs#d <- (Val.orl rs###s (Vlong i)))) m - | Pxoril d s i => - Next (nextinstr (rs#d <- (Val.xorl rs###s (Vlong i)))) m - | Psllil d s i => - Next (nextinstr (rs#d <- (Val.shll rs###s (Vint i)))) m - | Psrlil d s i => - Next (nextinstr (rs#d <- (Val.shrlu rs###s (Vint i)))) m - | Psrail d s i => - Next (nextinstr (rs#d <- (Val.shrl rs###s (Vint i)))) m - | Pluil d i => - Next (nextinstr (rs#d <- (Vlong (Int64.sign_ext 32 (Int64.shl i (Int64.repr 12)))))) m - -(** 64-bit integer register-register instructions *) - | Paddl d s1 s2 => - Next (nextinstr (rs#d <- (Val.addl rs###s1 rs###s2))) m - | Psubl d s1 s2 => - Next (nextinstr (rs#d <- (Val.subl rs###s1 rs###s2))) m - | Pmull d s1 s2 => - Next (nextinstr (rs#d <- (Val.mull rs###s1 rs###s2))) m - | Pmulhl d s1 s2 => - Next (nextinstr (rs#d <- (Val.mullhs rs###s1 rs###s2))) m - | Pmulhul d s1 s2 => - Next (nextinstr (rs#d <- (Val.mullhu rs###s1 rs###s2))) m - | Pdivl d s1 s2 => - Next (nextinstr (rs#d <- (Val.maketotal (Val.divls rs###s1 rs###s2)))) m - | Pdivul d s1 s2 => - Next (nextinstr (rs#d <- (Val.maketotal (Val.divlu rs###s1 rs###s2)))) m - | Preml d s1 s2 => - Next (nextinstr (rs#d <- (Val.maketotal (Val.modls rs###s1 rs###s2)))) m - | Premul d s1 s2 => - Next (nextinstr (rs#d <- (Val.maketotal (Val.modlu rs###s1 rs###s2)))) m - | Psltl d s1 s2 => - Next (nextinstr (rs#d <- (Val.maketotal (Val.cmpl Clt rs###s1 rs###s2)))) m - | Psltul d s1 s2 => - Next (nextinstr (rs#d <- (Val.maketotal (Val.cmplu (Mem.valid_pointer m) Clt rs###s1 rs###s2)))) m - | Pseql d s1 s2 => - Next (nextinstr (rs#d <- (Val.maketotal (Val.cmplu (Mem.valid_pointer m) Ceq rs###s1 rs###s2)))) m - | Psnel d s1 s2 => - Next (nextinstr (rs#d <- (Val.maketotal (Val.cmplu (Mem.valid_pointer m) Cne rs###s1 rs###s2)))) m - | Pandl d s1 s2 => - Next (nextinstr (rs#d <- (Val.andl rs###s1 rs###s2))) m - | Porl d s1 s2 => - Next (nextinstr (rs#d <- (Val.orl rs###s1 rs###s2))) m - | Pxorl d s1 s2 => - Next (nextinstr (rs#d <- (Val.xorl rs###s1 rs###s2))) m - | Pslll d s1 s2 => - Next (nextinstr (rs#d <- (Val.shll rs###s1 rs###s2))) m - | Psrll d s1 s2 => - Next (nextinstr (rs#d <- (Val.shrlu rs###s1 rs###s2))) m - | Psral d s1 s2 => - Next (nextinstr (rs#d <- (Val.shrl rs###s1 rs###s2))) m - - | Pcvtl2w d s => - Next (nextinstr (rs#d <- (Val.loword rs##s))) m - | Pcvtw2l r => - Next (nextinstr (rs#r <- (Val.longofint rs#r))) m - -(** Unconditional jumps. Links are always to X1/RA. *) - | Pj_l l => - goto_label f l rs m - | Pj_s s sg => - Next (rs#PC <- (Genv.symbol_address ge s Ptrofs.zero)) m - | Pj_r r sg => - Next (rs#PC <- (rs#r)) m - | Pjal_s s sg => - Next (rs#PC <- (Genv.symbol_address ge s Ptrofs.zero) - #RA <- (Val.offset_ptr rs#PC Ptrofs.one) - ) m - | Pjal_r r sg => - Next (rs#PC <- (rs#r) - #RA <- (Val.offset_ptr rs#PC Ptrofs.one) - ) m -(** Conditional branches, 32-bit comparisons *) - | Pbeqw s1 s2 l => - eval_branch f l rs m (Val.cmpu_bool (Mem.valid_pointer m) Ceq rs##s1 rs##s2) - | Pbnew s1 s2 l => - eval_branch f l rs m (Val.cmpu_bool (Mem.valid_pointer m) Cne rs##s1 rs##s2) - | Pbltw s1 s2 l => - eval_branch f l rs m (Val.cmp_bool Clt rs##s1 rs##s2) - | Pbltuw s1 s2 l => - eval_branch f l rs m (Val.cmpu_bool (Mem.valid_pointer m) Clt rs##s1 rs##s2) - | Pbgew s1 s2 l => - eval_branch f l rs m (Val.cmp_bool Cge rs##s1 rs##s2) - | Pbgeuw s1 s2 l => - eval_branch f l rs m (Val.cmpu_bool (Mem.valid_pointer m) Cge rs##s1 rs##s2) - -(** Conditional branches, 64-bit comparisons *) - | Pbeql s1 s2 l => - eval_branch f l rs m (Val.cmplu_bool (Mem.valid_pointer m) Ceq rs###s1 rs###s2) - | Pbnel s1 s2 l => - eval_branch f l rs m (Val.cmplu_bool (Mem.valid_pointer m) Cne rs###s1 rs###s2) - | Pbltl s1 s2 l => - eval_branch f l rs m (Val.cmpl_bool Clt rs###s1 rs###s2) - | Pbltul s1 s2 l => - eval_branch f l rs m (Val.cmplu_bool (Mem.valid_pointer m) Clt rs###s1 rs###s2) - | Pbgel s1 s2 l => - eval_branch f l rs m (Val.cmpl_bool Cge rs###s1 rs###s2) - | Pbgeul s1 s2 l => - eval_branch f l rs m (Val.cmplu_bool (Mem.valid_pointer m) Cge rs###s1 rs###s2) - -(** Loads and stores *) - | Plb d a ofs => - exec_load Mint8signed rs m d a ofs - | Plbu d a ofs => - exec_load Mint8unsigned rs m d a ofs - | Plh d a ofs => - exec_load Mint16signed rs m d a ofs - | Plhu d a ofs => - exec_load Mint16unsigned rs m d a ofs - | Plw d a ofs => - exec_load Mint32 rs m d a ofs - | Plw_a d a ofs => - exec_load Many32 rs m d a ofs - | Pld d a ofs => - exec_load Mint64 rs m d a ofs - | Pld_a d a ofs => - exec_load Many64 rs m d a ofs - | Psb s a ofs => - exec_store Mint8unsigned rs m s a ofs - | Psh s a ofs => - exec_store Mint16unsigned rs m s a ofs - | Psw s a ofs => - exec_store Mint32 rs m s a ofs - | Psw_a s a ofs => - exec_store Many32 rs m s a ofs - | Psd s a ofs => - exec_store Mint64 rs m s a ofs - | Psd_a s a ofs => - exec_store Many64 rs m s a ofs - -(** Floating point register move *) - | Pfmv d s => - Next (nextinstr (rs#d <- (rs#s))) m - -(** 32-bit (single-precision) floating point *) - | Pfls d a ofs => - exec_load Mfloat32 rs m d a ofs - | Pfss s a ofs => - exec_store Mfloat32 rs m s a ofs - - | Pfnegs d s => - Next (nextinstr (rs#d <- (Val.negfs rs#s))) m - | Pfabss d s => - Next (nextinstr (rs#d <- (Val.absfs rs#s))) m - - | Pfadds d s1 s2 => - Next (nextinstr (rs#d <- (Val.addfs rs#s1 rs#s2))) m - | Pfsubs d s1 s2 => - Next (nextinstr (rs#d <- (Val.subfs rs#s1 rs#s2))) m - | Pfmuls d s1 s2 => - Next (nextinstr (rs#d <- (Val.mulfs rs#s1 rs#s2))) m - | Pfdivs d s1 s2 => - Next (nextinstr (rs#d <- (Val.divfs rs#s1 rs#s2))) m - | Pfeqs d s1 s2 => - Next (nextinstr (rs#d <- (Val.cmpfs Ceq rs#s1 rs#s2))) m - | Pflts d s1 s2 => - Next (nextinstr (rs#d <- (Val.cmpfs Clt rs#s1 rs#s2))) m - | Pfles d s1 s2 => - Next (nextinstr (rs#d <- (Val.cmpfs Cle rs#s1 rs#s2))) m - - | Pfcvtws d s => - Next (nextinstr (rs#d <- (Val.maketotal (Val.intofsingle rs#s)))) m - | Pfcvtwus d s => - Next (nextinstr (rs#d <- (Val.maketotal (Val.intuofsingle rs#s)))) m - | Pfcvtsw d s => - Next (nextinstr (rs#d <- (Val.maketotal (Val.singleofint rs##s)))) m - | Pfcvtswu d s => - Next (nextinstr (rs#d <- (Val.maketotal (Val.singleofintu rs##s)))) m - - | Pfcvtls d s => - Next (nextinstr (rs#d <- (Val.maketotal (Val.longofsingle rs#s)))) m - | Pfcvtlus d s => - Next (nextinstr (rs#d <- (Val.maketotal (Val.longuofsingle rs#s)))) m - | Pfcvtsl d s => - Next (nextinstr (rs#d <- (Val.maketotal (Val.singleoflong rs###s)))) m - | Pfcvtslu d s => - Next (nextinstr (rs#d <- (Val.maketotal (Val.singleoflongu rs###s)))) m - -(** 64-bit (double-precision) floating point *) - | Pfld d a ofs => - exec_load Mfloat64 rs m d a ofs - | Pfld_a d a ofs => - exec_load Many64 rs m d a ofs - | Pfsd s a ofs => - exec_store Mfloat64 rs m s a ofs - | Pfsd_a s a ofs => - exec_store Many64 rs m s a ofs - - | Pfnegd d s => - Next (nextinstr (rs#d <- (Val.negf rs#s))) m - | Pfabsd d s => - Next (nextinstr (rs#d <- (Val.absf rs#s))) m - - | Pfaddd d s1 s2 => - Next (nextinstr (rs#d <- (Val.addf rs#s1 rs#s2))) m - | Pfsubd d s1 s2 => - Next (nextinstr (rs#d <- (Val.subf rs#s1 rs#s2))) m - | Pfmuld d s1 s2 => - Next (nextinstr (rs#d <- (Val.mulf rs#s1 rs#s2))) m - | Pfdivd d s1 s2 => - Next (nextinstr (rs#d <- (Val.divf rs#s1 rs#s2))) m - | Pfeqd d s1 s2 => - Next (nextinstr (rs#d <- (Val.cmpf Ceq rs#s1 rs#s2))) m - | Pfltd d s1 s2 => - Next (nextinstr (rs#d <- (Val.cmpf Clt rs#s1 rs#s2))) m - | Pfled d s1 s2 => - Next (nextinstr (rs#d <- (Val.cmpf Cle rs#s1 rs#s2))) m - - | Pfcvtwd d s => - Next (nextinstr (rs#d <- (Val.maketotal (Val.intoffloat rs#s)))) m - | Pfcvtwud d s => - Next (nextinstr (rs#d <- (Val.maketotal (Val.intuoffloat rs#s)))) m - | Pfcvtdw d s => - Next (nextinstr (rs#d <- (Val.maketotal (Val.floatofint rs##s)))) m - | Pfcvtdwu d s => - Next (nextinstr (rs#d <- (Val.maketotal (Val.floatofintu rs##s)))) m - - | Pfcvtld d s => - Next (nextinstr (rs#d <- (Val.maketotal (Val.longoffloat rs#s)))) m - | Pfcvtlud d s => - Next (nextinstr (rs#d <- (Val.maketotal (Val.longuoffloat rs#s)))) m - | Pfcvtdl d s => - Next (nextinstr (rs#d <- (Val.maketotal (Val.floatoflong rs###s)))) m - | Pfcvtdlu d s => - Next (nextinstr (rs#d <- (Val.maketotal (Val.floatoflongu rs###s)))) m - - | Pfcvtds d s => - Next (nextinstr (rs#d <- (Val.floatofsingle rs#s))) m - | Pfcvtsd d s => - Next (nextinstr (rs#d <- (Val.singleoffloat rs#s))) m - - | Pfmvxs d s => - Next (nextinstr (rs#d <- (ExtValues.bits_of_single rs#s))) m - | Pfmvxd d s => - Next (nextinstr (rs#d <- (ExtValues.bits_of_float rs#s))) m - - | Pfmvsx d s => - Next (nextinstr (rs#d <- (ExtValues.single_of_bits rs#s))) m - | Pfmvdx d s => - Next (nextinstr (rs#d <- (ExtValues.float_of_bits rs#s))) m - - -(** Pseudo-instructions *) - | Pallocframe sz pos => - let (m1, stk) := Mem.alloc m 0 sz in - let sp := (Vptr stk Ptrofs.zero) in - match Mem.storev Mptr m1 (Val.offset_ptr sp pos) rs#SP with - | None => Stuck - | Some m2 => Next (nextinstr (rs #X30 <- (rs SP) #SP <- sp #X31 <- Vundef)) m2 - end - | Pfreeframe sz pos => - match Mem.loadv Mptr m (Val.offset_ptr rs#SP pos) with - | None => Stuck - | Some v => - match rs SP with - | Vptr stk ofs => - match Mem.free m stk 0 sz with - | None => Stuck - | Some m' => Next (nextinstr (rs#SP <- v #X31 <- Vundef)) m' - end - | _ => Stuck - end - end - | Pselectl rd rb rt rf => - Next (nextinstr (rs#rd <- (ExtValues.select01_long - (rs###rb) (rs###rt) (rs###rf))) - #X31 <- Vundef) m - | Plabel lbl => - Next (nextinstr rs) m - | Ploadsymbol rd s ofs => - Next (nextinstr (rs#rd <- (Genv.symbol_address ge s ofs))) m - | Ploadsymbol_high rd s ofs => - Next (nextinstr (rs#rd <- (high_half ge s ofs))) m - | Ploadli rd i => - Next (nextinstr (rs#X31 <- Vundef #rd <- (Vlong i))) m - | Ploadfi rd f => - Next (nextinstr (rs#X31 <- Vundef #rd <- (Vfloat f))) m - | Ploadsi rd f => - Next (nextinstr (rs#X31 <- Vundef #rd <- (Vsingle f))) m - | Pbtbl r tbl => - match rs r with - | Vint n => - match list_nth_z tbl (Int.unsigned n) with - | None => Stuck - | Some lbl => goto_label f lbl (rs#X5 <- Vundef #X31 <- Vundef) m - end - | _ => Stuck - end - | Pbuiltin ef args res => - Stuck (**r treated specially below *) - - (** The following instructions and directives are not generated directly by Asmgen, - so we do not model them. *) - | Pfence - -<<<<<<< HEAD -======= - | Pfmvxs _ _ - | Pfmvsx _ _ - | Pfmvxd _ _ - | Pfmvdx _ _ - ->>>>>>> master - | Pfmins _ _ _ - | Pfmaxs _ _ _ - | Pfsqrts _ _ - | Pfmadds _ _ _ _ - | Pfmsubs _ _ _ _ - | Pfnmadds _ _ _ _ - | Pfnmsubs _ _ _ _ - - | Pfmind _ _ _ - | Pfmaxd _ _ _ - | Pfsqrtd _ _ - | Pfmaddd _ _ _ _ - | Pfmsubd _ _ _ _ - | Pfnmaddd _ _ _ _ - | Pfnmsubd _ _ _ _ - | Pnop - => Stuck - end. - -(** Translation of the LTL/Linear/Mach view of machine registers to - the RISC-V view. Note that no LTL register maps to [X31]. This - register is reserved as temporary, to be used by the generated RV32G - code. *) - -Definition preg_of (r: mreg) : preg := - match r with - | R5 => X5 | R6 => X6 | R7 => X7 - | R8 => X8 | R9 => X9 | R10 => X10 | R11 => X11 - | R12 => X12 | R13 => X13 | R14 => X14 | R15 => X15 - | R16 => X16 | R17 => X17 | R18 => X18 | R19 => X19 - | R20 => X20 | R21 => X21 | R22 => X22 | R23 => X23 - | R24 => X24 | R25 => X25 | R26 => X26 | R27 => X27 - | R28 => X28 | R29 => X29 | R30 => X30 - - | Machregs.F0 => F0 | Machregs.F1 => F1 | Machregs.F2 => F2 | Machregs.F3 => F3 - | Machregs.F4 => F4 | Machregs.F5 => F5 | Machregs.F6 => F6 | Machregs.F7 => F7 - | Machregs.F8 => F8 | Machregs.F9 => F9 | Machregs.F10 => F10 | Machregs.F11 => F11 - | Machregs.F12 => F12 | Machregs.F13 => F13 | Machregs.F14 => F14 | Machregs.F15 => F15 - | Machregs.F16 => F16 | Machregs.F17 => F17 | Machregs.F18 => F18 | Machregs.F19 => F19 - | Machregs.F20 => F20 | Machregs.F21 => F21 | Machregs.F22 => F22 | Machregs.F23 => F23 - | Machregs.F24 => F24 | Machregs.F25 => F25 | Machregs.F26 => F26 | Machregs.F27 => F27 - | Machregs.F28 => F28 | Machregs.F29 => F29 | Machregs.F30 => F30 | Machregs.F31 => F31 - end. - -(** Undefine all registers except SP and callee-save registers *) - -Definition undef_caller_save_regs (rs: regset) : regset := - fun r => - if preg_eq r SP - || In_dec preg_eq r (List.map preg_of (List.filter is_callee_save all_mregs)) - then rs r - else Vundef. - -(** Extract the values of the arguments of an external call. - We exploit the calling conventions from module [Conventions], except that - we use RISC-V registers instead of locations. *) - -Inductive extcall_arg (rs: regset) (m: mem): loc -> val -> Prop := - | extcall_arg_reg: forall r, - extcall_arg rs m (R r) (rs (preg_of r)) - | extcall_arg_stack: forall ofs ty bofs v, - bofs = Stacklayout.fe_ofs_arg + 4 * ofs -> - Mem.loadv (chunk_of_type ty) m - (Val.offset_ptr rs#SP (Ptrofs.repr bofs)) = Some v -> - extcall_arg rs m (S Outgoing ofs ty) v. - -Inductive extcall_arg_pair (rs: regset) (m: mem): rpair loc -> val -> Prop := - | extcall_arg_one: forall l v, - extcall_arg rs m l v -> - extcall_arg_pair rs m (One l) v - | extcall_arg_twolong: forall hi lo vhi vlo, - extcall_arg rs m hi vhi -> - extcall_arg rs m lo vlo -> - extcall_arg_pair rs m (Twolong hi lo) (Val.longofwords vhi vlo). - -Definition extcall_arguments - (rs: regset) (m: mem) (sg: signature) (args: list val) : Prop := - list_forall2 (extcall_arg_pair rs m) (loc_arguments sg) args. - -Definition loc_external_result (sg: signature) : rpair preg := - map_rpair preg_of (loc_result sg). - -(** Execution of the instruction at [rs PC]. *) - -Inductive state: Type := - | State: regset -> mem -> state. - -Inductive step: state -> trace -> state -> Prop := - | exec_step_internal: - forall b ofs f i rs m rs' m', - rs PC = Vptr b ofs -> - Genv.find_funct_ptr ge b = Some (Internal f) -> - find_instr (Ptrofs.unsigned ofs) (fn_code f) = Some i -> - exec_instr f i rs m = Next rs' m' -> - step (State rs m) E0 (State rs' m') - | exec_step_builtin: - forall b ofs f ef args res rs m vargs t vres rs' m', - rs PC = Vptr b ofs -> - Genv.find_funct_ptr ge b = Some (Internal f) -> - find_instr (Ptrofs.unsigned ofs) f.(fn_code) = Some (Pbuiltin ef args res) -> - eval_builtin_args ge rs (rs SP) m args vargs -> - external_call ef ge vargs m t vres m' -> - rs' = nextinstr - (set_res res vres - (undef_regs (map preg_of (destroyed_by_builtin ef)) - (rs#X31 <- Vundef))) -> - step (State rs m) t (State rs' m') - | exec_step_external: - forall b ef args res rs m t rs' m', - rs PC = Vptr b Ptrofs.zero -> - Genv.find_funct_ptr ge b = Some (External ef) -> - external_call ef ge args m t res m' -> - extcall_arguments rs m (ef_sig ef) args -> - rs' = (set_pair (loc_external_result (ef_sig ef) ) res (undef_caller_save_regs rs))#PC <- (rs RA) -> - step (State rs m) t (State rs' m'). - -End RELSEM. - -(** Execution of whole programs. *) - -Inductive initial_state (p: program): state -> Prop := - | initial_state_intro: forall m0, - let ge := Genv.globalenv p in - let rs0 := - (Pregmap.init Vundef) - # PC <- (Genv.symbol_address ge p.(prog_main) Ptrofs.zero) - # SP <- Vnullptr - # RA <- Vnullptr in - Genv.init_mem p = Some m0 -> - initial_state p (State rs0 m0). - -Inductive final_state: state -> int -> Prop := - | final_state_intro: forall rs m r, - rs PC = Vnullptr -> - rs X10 = Vint r -> - final_state (State rs m) r. - -Definition semantics (p: program) := - Semantics step (initial_state p) final_state (Genv.globalenv p). - -(** Determinacy of the [Asm] semantics. *) - -Remark extcall_arguments_determ: - forall rs m sg args1 args2, - extcall_arguments rs m sg args1 -> extcall_arguments rs m sg args2 -> args1 = args2. -Proof. - intros until m. - assert (A: forall l v1 v2, - extcall_arg rs m l v1 -> extcall_arg rs m l v2 -> v1 = v2). - { intros. inv H; inv H0; congruence. } - assert (B: forall p v1 v2, - extcall_arg_pair rs m p v1 -> extcall_arg_pair rs m p v2 -> v1 = v2). - { intros. inv H; inv H0. - eapply A; eauto. - f_equal; eapply A; eauto. } - assert (C: forall ll vl1, list_forall2 (extcall_arg_pair rs m) ll vl1 -> - forall vl2, list_forall2 (extcall_arg_pair rs m) ll vl2 -> vl1 = vl2). - { - induction 1; intros vl2 EA; inv EA. - auto. - f_equal; eauto. } - intros. eapply C; eauto. -Qed. - -Lemma semantics_determinate: forall p, determinate (semantics p). -Proof. -Ltac Equalities := - match goal with - | [ H1: ?a = ?b, H2: ?a = ?c |- _ ] => - rewrite H1 in H2; inv H2; Equalities - | _ => idtac - end. - intros; constructor; simpl; intros. -- (* determ *) - inv H; inv H0; Equalities. - split. constructor. auto. - discriminate. - discriminate. - assert (vargs0 = vargs) by (eapply eval_builtin_args_determ; eauto). subst vargs0. - exploit external_call_determ. eexact H5. eexact H11. intros [A B]. - split. auto. intros. destruct B; auto. subst. auto. - assert (args0 = args) by (eapply extcall_arguments_determ; eauto). subst args0. - exploit external_call_determ. eexact H3. eexact H8. intros [A B]. - split. auto. intros. destruct B; auto. subst. auto. -- (* trace length *) - red; intros. inv H; simpl. - lia. - eapply external_call_trace_length; eauto. - eapply external_call_trace_length; eauto. -- (* initial states *) - inv H; inv H0. f_equal. congruence. -- (* final no step *) - assert (NOTNULL: forall b ofs, Vnullptr <> Vptr b ofs). - { intros; unfold Vnullptr; destruct Archi.ptr64; congruence. } - inv H. unfold Vzero in H0. red; intros; red; intros. - inv H; rewrite H0 in *; eelim NOTNULL; eauto. -- (* final states *) - inv H; inv H0. congruence. -Qed. - -(** Classification functions for processor registers (used in Asmgenproof). *) - -Definition data_preg (r: preg) : bool := - match r with - | IR RA => false - | IR X31 => false - | IR _ => true - | FR _ => true - | PC => false - end. diff --git a/riscV/TO_MERGE/Asmgenproof1.v b/riscV/TO_MERGE/Asmgenproof1.v deleted file mode 100644 index 1a8ce27d..00000000 --- a/riscV/TO_MERGE/Asmgenproof1.v +++ /dev/null @@ -1,1367 +0,0 @@ -(* *********************************************************************) -(* *) -(* The Compcert verified compiler *) -(* *) -(* Xavier Leroy, INRIA Paris-Rocquencourt *) -(* Prashanth Mundkur, SRI International *) -(* *) -(* 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. *) -(* *) -(* The contributions by Prashanth Mundkur are reused and adapted *) -(* under the terms of a Contributor License Agreement between *) -(* SRI International and INRIA. *) -(* *) -(* *********************************************************************) - -Require Import Coqlib Errors Maps. -Require Import AST Zbits Integers Floats Values Memory Globalenvs. -Require Import Op Locations Mach Conventions. -Require Import Asm Asmgen Asmgenproof0. - -(** Decomposition of integer constants. *) - -Lemma make_immed32_sound: - forall n, - match make_immed32 n with - | Imm32_single imm => n = imm - | Imm32_pair hi lo => n = Int.add (Int.shl hi (Int.repr 12)) lo - end. -Proof. - intros; unfold make_immed32. set (lo := Int.sign_ext 12 n). - predSpec Int.eq Int.eq_spec n lo. -- auto. -- set (m := Int.sub n lo). - assert (A: eqmod (two_p 12) (Int.unsigned lo) (Int.unsigned n)) by (apply Int.eqmod_sign_ext'; compute; auto). - assert (B: eqmod (two_p 12) (Int.unsigned n - Int.unsigned lo) 0). - { replace 0 with (Int.unsigned n - Int.unsigned n) by lia. - auto using eqmod_sub, eqmod_refl. } - assert (C: eqmod (two_p 12) (Int.unsigned m) 0). - { apply eqmod_trans with (Int.unsigned n - Int.unsigned lo); auto. - apply eqmod_divides with Int.modulus. apply Int.eqm_sym; apply Int.eqm_unsigned_repr. - exists (two_p (32-12)); auto. } - assert (D: Int.modu m (Int.repr 4096) = Int.zero). - { apply eqmod_mod_eq in C. unfold Int.modu. - change (Int.unsigned (Int.repr 4096)) with (two_p 12). rewrite C. - reflexivity. - apply two_p_gt_ZERO; lia. } - rewrite <- (Int.divu_pow2 m (Int.repr 4096) (Int.repr 12)) by auto. - rewrite Int.shl_mul_two_p. - change (two_p (Int.unsigned (Int.repr 12))) with 4096. - replace (Int.mul (Int.divu m (Int.repr 4096)) (Int.repr 4096)) with m. - unfold m. rewrite Int.sub_add_opp. rewrite Int.add_assoc. rewrite <- (Int.add_commut lo). - rewrite Int.add_neg_zero. rewrite Int.add_zero. auto. - rewrite (Int.modu_divu_Euclid m (Int.repr 4096)) at 1 by (vm_compute; congruence). - rewrite D. apply Int.add_zero. -Qed. - -Lemma make_immed64_sound: - forall n, - match make_immed64 n with - | Imm64_single imm => n = imm - | Imm64_pair hi lo => n = Int64.add (Int64.sign_ext 32 (Int64.shl hi (Int64.repr 12))) lo - | Imm64_large imm => n = imm - end. -Proof. - intros; unfold make_immed64. set (lo := Int64.sign_ext 12 n). - predSpec Int64.eq Int64.eq_spec n lo. -- auto. -- set (m := Int64.sub n lo). - set (p := Int64.zero_ext 20 (Int64.shru m (Int64.repr 12))). - predSpec Int64.eq Int64.eq_spec n (Int64.add (Int64.sign_ext 32 (Int64.shl p (Int64.repr 12))) lo). - auto. - auto. -Qed. - -(** Properties of registers *) - -Lemma ireg_of_not_X31: - forall m r, ireg_of m = OK r -> IR r <> IR X31. -Proof. - intros. erewrite <- ireg_of_eq; eauto with asmgen. -Qed. - -Lemma ireg_of_not_X31': - forall m r, ireg_of m = OK r -> r <> X31. -Proof. - intros. apply ireg_of_not_X31 in H. congruence. -Qed. - -Global Hint Resolve ireg_of_not_X31 ireg_of_not_X31': asmgen. - -(** Useful simplification tactic *) - -Ltac Simplif := - ((rewrite nextinstr_inv by eauto with asmgen) - || (rewrite nextinstr_inv1 by eauto with asmgen) - || (rewrite Pregmap.gss) - || (rewrite nextinstr_pc) - || (rewrite Pregmap.gso by eauto with asmgen)); auto with asmgen. - -Ltac Simpl := repeat Simplif. - -(** * Correctness of RISC-V constructor functions *) - -Section CONSTRUCTORS. - -Variable ge: genv. -Variable fn: function. - -(** 32-bit integer constants and arithmetic *) - -Lemma load_hilo32_correct: - forall rd hi lo k rs m, - exists rs', - exec_straight ge fn (load_hilo32 rd hi lo k) rs m k rs' m - /\ rs'#rd = Vint (Int.add (Int.shl hi (Int.repr 12)) lo) - /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r. -Proof. - unfold load_hilo32; intros. - predSpec Int.eq Int.eq_spec lo Int.zero. -- subst lo. econstructor; split. - apply exec_straight_one. simpl; eauto. auto. - split. rewrite Int.add_zero. Simpl. - intros; Simpl. -- econstructor; split. - eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto. - split. Simpl. - intros; Simpl. -Qed. - -Lemma loadimm32_correct: - forall rd n k rs m, - exists rs', - exec_straight ge fn (loadimm32 rd n k) rs m k rs' m - /\ rs'#rd = Vint n - /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r. -Proof. - unfold loadimm32; intros. generalize (make_immed32_sound n); intros E. - destruct (make_immed32 n). -- subst imm. econstructor; split. - apply exec_straight_one. simpl; eauto. auto. - split. rewrite Int.add_zero_l; Simpl. - intros; Simpl. -- rewrite E. apply load_hilo32_correct. -Qed. - -Lemma opimm32_correct: - forall (op: ireg -> ireg0 -> ireg0 -> instruction) - (opi: ireg -> ireg0 -> int -> instruction) - (sem: val -> val -> val) m, - (forall d s1 s2 rs, - exec_instr ge fn (op d s1 s2) rs m = Next (nextinstr (rs#d <- (sem rs##s1 rs##s2))) m) -> - (forall d s n rs, - exec_instr ge fn (opi d s n) rs m = Next (nextinstr (rs#d <- (sem rs##s (Vint n)))) m) -> - forall rd r1 n k rs, - r1 <> X31 -> - exists rs', - exec_straight ge fn (opimm32 op opi rd r1 n k) rs m k rs' m - /\ rs'#rd = sem rs##r1 (Vint n) - /\ forall r, r <> PC -> r <> rd -> r <> X31 -> rs'#r = rs#r. -Proof. - intros. unfold opimm32. generalize (make_immed32_sound n); intros E. - destruct (make_immed32 n). -- subst imm. econstructor; split. - apply exec_straight_one. rewrite H0. simpl; eauto. auto. - split. Simpl. intros; Simpl. -- destruct (load_hilo32_correct X31 hi lo (op rd r1 X31 :: k) rs m) - as (rs' & A & B & C). - econstructor; split. - eapply exec_straight_trans. eexact A. apply exec_straight_one. - rewrite H; eauto. auto. - split. Simpl. simpl. rewrite B, C, E. auto. congruence. congruence. - intros; Simpl. -Qed. - -(** 64-bit integer constants and arithmetic *) - -Lemma load_hilo64_correct: - forall rd hi lo k rs m, - exists rs', - exec_straight ge fn (load_hilo64 rd hi lo k) rs m k rs' m - /\ rs'#rd = Vlong (Int64.add (Int64.sign_ext 32 (Int64.shl hi (Int64.repr 12))) lo) - /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r. -Proof. - unfold load_hilo64; intros. - predSpec Int64.eq Int64.eq_spec lo Int64.zero. -- subst lo. econstructor; split. - apply exec_straight_one. simpl; eauto. auto. - split. rewrite Int64.add_zero. Simpl. - intros; Simpl. -- econstructor; split. - eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto. - split. Simpl. - intros; Simpl. -Qed. - -Lemma loadimm64_correct: - forall rd n k rs m, - exists rs', - exec_straight ge fn (loadimm64 rd n k) rs m k rs' m - /\ rs'#rd = Vlong n - /\ forall r, r <> PC -> r <> rd -> r <> X31 -> rs'#r = rs#r. -Proof. - unfold loadimm64; intros. generalize (make_immed64_sound n); intros E. - destruct (make_immed64 n). -- subst imm. econstructor; split. - apply exec_straight_one. simpl; eauto. auto. - split. rewrite Int64.add_zero_l; Simpl. - intros; Simpl. -- exploit load_hilo64_correct; eauto. intros (rs' & A & B & C). - rewrite E. exists rs'; eauto. -- subst imm. econstructor; split. - apply exec_straight_one. simpl; eauto. auto. - split. Simpl. - intros; Simpl. -Qed. - -Lemma opimm64_correct: - forall (op: ireg -> ireg0 -> ireg0 -> instruction) - (opi: ireg -> ireg0 -> int64 -> instruction) - (sem: val -> val -> val) m, - (forall d s1 s2 rs, - exec_instr ge fn (op d s1 s2) rs m = Next (nextinstr (rs#d <- (sem rs###s1 rs###s2))) m) -> - (forall d s n rs, - exec_instr ge fn (opi d s n) rs m = Next (nextinstr (rs#d <- (sem rs###s (Vlong n)))) m) -> - forall rd r1 n k rs, - r1 <> X31 -> - exists rs', - exec_straight ge fn (opimm64 op opi rd r1 n k) rs m k rs' m - /\ rs'#rd = sem rs##r1 (Vlong n) - /\ forall r, r <> PC -> r <> rd -> r <> X31 -> rs'#r = rs#r. -Proof. - intros. unfold opimm64. generalize (make_immed64_sound n); intros E. - destruct (make_immed64 n). -- subst imm. econstructor; split. - apply exec_straight_one. rewrite H0. simpl; eauto. auto. - split. Simpl. intros; Simpl. -- destruct (load_hilo64_correct X31 hi lo (op rd r1 X31 :: k) rs m) - as (rs' & A & B & C). - econstructor; split. - eapply exec_straight_trans. eexact A. apply exec_straight_one. - rewrite H; eauto. auto. - split. Simpl. simpl. rewrite B, C, E. auto. congruence. congruence. - intros; Simpl. -- subst imm. econstructor; split. - eapply exec_straight_two. simpl; eauto. rewrite H. simpl; eauto. auto. auto. - split. Simpl. intros; Simpl. -Qed. - -(** Add offset to pointer *) - -Lemma addptrofs_correct: - forall rd r1 n k rs m, - r1 <> X31 -> - exists rs', - exec_straight ge fn (addptrofs rd r1 n k) rs m k rs' m - /\ Val.lessdef (Val.offset_ptr rs#r1 n) rs'#rd - /\ forall r, r <> PC -> r <> rd -> r <> X31 -> rs'#r = rs#r. -Proof. - unfold addptrofs; intros. - destruct (Ptrofs.eq_dec n Ptrofs.zero). -- subst n. econstructor; split. - apply exec_straight_one. simpl; eauto. auto. - split. Simpl. destruct (rs r1); simpl; auto. rewrite Ptrofs.add_zero; auto. - intros; Simpl. -- destruct Archi.ptr64 eqn:SF. -+ unfold addimm64. - exploit (opimm64_correct Paddl Paddil Val.addl); eauto. intros (rs' & A & B & C). - exists rs'; split. eexact A. split; auto. - rewrite B. simpl. destruct (rs r1); simpl; auto. rewrite SF. - rewrite Ptrofs.of_int64_to_int64 by auto. auto. -+ unfold addimm32. - exploit (opimm32_correct Paddw Paddiw Val.add); eauto. intros (rs' & A & B & C). - exists rs'; split. eexact A. split; auto. - rewrite B. simpl. destruct (rs r1); simpl; auto. rewrite SF. - rewrite Ptrofs.of_int_to_int by auto. auto. -Qed. - -Lemma addptrofs_correct_2: - forall rd r1 n k (rs: regset) m b ofs, - r1 <> X31 -> rs#r1 = Vptr b ofs -> - exists rs', - exec_straight ge fn (addptrofs rd r1 n k) rs m k rs' m - /\ rs'#rd = Vptr b (Ptrofs.add ofs n) - /\ forall r, r <> PC -> r <> rd -> r <> X31 -> rs'#r = rs#r. -Proof. - intros. exploit (addptrofs_correct rd r1 n); eauto. intros (rs' & A & B & C). - exists rs'; intuition eauto. - rewrite H0 in B. inv B. auto. -Qed. - -Ltac ArgsInv := - repeat (match goal with - | [ H: Error _ = OK _ |- _ ] => discriminate - | [ H: match ?args with nil => _ | _ :: _ => _ end = OK _ |- _ ] => destruct args - | [ H: bind _ _ = OK _ |- _ ] => monadInv H - | [ H: match _ with left _ => _ | right _ => assertion_failed end = OK _ |- _ ] => monadInv H; ArgsInv - | [ H: match _ with true => _ | false => assertion_failed end = OK _ |- _ ] => monadInv H; ArgsInv - end); - subst; - repeat (match goal with - | [ H: ireg_of _ = OK _ |- _ ] => simpl in *; rewrite (ireg_of_eq _ _ H) in * - | [ H: freg_of _ = OK _ |- _ ] => simpl in *; rewrite (freg_of_eq _ _ H) in * - end). - -Lemma transl_cbranch_correct_1: - forall cond args lbl k c m ms b sp rs m', - transl_cbranch cond args lbl k = OK c -> - eval_condition cond (List.map ms args) m = Some b -> - agree ms sp rs -> - Mem.extends m m' -> - exists rs', exists insn, - exec_straight_opt ge fn c rs m' (insn :: k) rs' m' - /\ exec_instr ge fn insn rs' m' = eval_branch fn lbl rs' m' (Some b) - /\ forall r, r <> PC -> r <> X31 -> rs'#r = rs#r. -Proof. - intros until m'; intros TRANSL EVAL AG MEXT. - set (vl' := map rs (map preg_of args)). - assert (EVAL': eval_condition cond vl' m' = Some b). - { apply eval_condition_lessdef with (map ms args) m; auto. eapply preg_vals; eauto. } - clear EVAL MEXT AG. - destruct cond; simpl in TRANSL; ArgsInv. - (* Pbeqw / Cmp *) - { destruct optR0 as [[]|]; - unfold apply_bin_r0, apply_bin_r0_r0r0lbl in *; - unfold zero32, Op.zero32 in *; - eexists; eexists; eauto; split; constructor; auto; - simpl in *. - + destruct (rs x); simpl in *; try congruence. - assert (HB: (Int.eq Int.zero i) = b) by congruence. - rewrite HB; destruct b; simpl; auto. - + destruct (rs x); simpl in *; try congruence. - assert (HB: (Int.eq i Int.zero) = b) by congruence. - rewrite HB; destruct b; simpl; auto. - + destruct (rs x); simpl in *; try congruence. - destruct (rs x0); try congruence. - assert (HB: (Int.eq i i0) = b) by congruence. - rewrite HB; destruct b; simpl; auto. } - (* Pbnew / Cmp *) - { destruct optR0 as [[]|]; - unfold apply_bin_r0, apply_bin_r0_r0r0lbl in *; - unfold zero32, Op.zero32 in *; - eexists; eexists; eauto; split; constructor; auto; - simpl in *. - + destruct (rs x); simpl in *; try congruence. - assert (HB: negb (Int.eq Int.zero i) = b) by congruence. - rewrite HB; destruct b; simpl; auto. - + destruct (rs x); simpl in *; try congruence. - assert (HB: negb (Int.eq i Int.zero) = b) by congruence. - rewrite HB; destruct b; simpl; auto. - + destruct (rs x); simpl in *; try congruence. - destruct (rs x0); try congruence. - assert (HB: negb (Int.eq i i0) = b) by congruence. - rewrite HB; destruct b; simpl; auto. } - (* Pbeqw, Pbnew, Pbltw, Pbtluw, Pbgew, Pbgeuw / Cmpu *) - 1-6: - destruct optR0 as [[]|]; - unfold apply_bin_r0, apply_bin_r0_r0r0lbl in *; - unfold zero32, Op.zero32 in *; - eexists; eexists; eauto; split; constructor; - simpl in *; try rewrite EVAL'; auto. - (* Pbeql / Cmpl *) - { destruct optR0 as [[]|]; - unfold apply_bin_r0, apply_bin_r0_r0r0lbl in *; - unfold zero64, Op.zero64 in *; - eexists; eexists; eauto; split; constructor; - simpl in *; auto. - + destruct (rs x); simpl in *; try congruence. - assert (HB: (Int64.eq Int64.zero i) = b) by congruence. - rewrite HB; destruct b; simpl; auto. - + destruct (rs x); simpl in *; try congruence. - assert (HB: (Int64.eq i Int64.zero) = b) by congruence. - rewrite HB; destruct b; simpl; auto. - + destruct (rs x); simpl in *; try congruence. - destruct (rs x0); try congruence. - assert (HB: (Int64.eq i i0) = b) by congruence. - rewrite HB; destruct b; simpl; auto. } - (* Pbnel / Cmpl *) - { destruct optR0 as [[]|]; - unfold apply_bin_r0, apply_bin_r0_r0r0lbl in *; - unfold zero64, Op.zero64 in *; - eexists; eexists; eauto; split; constructor; - simpl in *; auto. - + destruct (rs x); simpl in *; try congruence. - assert (HB: negb (Int64.eq Int64.zero i) = b) by congruence. - rewrite HB; destruct b; simpl; auto. - + destruct (rs x); simpl in *; try congruence. - assert (HB: negb (Int64.eq i Int64.zero) = b) by congruence. - rewrite HB; destruct b; simpl; auto. - + destruct (rs x); simpl in *; try congruence. - destruct (rs x0); try congruence. - assert (HB: negb (Int64.eq i i0) = b) by congruence. - rewrite HB; destruct b; simpl; auto. } - (* Pbeql, Pbnel, Pbltl, Pbtlul, Pbgel, Pbgeul / Cmplu *) - 1-6: - destruct optR0 as [[]|]; - unfold apply_bin_r0, apply_bin_r0_r0r0lbl in *; - unfold zero64, Op.zero64 in *; - eexists; eexists; eauto; split; constructor; - simpl in *; try rewrite EVAL'; auto. -Qed. - -Lemma transl_cbranch_correct_true: - forall cond args lbl k c m ms sp rs m', - transl_cbranch cond args lbl k = OK c -> - eval_condition cond (List.map ms args) m = Some true -> - agree ms sp rs -> - Mem.extends m m' -> - exists rs', exists insn, - exec_straight_opt ge fn c rs m' (insn :: k) rs' m' - /\ exec_instr ge fn insn rs' m' = goto_label fn lbl rs' m' - /\ forall r, r <> PC -> r <> X31 -> rs'#r = rs#r. -Proof. - intros. eapply transl_cbranch_correct_1 with (b := true); eauto. -Qed. - -Lemma transl_cbranch_correct_false: - forall cond args lbl k c m ms sp rs m', - transl_cbranch cond args lbl k = OK c -> - eval_condition cond (List.map ms args) m = Some false -> - agree ms sp rs -> - Mem.extends m m' -> - exists rs', - exec_straight ge fn c rs m' k rs' m' - /\ forall r, r <> PC -> r <> X31 -> rs'#r = rs#r. -Proof. - intros. exploit transl_cbranch_correct_1; eauto. simpl. - intros (rs' & insn & A & B & C). - exists (nextinstr rs'). - split. eapply exec_straight_opt_right; eauto. apply exec_straight_one; auto. - intros; Simpl. -Qed. - -<<<<<<< HEAD -======= -(** Translation of condition operators *) - -Lemma transl_cond_int32s_correct: - forall cmp rd r1 r2 k rs m, - exists rs', - exec_straight ge fn (transl_cond_int32s cmp rd r1 r2 k) rs m k rs' m - /\ Val.lessdef (Val.cmp cmp rs##r1 rs##r2) rs'#rd - /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r. -Proof. - intros. destruct cmp; simpl. -- econstructor; split. apply exec_straight_one; [simpl; eauto|auto]. - split; intros; Simpl. destruct (rs##r1); auto. destruct (rs##r2); auto. -- econstructor; split. apply exec_straight_one; [simpl; eauto|auto]. - split; intros; Simpl. destruct (rs##r1); auto. destruct (rs##r2); auto. -- econstructor; split. apply exec_straight_one; [simpl; eauto|auto]. - split; intros; Simpl. -- econstructor; split. - eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto. - split; intros; Simpl. unfold Val.cmp. rewrite <- Val.swap_cmp_bool. - simpl. rewrite (Val.negate_cmp_bool Clt). - destruct (Val.cmp_bool Clt rs##r2 rs##r1) as [[]|]; auto. -- econstructor; split. apply exec_straight_one; [simpl; eauto|auto]. - split; intros; Simpl. unfold Val.cmp. rewrite <- Val.swap_cmp_bool. auto. -- econstructor; split. - eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto. - split; intros; Simpl. unfold Val.cmp. rewrite (Val.negate_cmp_bool Clt). - destruct (Val.cmp_bool Clt rs##r1 rs##r2) as [[]|]; auto. -Qed. - -Lemma transl_cond_int32u_correct: - forall cmp rd r1 r2 k rs m, - exists rs', - exec_straight ge fn (transl_cond_int32u cmp rd r1 r2 k) rs m k rs' m - /\ rs'#rd = Val.cmpu (Mem.valid_pointer m) cmp rs##r1 rs##r2 - /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r. -Proof. - intros. destruct cmp; simpl. -- econstructor; split. apply exec_straight_one; [simpl; eauto|auto]. - split; intros; Simpl. -- econstructor; split. apply exec_straight_one; [simpl; eauto|auto]. - split; intros; Simpl. -- econstructor; split. apply exec_straight_one; [simpl; eauto|auto]. - split; intros; Simpl. -- econstructor; split. - eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto. - split; intros; Simpl. unfold Val.cmpu. rewrite <- Val.swap_cmpu_bool. - simpl. rewrite (Val.negate_cmpu_bool (Mem.valid_pointer m) Cle). - destruct (Val.cmpu_bool (Mem.valid_pointer m) Cle rs##r1 rs##r2) as [[]|]; auto. -- econstructor; split. apply exec_straight_one; [simpl; eauto|auto]. - split; intros; Simpl. unfold Val.cmpu. rewrite <- Val.swap_cmpu_bool. auto. -- econstructor; split. - eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto. - split; intros; Simpl. unfold Val.cmpu. rewrite (Val.negate_cmpu_bool (Mem.valid_pointer m) Clt). - destruct (Val.cmpu_bool (Mem.valid_pointer m) Clt rs##r1 rs##r2) as [[]|]; auto. -Qed. - -Lemma transl_cond_int64s_correct: - forall cmp rd r1 r2 k rs m, - exists rs', - exec_straight ge fn (transl_cond_int64s cmp rd r1 r2 k) rs m k rs' m - /\ Val.lessdef (Val.maketotal (Val.cmpl cmp rs###r1 rs###r2)) rs'#rd - /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r. -Proof. - intros. destruct cmp; simpl. -- econstructor; split. apply exec_straight_one; [simpl; eauto|auto]. - split; intros; Simpl. destruct (rs###r1); auto. destruct (rs###r2); auto. -- econstructor; split. apply exec_straight_one; [simpl; eauto|auto]. - split; intros; Simpl. destruct (rs###r1); auto. destruct (rs###r2); auto. -- econstructor; split. apply exec_straight_one; [simpl; eauto|auto]. - split; intros; Simpl. -- econstructor; split. - eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto. - split; intros; Simpl. unfold Val.cmpl. rewrite <- Val.swap_cmpl_bool. - simpl. rewrite (Val.negate_cmpl_bool Clt). - destruct (Val.cmpl_bool Clt rs###r2 rs###r1) as [[]|]; auto. -- econstructor; split. apply exec_straight_one; [simpl; eauto|auto]. - split; intros; Simpl. unfold Val.cmpl. rewrite <- Val.swap_cmpl_bool. auto. -- econstructor; split. - eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto. - split; intros; Simpl. unfold Val.cmpl. rewrite (Val.negate_cmpl_bool Clt). - destruct (Val.cmpl_bool Clt rs###r1 rs###r2) as [[]|]; auto. -Qed. - -Lemma transl_cond_int64u_correct: - forall cmp rd r1 r2 k rs m, - exists rs', - exec_straight ge fn (transl_cond_int64u cmp rd r1 r2 k) rs m k rs' m - /\ rs'#rd = Val.maketotal (Val.cmplu (Mem.valid_pointer m) cmp rs###r1 rs###r2) - /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r. -Proof. - intros. destruct cmp; simpl. -- econstructor; split. apply exec_straight_one; [simpl; eauto|auto]. - split; intros; Simpl. -- econstructor; split. apply exec_straight_one; [simpl; eauto|auto]. - split; intros; Simpl. -- econstructor; split. apply exec_straight_one; [simpl; eauto|auto]. - split; intros; Simpl. -- econstructor; split. - eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto. - split; intros; Simpl. unfold Val.cmplu. rewrite <- Val.swap_cmplu_bool. - simpl. rewrite (Val.negate_cmplu_bool (Mem.valid_pointer m) Cle). - destruct (Val.cmplu_bool (Mem.valid_pointer m) Cle rs###r1 rs###r2) as [[]|]; auto. -- econstructor; split. apply exec_straight_one; [simpl; eauto|auto]. - split; intros; Simpl. unfold Val.cmplu. rewrite <- Val.swap_cmplu_bool. auto. -- econstructor; split. - eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto. - split; intros; Simpl. unfold Val.cmplu. rewrite (Val.negate_cmplu_bool (Mem.valid_pointer m) Clt). - destruct (Val.cmplu_bool (Mem.valid_pointer m) Clt rs###r1 rs###r2) as [[]|]; auto. -Qed. - -Lemma transl_condimm_int32s_correct: - forall cmp rd r1 n k rs m, - r1 <> X31 -> - exists rs', - exec_straight ge fn (transl_condimm_int32s cmp rd r1 n k) rs m k rs' m - /\ Val.lessdef (Val.cmp cmp rs#r1 (Vint n)) rs'#rd - /\ forall r, r <> PC -> r <> rd -> r <> X31 -> rs'#r = rs#r. -Proof. - intros. unfold transl_condimm_int32s. - predSpec Int.eq Int.eq_spec n Int.zero. -- subst n. exploit transl_cond_int32s_correct. intros (rs' & A & B & C). - exists rs'; eauto. -- assert (DFL: - exists rs', - exec_straight ge fn (loadimm32 X31 n (transl_cond_int32s cmp rd r1 X31 k)) rs m k rs' m - /\ Val.lessdef (Val.cmp cmp rs#r1 (Vint n)) rs'#rd - /\ forall r, r <> PC -> r <> rd -> r <> X31 -> rs'#r = rs#r). - { exploit loadimm32_correct; eauto. intros (rs1 & A1 & B1 & C1). - exploit transl_cond_int32s_correct; eauto. intros (rs2 & A2 & B2 & C2). - exists rs2; split. - eapply exec_straight_trans. eexact A1. eexact A2. - split. simpl in B2. rewrite B1, C1 in B2 by auto with asmgen. auto. - intros; transitivity (rs1 r); auto. } - destruct cmp. -+ unfold xorimm32. - exploit (opimm32_correct Pxorw Pxoriw Val.xor); eauto. intros (rs1 & A1 & B1 & C1). - exploit transl_cond_int32s_correct; eauto. intros (rs2 & A2 & B2 & C2). - exists rs2; split. - eapply exec_straight_trans. eexact A1. eexact A2. - split. simpl in B2; rewrite B1 in B2; simpl in B2. destruct (rs#r1); auto. - unfold Val.cmp in B2; simpl in B2; rewrite Int.xor_is_zero in B2. exact B2. - intros; transitivity (rs1 r); auto. -+ unfold xorimm32. - exploit (opimm32_correct Pxorw Pxoriw Val.xor); eauto. intros (rs1 & A1 & B1 & C1). - exploit transl_cond_int32s_correct; eauto. intros (rs2 & A2 & B2 & C2). - exists rs2; split. - eapply exec_straight_trans. eexact A1. eexact A2. - split. simpl in B2; rewrite B1 in B2; simpl in B2. destruct (rs#r1); auto. - unfold Val.cmp in B2; simpl in B2; rewrite Int.xor_is_zero in B2. exact B2. - intros; transitivity (rs1 r); auto. -+ exploit (opimm32_correct Psltw Psltiw (Val.cmp Clt)); eauto. intros (rs1 & A1 & B1 & C1). - exists rs1; split. eexact A1. split; auto. rewrite B1; auto. -+ predSpec Int.eq Int.eq_spec n (Int.repr Int.max_signed). -* subst n. exploit loadimm32_correct; eauto. intros (rs1 & A1 & B1 & C1). - exists rs1; split. eexact A1. split; auto. - unfold Val.cmp; destruct (rs#r1); simpl; auto. rewrite B1. - unfold Int.lt. rewrite zlt_false. auto. - change (Int.signed (Int.repr Int.max_signed)) with Int.max_signed. - generalize (Int.signed_range i); lia. -* exploit (opimm32_correct Psltw Psltiw (Val.cmp Clt)); eauto. intros (rs1 & A1 & B1 & C1). - exists rs1; split. eexact A1. split; auto. - rewrite B1. unfold Val.cmp; simpl; destruct (rs#r1); simpl; auto. - unfold Int.lt. replace (Int.signed (Int.add n Int.one)) with (Int.signed n + 1). - destruct (zlt (Int.signed n) (Int.signed i)). - rewrite zlt_false by lia. auto. - rewrite zlt_true by lia. auto. - rewrite Int.add_signed. symmetry; apply Int.signed_repr. - assert (Int.signed n <> Int.max_signed). - { red; intros E. elim H1. rewrite <- (Int.repr_signed n). rewrite E. auto. } - generalize (Int.signed_range n); lia. -+ apply DFL. -+ apply DFL. -Qed. - -Lemma transl_condimm_int32u_correct: - forall cmp rd r1 n k rs m, - r1 <> X31 -> - exists rs', - exec_straight ge fn (transl_condimm_int32u cmp rd r1 n k) rs m k rs' m - /\ Val.lessdef (Val.cmpu (Mem.valid_pointer m) cmp rs#r1 (Vint n)) rs'#rd - /\ forall r, r <> PC -> r <> rd -> r <> X31 -> rs'#r = rs#r. -Proof. - intros. unfold transl_condimm_int32u. - predSpec Int.eq Int.eq_spec n Int.zero. -- subst n. exploit transl_cond_int32u_correct. intros (rs' & A & B & C). - exists rs'; split. eexact A. split; auto. rewrite B; auto. -- assert (DFL: - exists rs', - exec_straight ge fn (loadimm32 X31 n (transl_cond_int32u cmp rd r1 X31 k)) rs m k rs' m - /\ Val.lessdef (Val.cmpu (Mem.valid_pointer m) cmp rs#r1 (Vint n)) rs'#rd - /\ forall r, r <> PC -> r <> rd -> r <> X31 -> rs'#r = rs#r). - { exploit loadimm32_correct; eauto. intros (rs1 & A1 & B1 & C1). - exploit transl_cond_int32u_correct; eauto. intros (rs2 & A2 & B2 & C2). - exists rs2; split. - eapply exec_straight_trans. eexact A1. eexact A2. - split. simpl in B2. rewrite B1, C1 in B2 by auto with asmgen. rewrite B2; auto. - intros; transitivity (rs1 r); auto. } - destruct cmp. -+ apply DFL. -+ apply DFL. -+ exploit (opimm32_correct Psltuw Psltiuw (Val.cmpu (Mem.valid_pointer m) Clt) m); eauto. - intros (rs1 & A1 & B1 & C1). - exists rs1; split. eexact A1. split; auto. rewrite B1; auto. -+ apply DFL. -+ apply DFL. -+ apply DFL. -Qed. - -Lemma transl_condimm_int64s_correct: - forall cmp rd r1 n k rs m, - r1 <> X31 -> - exists rs', - exec_straight ge fn (transl_condimm_int64s cmp rd r1 n k) rs m k rs' m - /\ Val.lessdef (Val.maketotal (Val.cmpl cmp rs#r1 (Vlong n))) rs'#rd - /\ forall r, r <> PC -> r <> rd -> r <> X31 -> rs'#r = rs#r. -Proof. - intros. unfold transl_condimm_int64s. - predSpec Int64.eq Int64.eq_spec n Int64.zero. -- subst n. exploit transl_cond_int64s_correct. intros (rs' & A & B & C). - exists rs'; eauto. -- assert (DFL: - exists rs', - exec_straight ge fn (loadimm64 X31 n (transl_cond_int64s cmp rd r1 X31 k)) rs m k rs' m - /\ Val.lessdef (Val.maketotal (Val.cmpl cmp rs#r1 (Vlong n))) rs'#rd - /\ forall r, r <> PC -> r <> rd -> r <> X31 -> rs'#r = rs#r). - { exploit loadimm64_correct; eauto. intros (rs1 & A1 & B1 & C1). - exploit transl_cond_int64s_correct; eauto. intros (rs2 & A2 & B2 & C2). - exists rs2; split. - eapply exec_straight_trans. eexact A1. eexact A2. - split. simpl in B2. rewrite B1, C1 in B2 by auto with asmgen. auto. - intros; transitivity (rs1 r); auto. } - destruct cmp. -+ unfold xorimm64. - exploit (opimm64_correct Pxorl Pxoril Val.xorl); eauto. intros (rs1 & A1 & B1 & C1). - exploit transl_cond_int64s_correct; eauto. intros (rs2 & A2 & B2 & C2). - exists rs2; split. - eapply exec_straight_trans. eexact A1. eexact A2. - split. simpl in B2; rewrite B1 in B2; simpl in B2. destruct (rs#r1); auto. - unfold Val.cmpl in B2; simpl in B2; rewrite Int64.xor_is_zero in B2. exact B2. - intros; transitivity (rs1 r); auto. -+ unfold xorimm64. - exploit (opimm64_correct Pxorl Pxoril Val.xorl); eauto. intros (rs1 & A1 & B1 & C1). - exploit transl_cond_int64s_correct; eauto. intros (rs2 & A2 & B2 & C2). - exists rs2; split. - eapply exec_straight_trans. eexact A1. eexact A2. - split. simpl in B2; rewrite B1 in B2; simpl in B2. destruct (rs#r1); auto. - unfold Val.cmpl in B2; simpl in B2; rewrite Int64.xor_is_zero in B2. exact B2. - intros; transitivity (rs1 r); auto. -+ exploit (opimm64_correct Psltl Psltil (fun v1 v2 => Val.maketotal (Val.cmpl Clt v1 v2))); eauto. intros (rs1 & A1 & B1 & C1). - exists rs1; split. eexact A1. split; auto. rewrite B1; auto. -+ predSpec Int64.eq Int64.eq_spec n (Int64.repr Int64.max_signed). -* subst n. exploit loadimm32_correct; eauto. intros (rs1 & A1 & B1 & C1). - exists rs1; split. eexact A1. split; auto. - unfold Val.cmpl; destruct (rs#r1); simpl; auto. rewrite B1. - unfold Int64.lt. rewrite zlt_false. auto. - change (Int64.signed (Int64.repr Int64.max_signed)) with Int64.max_signed. - generalize (Int64.signed_range i); lia. -* exploit (opimm64_correct Psltl Psltil (fun v1 v2 => Val.maketotal (Val.cmpl Clt v1 v2))); eauto. intros (rs1 & A1 & B1 & C1). - exists rs1; split. eexact A1. split; auto. - rewrite B1. unfold Val.cmpl; simpl; destruct (rs#r1); simpl; auto. - unfold Int64.lt. replace (Int64.signed (Int64.add n Int64.one)) with (Int64.signed n + 1). - destruct (zlt (Int64.signed n) (Int64.signed i)). - rewrite zlt_false by lia. auto. - rewrite zlt_true by lia. auto. - rewrite Int64.add_signed. symmetry; apply Int64.signed_repr. - assert (Int64.signed n <> Int64.max_signed). - { red; intros E. elim H1. rewrite <- (Int64.repr_signed n). rewrite E. auto. } - generalize (Int64.signed_range n); lia. -+ apply DFL. -+ apply DFL. -Qed. - -Lemma transl_condimm_int64u_correct: - forall cmp rd r1 n k rs m, - r1 <> X31 -> - exists rs', - exec_straight ge fn (transl_condimm_int64u cmp rd r1 n k) rs m k rs' m - /\ Val.lessdef (Val.maketotal (Val.cmplu (Mem.valid_pointer m) cmp rs#r1 (Vlong n))) rs'#rd - /\ forall r, r <> PC -> r <> rd -> r <> X31 -> rs'#r = rs#r. -Proof. - intros. unfold transl_condimm_int64u. - predSpec Int64.eq Int64.eq_spec n Int64.zero. -- subst n. exploit transl_cond_int64u_correct. intros (rs' & A & B & C). - exists rs'; split. eexact A. split; auto. rewrite B; auto. -- assert (DFL: - exists rs', - exec_straight ge fn (loadimm64 X31 n (transl_cond_int64u cmp rd r1 X31 k)) rs m k rs' m - /\ Val.lessdef (Val.maketotal (Val.cmplu (Mem.valid_pointer m) cmp rs#r1 (Vlong n))) rs'#rd - /\ forall r, r <> PC -> r <> rd -> r <> X31 -> rs'#r = rs#r). - { exploit loadimm64_correct; eauto. intros (rs1 & A1 & B1 & C1). - exploit transl_cond_int64u_correct; eauto. intros (rs2 & A2 & B2 & C2). - exists rs2; split. - eapply exec_straight_trans. eexact A1. eexact A2. - split. simpl in B2. rewrite B1, C1 in B2 by auto with asmgen. rewrite B2; auto. - intros; transitivity (rs1 r); auto. } - destruct cmp. -+ apply DFL. -+ apply DFL. -+ exploit (opimm64_correct Psltul Psltiul (fun v1 v2 => Val.maketotal (Val.cmplu (Mem.valid_pointer m) Clt v1 v2)) m); eauto. - intros (rs1 & A1 & B1 & C1). - exists rs1; split. eexact A1. split; auto. rewrite B1; auto. -+ apply DFL. -+ apply DFL. -+ apply DFL. -Qed. - -Lemma transl_cond_op_correct: - forall cond rd args k c rs m, - transl_cond_op cond rd args k = OK c -> - exists rs', - exec_straight ge fn c rs m k rs' m - /\ Val.lessdef (Val.of_optbool (eval_condition cond (map rs (map preg_of args)) m)) rs'#rd - /\ forall r, r <> PC -> r <> rd -> r <> X31 -> rs'#r = rs#r. -Proof. - assert (MKTOT: forall ob, Val.of_optbool ob = Val.maketotal (option_map Val.of_bool ob)). - { destruct ob as [[]|]; reflexivity. } - intros until m; intros TR. - destruct cond; simpl in TR; ArgsInv. -+ (* cmp *) - exploit transl_cond_int32s_correct; eauto. intros (rs' & A & B & C). exists rs'; eauto. -+ (* cmpu *) - exploit transl_cond_int32u_correct; eauto. intros (rs' & A & B & C). - exists rs'; repeat split; eauto. rewrite B; auto. -+ (* cmpimm *) - apply transl_condimm_int32s_correct; eauto with asmgen. -+ (* cmpuimm *) - apply transl_condimm_int32u_correct; eauto with asmgen. -+ (* cmpl *) - exploit transl_cond_int64s_correct; eauto. intros (rs' & A & B & C). - exists rs'; repeat split; eauto. rewrite MKTOT; eauto. -+ (* cmplu *) - exploit transl_cond_int64u_correct; eauto. intros (rs' & A & B & C). - exists rs'; repeat split; eauto. rewrite B, MKTOT; eauto. -+ (* cmplimm *) - exploit transl_condimm_int64s_correct; eauto. instantiate (1 := x); eauto with asmgen. - intros (rs' & A & B & C). - exists rs'; repeat split; eauto. rewrite MKTOT; eauto. -+ (* cmpluimm *) - exploit transl_condimm_int64u_correct; eauto. instantiate (1 := x); eauto with asmgen. - intros (rs' & A & B & C). - exists rs'; repeat split; eauto. rewrite MKTOT; eauto. -+ (* cmpf *) - destruct (transl_cond_float c0 rd x x0) as [insn normal] eqn:TR. - fold (Val.cmpf c0 (rs x) (rs x0)). - set (v := Val.cmpf c0 (rs x) (rs x0)). - destruct normal; inv EQ2. -* econstructor; split. - apply exec_straight_one. eapply transl_cond_float_correct with (v := v); eauto. auto. - split; intros; Simpl. -* econstructor; split. - eapply exec_straight_two. - eapply transl_cond_float_correct with (v := Val.notbool v); eauto. - simpl; reflexivity. - auto. auto. - split; intros; Simpl. unfold v, Val.cmpf. destruct (Val.cmpf_bool c0 (rs x) (rs x0)) as [[]|]; auto. -+ (* notcmpf *) - destruct (transl_cond_float c0 rd x x0) as [insn normal] eqn:TR. - rewrite Val.notbool_negb_3. fold (Val.cmpf c0 (rs x) (rs x0)). - set (v := Val.cmpf c0 (rs x) (rs x0)). - destruct normal; inv EQ2. -* econstructor; split. - eapply exec_straight_two. - eapply transl_cond_float_correct with (v := v); eauto. - simpl; reflexivity. - auto. auto. - split; intros; Simpl. unfold v, Val.cmpf. destruct (Val.cmpf_bool c0 (rs x) (rs x0)) as [[]|]; auto. -* econstructor; split. - apply exec_straight_one. eapply transl_cond_float_correct with (v := Val.notbool v); eauto. auto. - split; intros; Simpl. -+ (* cmpfs *) - destruct (transl_cond_single c0 rd x x0) as [insn normal] eqn:TR. - fold (Val.cmpfs c0 (rs x) (rs x0)). - set (v := Val.cmpfs c0 (rs x) (rs x0)). - destruct normal; inv EQ2. -* econstructor; split. - apply exec_straight_one. eapply transl_cond_single_correct with (v := v); eauto. auto. - split; intros; Simpl. -* econstructor; split. - eapply exec_straight_two. - eapply transl_cond_single_correct with (v := Val.notbool v); eauto. - simpl; reflexivity. - auto. auto. - split; intros; Simpl. unfold v, Val.cmpfs. destruct (Val.cmpfs_bool c0 (rs x) (rs x0)) as [[]|]; auto. -+ (* notcmpfs *) - destruct (transl_cond_single c0 rd x x0) as [insn normal] eqn:TR. - rewrite Val.notbool_negb_3. fold (Val.cmpfs c0 (rs x) (rs x0)). - set (v := Val.cmpfs c0 (rs x) (rs x0)). - destruct normal; inv EQ2. -* econstructor; split. - eapply exec_straight_two. - eapply transl_cond_single_correct with (v := v); eauto. - simpl; reflexivity. - auto. auto. - split; intros; Simpl. unfold v, Val.cmpfs. destruct (Val.cmpfs_bool c0 (rs x) (rs x0)) as [[]|]; auto. -* econstructor; split. - apply exec_straight_one. eapply transl_cond_single_correct with (v := Val.notbool v); eauto. auto. - split; intros; Simpl. -Qed. - ->>>>>>> master -(** Some arithmetic properties. *) - -Remark cast32unsigned_from_cast32signed: - forall i, Int64.repr (Int.unsigned i) = Int64.zero_ext 32 (Int64.repr (Int.signed i)). -Proof. - intros. apply Int64.same_bits_eq; intros. - rewrite Int64.bits_zero_ext, !Int64.testbit_repr by tauto. - rewrite Int.bits_signed by tauto. fold (Int.testbit i i0). - change Int.zwordsize with 32. - destruct (zlt i0 32). auto. apply Int.bits_above. auto. -Qed. - -(* Translation of arithmetic operations *) - -Ltac SimplEval H := - match type of H with - | Some _ = None _ => discriminate - | Some _ = Some _ => inv H - | ?a = Some ?b => let A := fresh in assert (A: Val.maketotal a = b) by (rewrite H; reflexivity) -end. - -Ltac TranslOpSimpl := - econstructor; split; - [ apply exec_straight_one; [simpl; eauto | reflexivity] - | split; [ apply Val.lessdef_same; Simpl; fail | intros; Simpl; fail ] ]. - -Lemma transl_op_correct: - forall op args res k (rs: regset) m v c, - transl_op op args res k = OK c -> - eval_operation ge (rs#SP) op (map rs (map preg_of args)) m = Some v -> - exists rs', - exec_straight ge fn c rs m k rs' m - /\ Val.lessdef v rs'#(preg_of res) - /\ forall r, data_preg r = true -> r <> preg_of res -> preg_notin r (destroyed_by_op op) -> rs' r = rs r. -Proof. - assert (SAME: forall v1 v2, v1 = v2 -> Val.lessdef v2 v1). { intros; subst; auto. } -Opaque Int.eq. - intros until c; intros TR EV. - unfold transl_op in TR; destruct op; ArgsInv; simpl in EV; SimplEval EV; try TranslOpSimpl. - (* move *) - { destruct (preg_of res), (preg_of m0); inv TR; TranslOpSimpl. } - (* intconst *) - { exploit loadimm32_correct; eauto. intros (rs' & A & B & C). - exists rs'; split; eauto. rewrite B; auto with asmgen. } - (* longconst *) - { exploit loadimm64_correct; eauto. intros (rs' & A & B & C). - exists rs'; split; eauto. rewrite B; auto with asmgen. } - (* floatconst *) - { destruct (Float.eq_dec n Float.zero). - + subst n. econstructor; split. - apply exec_straight_one. simpl; eauto. auto. - split; intros; Simpl. - + econstructor; split. - apply exec_straight_one. simpl; eauto. auto. - split; intros; Simpl. } - (* singleconst *) - { destruct (Float32.eq_dec n Float32.zero). - + subst n. econstructor; split. - apply exec_straight_one. simpl; eauto. auto. - split; intros; Simpl. - + econstructor; split. - apply exec_straight_one. simpl; eauto. auto. - split; intros; Simpl. } - (* addrsymbol *) - { destruct (Archi.pic_code tt && negb (Ptrofs.eq ofs Ptrofs.zero)). - + set (rs1 := nextinstr (rs#x <- (Genv.symbol_address ge id Ptrofs.zero))). - exploit (addptrofs_correct x x ofs k rs1 m); eauto with asmgen. - intros (rs2 & A & B & C). - exists rs2; split. - apply exec_straight_step with rs1 m; auto. - split. replace ofs with (Ptrofs.add Ptrofs.zero ofs) by (apply Ptrofs.add_zero_l). - rewrite Genv.shift_symbol_address. - replace (rs1 x) with (Genv.symbol_address ge id Ptrofs.zero) in B by (unfold rs1; Simpl). - exact B. - intros. rewrite C by eauto with asmgen. unfold rs1; Simpl. - + TranslOpSimpl. } - (* stackoffset *) - { exploit addptrofs_correct. instantiate (1 := X2); auto with asmgen. intros (rs' & A & B & C). - exists rs'; split; eauto. auto with asmgen. } - (* cast8signed *) - { econstructor; split. - eapply exec_straight_two. simpl;eauto. simpl;eauto. auto. auto. - split; intros; Simpl. - assert (A: Int.ltu (Int.repr 24) Int.iwordsize = true) by auto. - destruct (rs x0); auto; simpl. rewrite A; simpl. rewrite A. - apply Val.lessdef_same. f_equal. apply Int.sign_ext_shr_shl. split; reflexivity. } - (* cast16signed *) - { econstructor; split. - eapply exec_straight_two. simpl;eauto. simpl;eauto. auto. auto. - split; intros; Simpl. - assert (A: Int.ltu (Int.repr 16) Int.iwordsize = true) by auto. - destruct (rs x0); auto; simpl. rewrite A; simpl. rewrite A. - apply Val.lessdef_same. f_equal. apply Int.sign_ext_shr_shl. split; reflexivity. } - (* addimm *) - { exploit (opimm32_correct Paddw Paddiw Val.add); auto. instantiate (1 := x0); eauto with asmgen. - intros (rs' & A & B & C). - exists rs'; split; eauto. rewrite B; auto with asmgen. } - (* andimm *) - { exploit (opimm32_correct Pandw Pandiw Val.and); auto. instantiate (1 := x0); eauto with asmgen. - intros (rs' & A & B & C). - exists rs'; split; eauto. rewrite B; auto with asmgen. } - (* orimm *) - exploit (opimm32_correct Porw Poriw Val.or); auto. instantiate (1 := x0); eauto with asmgen. - { intros (rs' & A & B & C). - exists rs'; split; eauto. rewrite B; auto with asmgen. } - (* xorimm *) - { exploit (opimm32_correct Pxorw Pxoriw Val.xor); auto. instantiate (1 := x0); eauto with asmgen. - intros (rs' & A & B & C). - exists rs'; split; eauto. rewrite B; auto with asmgen. } - (* shrximm *) - { destruct (Val.shrx (rs x0) (Vint n)) eqn:TOTAL; cbn. - { - exploit Val.shrx_shr_3; eauto. intros E; subst v. - destruct (Int.eq n Int.zero). - + econstructor; split. apply exec_straight_one. simpl; eauto. auto. - split; intros; Simpl. - + destruct (Int.eq n Int.one). - * econstructor; split. - eapply exec_straight_step. simpl; reflexivity. auto. - eapply exec_straight_step. simpl; reflexivity. auto. - apply exec_straight_one. simpl; reflexivity. auto. - split; intros; Simpl. - * change (Int.repr 32) with Int.iwordsize. set (n' := Int.sub Int.iwordsize n). - econstructor; split. - eapply exec_straight_step. simpl; reflexivity. auto. - eapply exec_straight_step. simpl; reflexivity. auto. - eapply exec_straight_step. simpl; reflexivity. auto. - apply exec_straight_one. simpl; reflexivity. auto. - split; intros; Simpl. - } - destruct (Int.eq n Int.zero). - + econstructor; split. apply exec_straight_one. simpl; eauto. auto. - split; intros; Simpl. - + destruct (Int.eq n Int.one). - * econstructor; split. - eapply exec_straight_step. simpl; reflexivity. auto. - eapply exec_straight_step. simpl; reflexivity. auto. - apply exec_straight_one. simpl; reflexivity. auto. - split; intros; Simpl. - * change (Int.repr 32) with Int.iwordsize. set (n' := Int.sub Int.iwordsize n). - econstructor; split. - eapply exec_straight_step. simpl; reflexivity. auto. - eapply exec_straight_step. simpl; reflexivity. auto. - eapply exec_straight_step. simpl; reflexivity. auto. - apply exec_straight_one. simpl; reflexivity. auto. - split; intros; Simpl. } - (* longofintu *) - { econstructor; split. - eapply exec_straight_three. simpl; eauto. simpl; eauto. simpl; eauto. auto. auto. auto. - split; intros; Simpl. destruct (rs x0); auto. simpl. - assert (A: Int.ltu (Int.repr 32) Int64.iwordsize' = true) by auto. - rewrite A; simpl. rewrite A. apply Val.lessdef_same. f_equal. - rewrite cast32unsigned_from_cast32signed. apply Int64.zero_ext_shru_shl. compute; auto. } - (* addlimm *) - { exploit (opimm64_correct Paddl Paddil Val.addl); auto. instantiate (1 := x0); eauto with asmgen. - intros (rs' & A & B & C). - exists rs'; split; eauto. rewrite B; auto with asmgen. } - (* andimm *) - { exploit (opimm64_correct Pandl Pandil Val.andl); auto. instantiate (1 := x0); eauto with asmgen. - intros (rs' & A & B & C). - exists rs'; split; eauto. rewrite B; auto with asmgen. } - (* orimm *) - { exploit (opimm64_correct Porl Poril Val.orl); auto. instantiate (1 := x0); eauto with asmgen. - intros (rs' & A & B & C). - exists rs'; split; eauto. rewrite B; auto with asmgen. } - (* xorimm *) - { exploit (opimm64_correct Pxorl Pxoril Val.xorl); auto. instantiate (1 := x0); eauto with asmgen. - intros (rs' & A & B & C). - exists rs'; split; eauto. rewrite B; auto with asmgen. } - (* shrxlimm *) - { destruct (Val.shrxl (rs x0) (Vint n)) eqn:TOTAL. - { - exploit Val.shrxl_shrl_3; eauto. intros E; subst v. - destruct (Int.eq n Int.zero). - + econstructor; split. apply exec_straight_one. simpl; eauto. auto. - split; intros; Simpl. - + destruct (Int.eq n Int.one). - * econstructor; split. - eapply exec_straight_step. simpl; reflexivity. auto. - eapply exec_straight_step. simpl; reflexivity. auto. - apply exec_straight_one. simpl; reflexivity. auto. - split; intros; Simpl. - - * change (Int.repr 64) with Int64.iwordsize'. set (n' := Int.sub Int64.iwordsize' n). - econstructor; split. - eapply exec_straight_step. simpl; reflexivity. auto. - eapply exec_straight_step. simpl; reflexivity. auto. - eapply exec_straight_step. simpl; reflexivity. auto. - apply exec_straight_one. simpl; reflexivity. auto. - split; intros; Simpl. - } - destruct (Int.eq n Int.zero). - + econstructor; split. apply exec_straight_one. simpl; eauto. auto. - split; intros; Simpl. - + destruct (Int.eq n Int.one). - * econstructor; split. - eapply exec_straight_step. simpl; reflexivity. auto. - eapply exec_straight_step. simpl; reflexivity. auto. - apply exec_straight_one. simpl; reflexivity. auto. - split; intros; Simpl. - - * change (Int.repr 64) with Int64.iwordsize'. set (n' := Int.sub Int64.iwordsize' n). - econstructor; split. - eapply exec_straight_step. simpl; reflexivity. auto. - eapply exec_straight_step. simpl; reflexivity. auto. - eapply exec_straight_step. simpl; reflexivity. auto. - apply exec_straight_one. simpl; reflexivity. auto. - split; intros; Simpl. } - (* Expanded instructions from RTL *) - 7,8,15,16: - econstructor; split; try apply exec_straight_one; simpl; eauto; - split; intros; Simpl; unfold may_undef_int; try destruct is_long; simpl; - try rewrite Int.add_commut; try rewrite Int64.add_commut; - destruct (rs (preg_of m0)); try discriminate; eauto. - 1-12: - destruct optR0 as [[]|]; unfold apply_bin_r0_r0r0, apply_bin_r0; - econstructor; split; try apply exec_straight_one; simpl; eauto; - split; intros; Simpl; - destruct (rs x0); auto; - destruct (rs x1); auto. - (* select *) - { econstructor; split. apply exec_straight_one. simpl; eauto. auto. - split; intros; Simpl. - apply Val.lessdef_normalize. } -Qed. - -(** Memory accesses *) - -Lemma indexed_memory_access_correct: - forall mk_instr base ofs k rs m, - base <> X31 -> - exists base' ofs' rs', - exec_straight_opt ge fn (indexed_memory_access mk_instr base ofs k) rs m - (mk_instr base' ofs' :: k) rs' m - /\ Val.offset_ptr rs'#base' (eval_offset ge ofs') = Val.offset_ptr rs#base ofs - /\ forall r, r <> PC -> r <> X31 -> rs'#r = rs#r. -Proof. - unfold indexed_memory_access; intros. - destruct Archi.ptr64 eqn:SF. -- generalize (make_immed64_sound (Ptrofs.to_int64 ofs)); intros EQ. - destruct (make_immed64 (Ptrofs.to_int64 ofs)). -+ econstructor; econstructor; econstructor; split. - apply exec_straight_opt_refl. - split; auto. simpl. subst imm. rewrite Ptrofs.of_int64_to_int64 by auto. auto. -+ econstructor; econstructor; econstructor; split. - constructor. eapply exec_straight_two. - simpl; eauto. simpl; eauto. auto. auto. - split; intros; Simpl. destruct (rs base); auto; simpl. rewrite SF. simpl. - rewrite Ptrofs.add_assoc. f_equal. f_equal. - rewrite <- (Ptrofs.of_int64_to_int64 SF ofs). rewrite EQ. - symmetry; auto with ptrofs. -+ econstructor; econstructor; econstructor; split. - constructor. eapply exec_straight_two. - simpl; eauto. simpl; eauto. auto. auto. - split; intros; Simpl. unfold eval_offset. destruct (rs base); auto; simpl. rewrite SF. simpl. - rewrite Ptrofs.add_zero. subst imm. rewrite Ptrofs.of_int64_to_int64 by auto. auto. -- generalize (make_immed32_sound (Ptrofs.to_int ofs)); intros EQ. - destruct (make_immed32 (Ptrofs.to_int ofs)). -+ econstructor; econstructor; econstructor; split. - apply exec_straight_opt_refl. - split; auto. simpl. subst imm. rewrite Ptrofs.of_int_to_int by auto. auto. -+ econstructor; econstructor; econstructor; split. - constructor. eapply exec_straight_two. - simpl; eauto. simpl; eauto. auto. auto. - split; intros; Simpl. destruct (rs base); auto; simpl. rewrite SF. simpl. - rewrite Ptrofs.add_assoc. f_equal. f_equal. - rewrite <- (Ptrofs.of_int_to_int SF ofs). rewrite EQ. - symmetry; auto with ptrofs. -Qed. - -Lemma indexed_load_access_correct: - forall chunk (mk_instr: ireg -> offset -> instruction) rd m, - (forall base ofs rs, - exec_instr ge fn (mk_instr base ofs) rs m = exec_load ge chunk rs m rd base ofs) -> - forall (base: ireg) ofs k (rs: regset) v, - Mem.loadv chunk m (Val.offset_ptr rs#base ofs) = Some v -> - base <> X31 -> rd <> PC -> - exists rs', - exec_straight ge fn (indexed_memory_access mk_instr base ofs k) rs m k rs' m - /\ rs'#rd = v - /\ forall r, r <> PC -> r <> X31 -> r <> rd -> rs'#r = rs#r. -Proof. - intros until m; intros EXEC; intros until v; intros LOAD NOT31 NOTPC. - exploit indexed_memory_access_correct; eauto. - intros (base' & ofs' & rs' & A & B & C). - econstructor; split. - eapply exec_straight_opt_right. eexact A. apply exec_straight_one. rewrite EXEC. - unfold exec_load. rewrite B, LOAD. eauto. Simpl. - split; intros; Simpl. -Qed. - -Lemma indexed_store_access_correct: - forall chunk (mk_instr: ireg -> offset -> instruction) r1 m, - (forall base ofs rs, - exec_instr ge fn (mk_instr base ofs) rs m = exec_store ge chunk rs m r1 base ofs) -> - forall (base: ireg) ofs k (rs: regset) m', - Mem.storev chunk m (Val.offset_ptr rs#base ofs) (rs#r1) = Some m' -> - base <> X31 -> r1 <> X31 -> r1 <> PC -> - exists rs', - exec_straight ge fn (indexed_memory_access mk_instr base ofs k) rs m k rs' m' - /\ forall r, r <> PC -> r <> X31 -> rs'#r = rs#r. -Proof. - intros until m; intros EXEC; intros until m'; intros STORE NOT31 NOT31' NOTPC. - exploit indexed_memory_access_correct; eauto. - intros (base' & ofs' & rs' & A & B & C). - econstructor; split. - eapply exec_straight_opt_right. eexact A. apply exec_straight_one. rewrite EXEC. - unfold exec_store. rewrite B, C, STORE by auto. eauto. auto. - intros; Simpl. -Qed. - -Lemma loadind_correct: - forall (base: ireg) ofs ty dst k c (rs: regset) m v, - loadind base ofs ty dst k = OK c -> - Mem.loadv (chunk_of_type ty) m (Val.offset_ptr rs#base ofs) = Some v -> - base <> X31 -> - exists rs', - exec_straight ge fn c rs m k rs' m - /\ rs'#(preg_of dst) = v - /\ forall r, r <> PC -> r <> X31 -> r <> preg_of dst -> rs'#r = rs#r. -Proof. - intros until v; intros TR LOAD NOT31. - assert (A: exists mk_instr, - c = indexed_memory_access mk_instr base ofs k - /\ forall base' ofs' rs', - exec_instr ge fn (mk_instr base' ofs') rs' m = - exec_load ge (chunk_of_type ty) rs' m (preg_of dst) base' ofs'). - { unfold loadind in TR. destruct ty, (preg_of dst); inv TR; econstructor; split; eauto. } - destruct A as (mk_instr & B & C). subst c. - eapply indexed_load_access_correct; eauto with asmgen. -Qed. - -Lemma storeind_correct: - forall (base: ireg) ofs ty src k c (rs: regset) m m', - storeind src base ofs ty k = OK c -> - Mem.storev (chunk_of_type ty) m (Val.offset_ptr rs#base ofs) rs#(preg_of src) = Some m' -> - base <> X31 -> - exists rs', - exec_straight ge fn c rs m k rs' m' - /\ forall r, r <> PC -> r <> X31 -> rs'#r = rs#r. -Proof. - intros until m'; intros TR STORE NOT31. - assert (A: exists mk_instr, - c = indexed_memory_access mk_instr base ofs k - /\ forall base' ofs' rs', - exec_instr ge fn (mk_instr base' ofs') rs' m = - exec_store ge (chunk_of_type ty) rs' m (preg_of src) base' ofs'). - { unfold storeind in TR. destruct ty, (preg_of src); inv TR; econstructor; split; eauto. } - destruct A as (mk_instr & B & C). subst c. - eapply indexed_store_access_correct; eauto with asmgen. -Qed. - -Lemma loadind_ptr_correct: - forall (base: ireg) ofs (dst: ireg) k (rs: regset) m v, - Mem.loadv Mptr m (Val.offset_ptr rs#base ofs) = Some v -> - base <> X31 -> - exists rs', - exec_straight ge fn (loadind_ptr base ofs dst k) rs m k rs' m - /\ rs'#dst = v - /\ forall r, r <> PC -> r <> X31 -> r <> dst -> rs'#r = rs#r. -Proof. - intros. eapply indexed_load_access_correct; eauto with asmgen. - intros. unfold Mptr. destruct Archi.ptr64; auto. -Qed. - -Lemma storeind_ptr_correct: - forall (base: ireg) ofs (src: ireg) k (rs: regset) m m', - Mem.storev Mptr m (Val.offset_ptr rs#base ofs) rs#src = Some m' -> - base <> X31 -> src <> X31 -> - exists rs', - exec_straight ge fn (storeind_ptr src base ofs k) rs m k rs' m' - /\ forall r, r <> PC -> r <> X31 -> rs'#r = rs#r. -Proof. - intros. eapply indexed_store_access_correct with (r1 := src); eauto with asmgen. - intros. unfold Mptr. destruct Archi.ptr64; auto. -Qed. - -Lemma transl_memory_access_correct: - forall mk_instr addr args k c (rs: regset) m v, - transl_memory_access mk_instr addr args k = OK c -> - eval_addressing ge rs#SP addr (map rs (map preg_of args)) = Some v -> - exists base ofs rs', - exec_straight_opt ge fn c rs m (mk_instr base ofs :: k) rs' m - /\ Val.offset_ptr rs'#base (eval_offset ge ofs) = v - /\ forall r, r <> PC -> r <> X31 -> rs'#r = rs#r. -Proof. - intros until v; intros TR EV. - unfold transl_memory_access in TR; destruct addr; ArgsInv. -- (* indexed *) - inv EV. apply indexed_memory_access_correct; eauto with asmgen. -- (* global *) - simpl in EV. inv EV. inv TR. econstructor; econstructor; econstructor; split. - constructor. apply exec_straight_one. simpl; eauto. auto. - split; intros; Simpl. unfold eval_offset. apply low_high_half. -- (* stack *) - inv TR. inv EV. apply indexed_memory_access_correct; eauto with asmgen. -Qed. - -Lemma transl_load_access_correct: - forall chunk (mk_instr: ireg -> offset -> instruction) addr args k c rd (rs: regset) m v v', - (forall base ofs rs, - exec_instr ge fn (mk_instr base ofs) rs m = exec_load ge chunk rs m rd base ofs) -> - transl_memory_access mk_instr addr args k = OK c -> - eval_addressing ge rs#SP addr (map rs (map preg_of args)) = Some v -> - Mem.loadv chunk m v = Some v' -> - rd <> PC -> - exists rs', - exec_straight ge fn c rs m k rs' m - /\ rs'#rd = v' - /\ forall r, r <> PC -> r <> X31 -> r <> rd -> rs'#r = rs#r. -Proof. - intros until v'; intros INSTR TR EV LOAD NOTPC. - exploit transl_memory_access_correct; eauto. - intros (base & ofs & rs' & A & B & C). - econstructor; split. - eapply exec_straight_opt_right. eexact A. apply exec_straight_one. - rewrite INSTR. unfold exec_load. rewrite B, LOAD. reflexivity. Simpl. - split; intros; Simpl. -Qed. - -Lemma transl_store_access_correct: - forall chunk (mk_instr: ireg -> offset -> instruction) addr args k c r1 (rs: regset) m v m', - (forall base ofs rs, - exec_instr ge fn (mk_instr base ofs) rs m = exec_store ge chunk rs m r1 base ofs) -> - transl_memory_access mk_instr addr args k = OK c -> - eval_addressing ge rs#SP addr (map rs (map preg_of args)) = Some v -> - Mem.storev chunk m v rs#r1 = Some m' -> - r1 <> PC -> r1 <> X31 -> - exists rs', - exec_straight ge fn c rs m k rs' m' - /\ forall r, r <> PC -> r <> X31 -> rs'#r = rs#r. -Proof. - intros until m'; intros INSTR TR EV STORE NOTPC NOT31. - exploit transl_memory_access_correct; eauto. - intros (base & ofs & rs' & A & B & C). - econstructor; split. - eapply exec_straight_opt_right. eexact A. apply exec_straight_one. - rewrite INSTR. unfold exec_store. rewrite B, C, STORE by auto. reflexivity. auto. - intros; Simpl. -Qed. - -Lemma transl_load_correct: - forall trap chunk addr args dst k c (rs: regset) m a v, - transl_load trap chunk addr args dst k = OK c -> - eval_addressing ge rs#SP addr (map rs (map preg_of args)) = Some a -> - Mem.loadv chunk m a = Some v -> - exists rs', - exec_straight ge fn c rs m k rs' m - /\ rs'#(preg_of dst) = v - /\ forall r, r <> PC -> r <> X31 -> r <> preg_of dst -> rs'#r = rs#r. -Proof. - intros until v; intros TR EV LOAD. - destruct trap; try (simpl in *; discriminate). - assert (A: exists mk_instr, - transl_memory_access mk_instr addr args k = OK c - /\ forall base ofs rs, - exec_instr ge fn (mk_instr base ofs) rs m = exec_load ge chunk rs m (preg_of dst) base ofs). - { unfold transl_load in TR; destruct chunk; ArgsInv; econstructor; (split; [eassumption|auto]). } - destruct A as (mk_instr & B & C). - eapply transl_load_access_correct; eauto with asmgen. -Qed. - -Lemma transl_store_correct: - forall chunk addr args src k c (rs: regset) m a m', - transl_store chunk addr args src k = OK c -> - eval_addressing ge rs#SP addr (map rs (map preg_of args)) = Some a -> - Mem.storev chunk m a rs#(preg_of src) = Some m' -> - exists rs', - exec_straight ge fn c rs m k rs' m' - /\ forall r, r <> PC -> r <> X31 -> rs'#r = rs#r. -Proof. - intros until m'; intros TR EV STORE. - assert (A: exists mk_instr chunk', - transl_memory_access mk_instr addr args k = OK c - /\ (forall base ofs rs, - exec_instr ge fn (mk_instr base ofs) rs m = exec_store ge chunk' rs m (preg_of src) base ofs) - /\ Mem.storev chunk m a rs#(preg_of src) = Mem.storev chunk' m a rs#(preg_of src)). - { unfold transl_store in TR; destruct chunk; ArgsInv; - (econstructor; econstructor; split; [eassumption | split; [ intros; simpl; reflexivity | auto]]). - destruct a; auto. apply Mem.store_signed_unsigned_8. - destruct a; auto. apply Mem.store_signed_unsigned_16. - } - destruct A as (mk_instr & chunk' & B & C & D). - rewrite D in STORE; clear D. - eapply transl_store_access_correct; eauto with asmgen. -Qed. - -(** Function epilogues *) - -Lemma make_epilogue_correct: - forall ge0 f m stk soff cs m' ms rs k tm, - load_stack m (Vptr stk soff) Tptr f.(fn_link_ofs) = Some (parent_sp cs) -> - load_stack m (Vptr stk soff) Tptr f.(fn_retaddr_ofs) = Some (parent_ra cs) -> - Mem.free m stk 0 f.(fn_stacksize) = Some m' -> - agree ms (Vptr stk soff) rs -> - Mem.extends m tm -> - match_stack ge0 cs -> - exists rs', exists tm', - exec_straight ge fn (make_epilogue f k) rs tm k rs' tm' - /\ agree ms (parent_sp cs) rs' - /\ Mem.extends m' tm' - /\ rs'#RA = parent_ra cs - /\ rs'#SP = parent_sp cs - /\ (forall r, r <> PC -> r <> RA -> r <> SP -> r <> X31 -> rs'#r = rs#r). -Proof. - intros until tm; intros LP LRA FREE AG MEXT MCS. - exploit Mem.loadv_extends. eauto. eexact LP. auto. simpl. intros (parent' & LP' & LDP'). - exploit Mem.loadv_extends. eauto. eexact LRA. auto. simpl. intros (ra' & LRA' & LDRA'). - exploit lessdef_parent_sp; eauto. intros EQ; subst parent'; clear LDP'. - exploit lessdef_parent_ra; eauto. intros EQ; subst ra'; clear LDRA'. - exploit Mem.free_parallel_extends; eauto. intros (tm' & FREE' & MEXT'). - unfold make_epilogue. - rewrite chunk_of_Tptr in *. - exploit (loadind_ptr_correct SP (fn_retaddr_ofs f) RA (Pfreeframe (fn_stacksize f) (fn_link_ofs f) :: k) rs tm). - rewrite <- (sp_val _ _ _ AG). simpl. eexact LRA'. congruence. - intros (rs1 & A1 & B1 & C1). - econstructor; econstructor; split. - eapply exec_straight_trans. eexact A1. apply exec_straight_one. simpl. - rewrite (C1 X2) by auto with asmgen. rewrite <- (sp_val _ _ _ AG). simpl; rewrite LP'. - rewrite FREE'. eauto. auto. - split. apply agree_nextinstr. apply agree_set_other; auto with asmgen. - apply agree_change_sp with (Vptr stk soff). - apply agree_exten with rs; auto. intros; apply C1; auto with asmgen. - eapply parent_sp_def; eauto. - split. auto. - split. Simpl. - split. Simpl. - intros. Simpl. -Qed. - -End CONSTRUCTORS. diff --git a/riscV/TO_MERGE/SelectLongproof.v b/riscV/TO_MERGE/SelectLongproof.v deleted file mode 100644 index 954dd134..00000000 --- a/riscV/TO_MERGE/SelectLongproof.v +++ /dev/null @@ -1,650 +0,0 @@ -(* *********************************************************************) -(* *) -(* The Compcert verified compiler *) -(* *) -(* Xavier Leroy, INRIA Paris *) -(* Prashanth Mundkur, SRI International *) -(* *) -(* 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. *) -(* *) -(* The contributions by Prashanth Mundkur are reused and adapted *) -(* under the terms of a Contributor License Agreement between *) -(* SRI International and INRIA. *) -(* *) -(* *********************************************************************) - -(** Correctness of instruction selection for 64-bit integer operations *) - -Require Import String Coqlib Maps Integers Floats Errors. -Require Archi. -Require Import AST Values Memory Globalenvs Events. -Require Import Cminor Op CminorSel. -Require Import OpHelpers OpHelpersproof. -Require Import SelectOp SelectOpproof SplitLong SplitLongproof. -Require Import SelectLong. - -Local Open Scope cminorsel_scope. -Local Open Scope string_scope. - -(** * Correctness of the instruction selection functions for 64-bit operators *) - -Section CMCONSTR. - -Variable prog: program. -Variable hf: helper_functions. -Hypothesis HELPERS: helper_functions_declared prog hf. -Let ge := Genv.globalenv prog. -Variable sp: val. -Variable e: env. -Variable m: mem. - -Definition unary_constructor_sound (cstr: expr -> expr) (sem: val -> val) : Prop := - forall le a x, - eval_expr ge sp e m le a x -> - exists v, eval_expr ge sp e m le (cstr a) v /\ Val.lessdef (sem x) v. - -Definition binary_constructor_sound (cstr: expr -> expr -> expr) (sem: val -> val -> val) : Prop := - forall le a x b y, - eval_expr ge sp e m le a x -> - eval_expr ge sp e m le b y -> - exists v, eval_expr ge sp e m le (cstr a b) v /\ Val.lessdef (sem x y) v. - -Definition partial_unary_constructor_sound (cstr: expr -> expr) (sem: val -> option val) : Prop := - forall le a x y, - eval_expr ge sp e m le a x -> - sem x = Some y -> - exists v, eval_expr ge sp e m le (cstr a) v /\ Val.lessdef y v. - -Definition partial_binary_constructor_sound (cstr: expr -> expr -> expr) (sem: val -> val -> option val) : Prop := - forall le a x b y z, - eval_expr ge sp e m le a x -> - eval_expr ge sp e m le b y -> - sem x y = Some z -> - exists v, eval_expr ge sp e m le (cstr a b) v /\ Val.lessdef z v. - -Theorem eval_longconst: - forall le n, eval_expr ge sp e m le (longconst n) (Vlong n). -Proof. - unfold longconst; intros; destruct Archi.splitlong. - apply SplitLongproof.eval_longconst. - EvalOp. -Qed. - -Lemma is_longconst_sound: - forall v a n le, - is_longconst a = Some n -> eval_expr ge sp e m le a v -> v = Vlong n. -Proof with (try discriminate). - intros. unfold is_longconst in *. destruct Archi.splitlong. - eapply SplitLongproof.is_longconst_sound; eauto. - assert (a = Eop (Olongconst n) Enil). - { destruct a... destruct o... destruct e0... congruence. } - subst a. InvEval. auto. -Qed. - -Theorem eval_intoflong: unary_constructor_sound intoflong Val.loword. -Proof. - unfold intoflong; destruct Archi.splitlong. apply SplitLongproof.eval_intoflong. - red; intros. destruct (is_longconst a) as [n|] eqn:C. -- TrivialExists. simpl. erewrite (is_longconst_sound x) by eauto. auto. -- TrivialExists. -Qed. - -Theorem eval_longofintu: unary_constructor_sound longofintu Val.longofintu. -Proof. - unfold longofintu; destruct Archi.splitlong. apply SplitLongproof.eval_longofintu. - red; intros. destruct (is_intconst a) as [n|] eqn:C. -- econstructor; split. apply eval_longconst. - exploit is_intconst_sound; eauto. intros; subst x. auto. -- TrivialExists. -Qed. - -Theorem eval_longofint: unary_constructor_sound longofint Val.longofint. -Proof. - unfold longofint; destruct Archi.splitlong. apply SplitLongproof.eval_longofint. - red; intros. destruct (is_intconst a) as [n|] eqn:C. -- econstructor; split. apply eval_longconst. - exploit is_intconst_sound; eauto. intros; subst x. auto. -- TrivialExists. -Qed. - -Theorem eval_negl: unary_constructor_sound negl Val.negl. -Proof. - unfold negl. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_negl; auto. - red; intros. destruct (is_longconst a) as [n|] eqn:C. -- exploit is_longconst_sound; eauto. intros EQ; subst x. - econstructor; split. apply eval_longconst. auto. -- TrivialExists. -Qed. - -Theorem eval_addlimm: forall n, unary_constructor_sound (addlimm n) (fun v => Val.addl v (Vlong n)). -Proof. - unfold addlimm; intros; red; intros. - predSpec Int64.eq Int64.eq_spec n Int64.zero. - subst. exists x; split; auto. - destruct x; simpl; rewrite ?Int64.add_zero, ?Ptrofs.add_zero; auto. - destruct Archi.ptr64; auto. - destruct (addlimm_match a); InvEval. -- econstructor; split. apply eval_longconst. rewrite Int64.add_commut; auto. -- econstructor; split. EvalOp. simpl; eauto. - unfold Genv.symbol_address. destruct (Genv.find_symbol ge s); simpl; auto. - destruct Archi.ptr64; auto. rewrite Ptrofs.add_commut; auto. -- econstructor; split. EvalOp. simpl; eauto. - destruct sp; simpl; auto. destruct Archi.ptr64; auto. - rewrite Ptrofs.add_assoc, (Ptrofs.add_commut m0). auto. -- subst x. rewrite Val.addl_assoc. rewrite Int64.add_commut. TrivialExists. -- TrivialExists. -Qed. - -Theorem eval_addl: binary_constructor_sound addl Val.addl. -Proof. - unfold addl. destruct Archi.splitlong eqn:SL. - apply SplitLongproof.eval_addl. apply Archi.splitlong_ptr32; auto. -(* - assert (SF: Archi.ptr64 = true). - { Local Transparent Archi.splitlong. unfold Archi.splitlong in SL. - destruct Archi.ptr64; simpl in *; congruence. } -*) -(* - assert (B: forall id ofs n, - Genv.symbol_address ge id (Ptrofs.add ofs (Ptrofs.repr n)) = - Val.addl (Genv.symbol_address ge id ofs) (Vlong (Int64.repr n))). - { intros. replace (Ptrofs.repr n) with (Ptrofs.of_int64 (Int64.repr n)) by auto with ptrofs. - apply Genv.shift_symbol_address_64; auto. } - -*) - red; intros until y. - case (addl_match a b); intros; InvEval. - - rewrite Val.addl_commut. apply eval_addlimm; auto. - - apply eval_addlimm; auto. - - subst. - replace (Val.addl (Val.addl v1 (Vlong n1)) (Val.addl v0 (Vlong n2))) - with (Val.addl (Val.addl v1 v0) (Val.addl (Vlong n1) (Vlong n2))). - apply eval_addlimm. EvalOp. - repeat rewrite Val.addl_assoc. decEq. apply Val.addl_permut. - - subst. econstructor; split. - EvalOp. constructor. EvalOp. simpl; eauto. constructor. eauto. constructor. simpl; eauto. - rewrite Val.addl_commut. destruct sp; simpl; auto. - destruct v1; simpl; auto. - destruct Archi.ptr64 eqn:SF; auto. - apply Val.lessdef_same. f_equal. rewrite ! Ptrofs.add_assoc. f_equal. - rewrite (Ptrofs.add_commut (Ptrofs.of_int64 n1)), Ptrofs.add_assoc. f_equal. auto with ptrofs. - destruct Archi.ptr64 eqn:SF; auto. - - subst. econstructor; split. - EvalOp. constructor. EvalOp. simpl; eauto. constructor. eauto. constructor. simpl; eauto. - destruct sp; simpl; auto. - destruct v1; simpl; auto. - destruct Archi.ptr64 eqn:SF; auto. - apply Val.lessdef_same. f_equal. rewrite ! Ptrofs.add_assoc. f_equal. f_equal. - rewrite Ptrofs.add_commut. auto with ptrofs. - destruct Archi.ptr64 eqn:SF; auto. - - subst. - replace (Val.addl (Val.addl v1 (Vlong n1)) y) - with (Val.addl (Val.addl v1 y) (Vlong n1)). - apply eval_addlimm. EvalOp. - repeat rewrite Val.addl_assoc. decEq. apply Val.addl_commut. - - subst. - replace (Val.addl x (Val.addl v1 (Vlong n2))) - with (Val.addl (Val.addl x v1) (Vlong n2)). - apply eval_addlimm. EvalOp. - repeat rewrite Val.addl_assoc. reflexivity. - - TrivialExists. -Qed. - -Theorem eval_subl: binary_constructor_sound subl Val.subl. -Proof. - unfold subl. destruct Archi.splitlong eqn:SL. - apply SplitLongproof.eval_subl. apply Archi.splitlong_ptr32; auto. - red; intros; destruct (subl_match a b); InvEval. -- rewrite Val.subl_addl_opp. apply eval_addlimm; auto. -- subst. rewrite Val.subl_addl_l. rewrite Val.subl_addl_r. - rewrite Val.addl_assoc. simpl. rewrite Int64.add_commut. rewrite <- Int64.sub_add_opp. - apply eval_addlimm; EvalOp. -- subst. rewrite Val.subl_addl_l. apply eval_addlimm; EvalOp. -- subst. rewrite Val.subl_addl_r. - apply eval_addlimm; EvalOp. -- TrivialExists. -Qed. - -Theorem eval_shllimm: forall n, unary_constructor_sound (fun e => shllimm e n) (fun v => Val.shll v (Vint n)). -Proof. - intros; unfold shllimm. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_shllimm; auto. - red; intros. - predSpec Int.eq Int.eq_spec n Int.zero. - exists x; split; auto. subst n; destruct x; simpl; auto. - destruct (Int.ltu Int.zero Int64.iwordsize'); auto. - change (Int64.shl' i Int.zero) with (Int64.shl i Int64.zero). rewrite Int64.shl_zero; auto. - destruct (Int.ltu n Int64.iwordsize') eqn:LT; simpl. - assert (DEFAULT: exists v, eval_expr ge sp e m le (Eop (Oshllimm n) (a:::Enil)) v - /\ Val.lessdef (Val.shll x (Vint n)) v) by TrivialExists. - destruct (shllimm_match a); InvEval. -- econstructor; split. apply eval_longconst. simpl; rewrite LT; auto. -- destruct (Int.ltu (Int.add n n1) Int64.iwordsize') eqn:LT'; auto. - subst. econstructor; split. EvalOp. simpl; eauto. - destruct v1; simpl; auto. rewrite LT'. - destruct (Int.ltu n1 Int64.iwordsize') eqn:LT1; auto. - simpl; rewrite LT. rewrite Int.add_commut, Int64.shl'_shl'; auto. rewrite Int.add_commut; auto. -- apply DEFAULT. -- TrivialExists. constructor; eauto. constructor. EvalOp. simpl; eauto. constructor. auto. -Qed. - -Theorem eval_shrluimm: forall n, unary_constructor_sound (fun e => shrluimm e n) (fun v => Val.shrlu v (Vint n)). -Proof. - intros; unfold shrluimm. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_shrluimm; auto. - red; intros. - predSpec Int.eq Int.eq_spec n Int.zero. - exists x; split; auto. subst n; destruct x; simpl; auto. - destruct (Int.ltu Int.zero Int64.iwordsize'); auto. - change (Int64.shru' i Int.zero) with (Int64.shru i Int64.zero). rewrite Int64.shru_zero; auto. - destruct (Int.ltu n Int64.iwordsize') eqn:LT; simpl. - assert (DEFAULT: exists v, eval_expr ge sp e m le (Eop (Oshrluimm n) (a:::Enil)) v - /\ Val.lessdef (Val.shrlu x (Vint n)) v) by TrivialExists. - destruct (shrluimm_match a); InvEval. -- econstructor; split. apply eval_longconst. simpl; rewrite LT; auto. -- destruct (Int.ltu (Int.add n n1) Int64.iwordsize') eqn:LT'; auto. - subst. econstructor; split. EvalOp. simpl; eauto. - destruct v1; simpl; auto. rewrite LT'. - destruct (Int.ltu n1 Int64.iwordsize') eqn:LT1; auto. - simpl; rewrite LT. rewrite Int.add_commut, Int64.shru'_shru'; auto. rewrite Int.add_commut; auto. -- apply DEFAULT. -- TrivialExists. constructor; eauto. constructor. EvalOp. simpl; eauto. constructor. auto. -Qed. - -Theorem eval_shrlimm: forall n, unary_constructor_sound (fun e => shrlimm e n) (fun v => Val.shrl v (Vint n)). -Proof. - intros; unfold shrlimm. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_shrlimm; auto. - red; intros. - predSpec Int.eq Int.eq_spec n Int.zero. - exists x; split; auto. subst n; destruct x; simpl; auto. - destruct (Int.ltu Int.zero Int64.iwordsize'); auto. - change (Int64.shr' i Int.zero) with (Int64.shr i Int64.zero). rewrite Int64.shr_zero; auto. - destruct (Int.ltu n Int64.iwordsize') eqn:LT; simpl. - assert (DEFAULT: exists v, eval_expr ge sp e m le (Eop (Oshrlimm n) (a:::Enil)) v - /\ Val.lessdef (Val.shrl x (Vint n)) v) by TrivialExists. - destruct (shrlimm_match a); InvEval. -- econstructor; split. apply eval_longconst. simpl; rewrite LT; auto. -- destruct (Int.ltu (Int.add n n1) Int64.iwordsize') eqn:LT'; auto. - subst. econstructor; split. EvalOp. simpl; eauto. - destruct v1; simpl; auto. rewrite LT'. - destruct (Int.ltu n1 Int64.iwordsize') eqn:LT1; auto. - simpl; rewrite LT. rewrite Int.add_commut, Int64.shr'_shr'; auto. rewrite Int.add_commut; auto. -- apply DEFAULT. -- TrivialExists. constructor; eauto. constructor. EvalOp. simpl; eauto. constructor. auto. -Qed. - -Theorem eval_shll: binary_constructor_sound shll Val.shll. -Proof. - unfold shll. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_shll; auto. - red; intros. destruct (is_intconst b) as [n2|] eqn:C. -- exploit is_intconst_sound; eauto. intros EQ; subst y. apply eval_shllimm; auto. -- TrivialExists. -Qed. - -Theorem eval_shrlu: binary_constructor_sound shrlu Val.shrlu. -Proof. - unfold shrlu. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_shrlu; auto. - red; intros. destruct (is_intconst b) as [n2|] eqn:C. -- exploit is_intconst_sound; eauto. intros EQ; subst y. apply eval_shrluimm; auto. -- TrivialExists. -Qed. - -Theorem eval_shrl: binary_constructor_sound shrl Val.shrl. -Proof. - unfold shrl. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_shrl; auto. - red; intros. destruct (is_intconst b) as [n2|] eqn:C. -- exploit is_intconst_sound; eauto. intros EQ; subst y. apply eval_shrlimm; auto. -- TrivialExists. -Qed. - -Theorem eval_mullimm_base: forall n, unary_constructor_sound (mullimm_base n) (fun v => Val.mull v (Vlong n)). -Proof. - intros; unfold mullimm_base. red; intros. - assert (DEFAULT: exists v, - eval_expr ge sp e m le (Eop Omull (a ::: longconst n ::: Enil)) v - /\ Val.lessdef (Val.mull x (Vlong n)) v). - { econstructor; split. EvalOp. constructor. eauto. constructor. apply eval_longconst. constructor. simpl; eauto. - auto. } - generalize (Int64.one_bits'_decomp n); intros D. - destruct (Int64.one_bits' n) as [ | i [ | j [ | ? ? ]]] eqn:B. -- apply DEFAULT. -- replace (Val.mull x (Vlong n)) with (Val.shll x (Vint i)). - apply eval_shllimm; auto. - simpl in D. rewrite D, Int64.add_zero. destruct x; simpl; auto. - rewrite (Int64.one_bits'_range n) by (rewrite B; auto with coqlib). - rewrite Int64.shl'_mul; auto. -- set (le' := x :: le). - assert (A0: eval_expr ge sp e m le' (Eletvar O) x) by (constructor; reflexivity). - exploit (eval_shllimm i). eexact A0. intros (v1 & A1 & B1). - exploit (eval_shllimm j). eexact A0. intros (v2 & A2 & B2). - exploit (eval_addl). eexact A1. eexact A2. intros (v3 & A3 & B3). - exists v3; split. econstructor; eauto. - rewrite D. simpl. rewrite Int64.add_zero. destruct x; auto. - simpl in *. - rewrite (Int64.one_bits'_range n) in B1 by (rewrite B; auto with coqlib). - rewrite (Int64.one_bits'_range n) in B2 by (rewrite B; auto with coqlib). - inv B1; inv B2. simpl in B3; inv B3. - rewrite Int64.mul_add_distr_r. rewrite <- ! Int64.shl'_mul. auto. -- apply DEFAULT. -Qed. - -Theorem eval_mullimm: forall n, unary_constructor_sound (mullimm n) (fun v => Val.mull v (Vlong n)). -Proof. - unfold mullimm. intros; red; intros. - destruct Archi.splitlong eqn:SL. - eapply SplitLongproof.eval_mullimm; eauto. - predSpec Int64.eq Int64.eq_spec n Int64.zero. - exists (Vlong Int64.zero); split. apply eval_longconst. - destruct x; simpl; auto. subst n; rewrite Int64.mul_zero; auto. - predSpec Int64.eq Int64.eq_spec n Int64.one. - exists x; split; auto. - destruct x; simpl; auto. subst n; rewrite Int64.mul_one; auto. - destruct (mullimm_match a); InvEval. -- econstructor; split. apply eval_longconst. rewrite Int64.mul_commut; auto. -- exploit (eval_mullimm_base n); eauto. intros (v2 & A2 & B2). - exploit (eval_addlimm (Int64.mul n n2)). eexact A2. intros (v3 & A3 & B3). - exists v3; split; auto. - subst x. destruct v1; simpl; auto. - simpl in B2; inv B2. simpl in B3; inv B3. rewrite Int64.mul_add_distr_l. - rewrite (Int64.mul_commut n). auto. - destruct Archi.ptr64; simpl; auto. -- apply eval_mullimm_base; auto. -Qed. - -Theorem eval_mull: binary_constructor_sound mull Val.mull. -Proof. - unfold mull. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_mull; auto. - red; intros; destruct (mull_match a b); InvEval. -- rewrite Val.mull_commut. apply eval_mullimm; auto. -- apply eval_mullimm; auto. -- TrivialExists. -Qed. - -Theorem eval_mullhu: - forall n, unary_constructor_sound (fun a => mullhu a n) (fun v => Val.mullhu v (Vlong n)). -Proof. - unfold mullhu; intros. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_mullhu; auto. - red; intros. TrivialExists. constructor. eauto. constructor. apply eval_longconst. constructor. auto. -Qed. - -Theorem eval_mullhs: - forall n, unary_constructor_sound (fun a => mullhs a n) (fun v => Val.mullhs v (Vlong n)). -Proof. - unfold mullhs; intros. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_mullhs; auto. - red; intros. TrivialExists. constructor. eauto. constructor. apply eval_longconst. constructor. auto. -Qed. - -Theorem eval_andlimm: forall n, unary_constructor_sound (andlimm n) (fun v => Val.andl v (Vlong n)). -Proof. - unfold andlimm; intros; red; intros. - predSpec Int64.eq Int64.eq_spec n Int64.zero. - exists (Vlong Int64.zero); split. apply eval_longconst. - subst. destruct x; simpl; auto. rewrite Int64.and_zero; auto. - predSpec Int64.eq Int64.eq_spec n Int64.mone. - exists x; split. assumption. - subst. destruct x; simpl; auto. rewrite Int64.and_mone; auto. - destruct (andlimm_match a); InvEval; subst. -- econstructor; split. apply eval_longconst. simpl. rewrite Int64.and_commut; auto. -- TrivialExists. simpl. rewrite Val.andl_assoc. rewrite Int64.and_commut; auto. -- TrivialExists. -Qed. - -Theorem eval_andl: binary_constructor_sound andl Val.andl. -Proof. - unfold andl; destruct Archi.splitlong. apply SplitLongproof.eval_andl. - red; intros. destruct (andl_match a b). -- InvEval. rewrite Val.andl_commut. apply eval_andlimm; auto. -- InvEval. apply eval_andlimm; auto. -- TrivialExists. -Qed. - -Theorem eval_orlimm: forall n, unary_constructor_sound (orlimm n) (fun v => Val.orl v (Vlong n)). -Proof. - unfold orlimm; intros; red; intros. - predSpec Int64.eq Int64.eq_spec n Int64.zero. - exists x; split; auto. subst. destruct x; simpl; auto. rewrite Int64.or_zero; auto. - predSpec Int64.eq Int64.eq_spec n Int64.mone. - econstructor; split. apply eval_longconst. subst. destruct x; simpl; auto. rewrite Int64.or_mone; auto. - destruct (orlimm_match a); InvEval; subst. -- econstructor; split. apply eval_longconst. simpl. rewrite Int64.or_commut; auto. -- TrivialExists. simpl. rewrite Val.orl_assoc. rewrite Int64.or_commut; auto. -- TrivialExists. -Qed. - -Theorem eval_orl: binary_constructor_sound orl Val.orl. -Proof. - unfold orl; destruct Archi.splitlong. apply SplitLongproof.eval_orl. - red; intros. - destruct (orl_match a b). -- InvEval. rewrite Val.orl_commut. apply eval_orlimm; auto. -- InvEval. apply eval_orlimm; auto. -- TrivialExists. -Qed. - -Theorem eval_xorlimm: forall n, unary_constructor_sound (xorlimm n) (fun v => Val.xorl v (Vlong n)). -Proof. - unfold xorlimm; intros; red; intros. - predSpec Int64.eq Int64.eq_spec n Int64.zero. - exists x; split; auto. subst. destruct x; simpl; auto. rewrite Int64.xor_zero; auto. - destruct (xorlimm_match a); InvEval; subst. -- econstructor; split. apply eval_longconst. simpl. rewrite Int64.xor_commut; auto. -- rewrite Val.xorl_assoc. simpl. rewrite (Int64.xor_commut n2). - predSpec Int64.eq Int64.eq_spec (Int64.xor n n2) Int64.zero. -+ rewrite H. exists v1; split; auto. destruct v1; simpl; auto. rewrite Int64.xor_zero; auto. -+ TrivialExists. -- TrivialExists. -Qed. - -Theorem eval_xorl: binary_constructor_sound xorl Val.xorl. -Proof. - unfold xorl; destruct Archi.splitlong. apply SplitLongproof.eval_xorl. - red; intros. destruct (xorl_match a b). -- InvEval. rewrite Val.xorl_commut. apply eval_xorlimm; auto. -- InvEval. apply eval_xorlimm; auto. -- TrivialExists. -Qed. - -Theorem eval_notl: unary_constructor_sound notl Val.notl. -Proof. - unfold notl; destruct Archi.splitlong. apply SplitLongproof.eval_notl. - red; intros. rewrite Val.notl_xorl. apply eval_xorlimm; auto. -Qed. - -Theorem eval_divls_base: partial_binary_constructor_sound divls_base Val.divls. -Proof. - unfold divls_base; red; intros. destruct Archi.splitlong eqn:SL. - eapply SplitLongproof.eval_divls_base; eauto. - TrivialExists. - cbn. - rewrite H1. - cbn. - trivial. -Qed. - -Theorem eval_modls_base: partial_binary_constructor_sound modls_base Val.modls. -Proof. - unfold modls_base; red; intros. destruct Archi.splitlong eqn:SL. - eapply SplitLongproof.eval_modls_base; eauto. - TrivialExists. - cbn. - rewrite H1. - cbn. - trivial. -Qed. - -Theorem eval_divlu_base: partial_binary_constructor_sound divlu_base Val.divlu. -Proof. - unfold divlu_base; red; intros. destruct Archi.splitlong eqn:SL. - eapply SplitLongproof.eval_divlu_base; eauto. - TrivialExists. - cbn. - rewrite H1. - cbn. - trivial. -Qed. - -Theorem eval_modlu_base: partial_binary_constructor_sound modlu_base Val.modlu. -Proof. - unfold modlu_base; red; intros. destruct Archi.splitlong eqn:SL. - eapply SplitLongproof.eval_modlu_base; eauto. - TrivialExists. - cbn. - rewrite H1. - cbn. - trivial. -Qed. - -Theorem eval_shrxlimm: - forall le a n x z, - eval_expr ge sp e m le a x -> - Val.shrxl x (Vint n) = Some z -> - exists v, eval_expr ge sp e m le (shrxlimm a n) v /\ Val.lessdef z v. -Proof. - unfold shrxlimm; intros. destruct Archi.splitlong eqn:SL. -+ eapply SplitLongproof.eval_shrxlimm; eauto using Archi.splitlong_ptr32. -+ predSpec Int.eq Int.eq_spec n Int.zero. -- subst n. destruct x; simpl in H0; inv H0. econstructor; split; eauto. - change (Int.ltu Int.zero (Int.repr 63)) with true. simpl. rewrite Int64.shrx'_zero; auto. -- TrivialExists. -<<<<<<< HEAD - cbn. - rewrite H0. - reflexivity. -======= -(* - intros. unfold shrxlimm. destruct Archi.splitlong eqn:SL. -+ eapply SplitLongproof.eval_shrxlimm; eauto using Archi.splitlong_ptr32. -+ destruct x; simpl in H0; try discriminate. - destruct (Int.ltu n (Int.repr 63)) eqn:LTU; inv H0. - predSpec Int.eq Int.eq_spec n Int.zero. - - subst n. exists (Vlong i); split; auto. rewrite Int64.shrx'_zero. auto. - - assert (NZ: Int.unsigned n <> 0). - { intro EQ; elim H0. rewrite <- (Int.repr_unsigned n). rewrite EQ; auto. } - assert (LT: 0 <= Int.unsigned n < 63) by (apply Int.ltu_inv in LTU; assumption). - assert (LTU2: Int.ltu (Int.sub Int64.iwordsize' n) Int64.iwordsize' = true). - { unfold Int.ltu; apply zlt_true. - unfold Int.sub. change (Int.unsigned Int64.iwordsize') with 64. - rewrite Int.unsigned_repr. lia. - assert (64 < Int.max_unsigned) by reflexivity. lia. } - assert (X: eval_expr ge sp e m le - (Eop (Oshrlimm (Int.repr (Int64.zwordsize - 1))) (a ::: Enil)) - (Vlong (Int64.shr' i (Int.repr (Int64.zwordsize - 1))))). - { EvalOp. } - assert (Y: eval_expr ge sp e m le (shrxlimm_inner a n) - (Vlong (Int64.shru' (Int64.shr' i (Int.repr (Int64.zwordsize - 1))) (Int.sub Int64.iwordsize' n)))). - { EvalOp. simpl. rewrite LTU2. auto. } - TrivialExists. - constructor. EvalOp. simpl; eauto. constructor. - simpl. unfold Int.ltu; rewrite zlt_true. rewrite Int64.shrx'_shr_2 by auto. reflexivity. - change (Int.unsigned Int64.iwordsize') with 64; lia. -*) ->>>>>>> master -Qed. - -Theorem eval_cmplu: - forall c le a x b y v, - eval_expr ge sp e m le a x -> - eval_expr ge sp e m le b y -> - Val.cmplu (Mem.valid_pointer m) c x y = Some v -> - eval_expr ge sp e m le (cmplu c a b) v. -Proof. - unfold cmplu; intros. destruct Archi.splitlong eqn:SL. - eapply SplitLongproof.eval_cmplu; eauto using Archi.splitlong_ptr32. - unfold Val.cmplu in H1. - destruct (Val.cmplu_bool (Mem.valid_pointer m) c x y) as [vb|] eqn:C; simpl in H1; inv H1. - destruct (is_longconst a) as [n1|] eqn:LC1; destruct (is_longconst b) as [n2|] eqn:LC2; - try (assert (x = Vlong n1) by (eapply is_longconst_sound; eauto)); - try (assert (y = Vlong n2) by (eapply is_longconst_sound; eauto)); - subst. -- simpl in C; inv C. EvalOp. destruct (Int64.cmpu c n1 n2); reflexivity. -- EvalOp. simpl. rewrite Val.swap_cmplu_bool. rewrite C; auto. -- EvalOp. simpl; rewrite C; auto. -- EvalOp. simpl; rewrite C; auto. -Qed. - -Theorem eval_cmpl: - forall c le a x b y v, - eval_expr ge sp e m le a x -> - eval_expr ge sp e m le b y -> - Val.cmpl c x y = Some v -> - eval_expr ge sp e m le (cmpl c a b) v. -Proof. - unfold cmpl; intros. destruct Archi.splitlong eqn:SL. - eapply SplitLongproof.eval_cmpl; eauto. - unfold Val.cmpl in H1. - destruct (Val.cmpl_bool c x y) as [vb|] eqn:C; simpl in H1; inv H1. - destruct (is_longconst a) as [n1|] eqn:LC1; destruct (is_longconst b) as [n2|] eqn:LC2; - try (assert (x = Vlong n1) by (eapply is_longconst_sound; eauto)); - try (assert (y = Vlong n2) by (eapply is_longconst_sound; eauto)); - subst. -- simpl in C; inv C. EvalOp. destruct (Int64.cmp c n1 n2); reflexivity. -- EvalOp. simpl. rewrite Val.swap_cmpl_bool. rewrite C; auto. -- EvalOp. simpl; rewrite C; auto. -- EvalOp. simpl; rewrite C; auto. -Qed. - -Theorem eval_longoffloat: partial_unary_constructor_sound longoffloat Val.longoffloat. -Proof. - unfold longoffloat; red; intros. destruct Archi.splitlong eqn:SL. - eapply SplitLongproof.eval_longoffloat; eauto. - TrivialExists. - cbn; rewrite H0; reflexivity. -Qed. - -Theorem eval_longuoffloat: partial_unary_constructor_sound longuoffloat Val.longuoffloat. -Proof. - unfold longuoffloat; red; intros. destruct Archi.splitlong eqn:SL. - eapply SplitLongproof.eval_longuoffloat; eauto. - TrivialExists. - cbn; rewrite H0; reflexivity. -Qed. - -Theorem eval_floatoflong: partial_unary_constructor_sound floatoflong Val.floatoflong. -Proof. - unfold floatoflong; red; intros. destruct Archi.splitlong eqn:SL. - eapply SplitLongproof.eval_floatoflong; eauto. - TrivialExists. - cbn; rewrite H0; reflexivity. -Qed. - -Theorem eval_floatoflongu: partial_unary_constructor_sound floatoflongu Val.floatoflongu. -Proof. - unfold floatoflongu; red; intros. destruct Archi.splitlong eqn:SL. - eapply SplitLongproof.eval_floatoflongu; eauto. - TrivialExists. - cbn; rewrite H0; reflexivity. -Qed. - -Theorem eval_longofsingle: partial_unary_constructor_sound longofsingle Val.longofsingle. -Proof. - unfold longofsingle; red; intros. destruct Archi.splitlong eqn:SL. - eapply SplitLongproof.eval_longofsingle; eauto. - TrivialExists. - cbn; rewrite H0; reflexivity. -Qed. - -Theorem eval_longuofsingle: partial_unary_constructor_sound longuofsingle Val.longuofsingle. -Proof. - unfold longuofsingle; red; intros. destruct Archi.splitlong eqn:SL. - eapply SplitLongproof.eval_longuofsingle; eauto. - TrivialExists. - cbn; rewrite H0; reflexivity. -Qed. - -Theorem eval_singleoflong: partial_unary_constructor_sound singleoflong Val.singleoflong. -Proof. - unfold singleoflong; red; intros. destruct Archi.splitlong eqn:SL. - eapply SplitLongproof.eval_singleoflong; eauto. - TrivialExists. - cbn; rewrite H0; reflexivity. -Qed. - -Theorem eval_singleoflongu: partial_unary_constructor_sound singleoflongu Val.singleoflongu. -Proof. - unfold singleoflongu; red; intros. destruct Archi.splitlong eqn:SL. - eapply SplitLongproof.eval_singleoflongu; eauto. - TrivialExists. - cbn; rewrite H0; reflexivity. -Qed. - -End CMCONSTR. diff --git a/riscV/TO_MERGE/SelectOpproof.v b/riscV/TO_MERGE/SelectOpproof.v deleted file mode 100644 index 9bd66213..00000000 --- a/riscV/TO_MERGE/SelectOpproof.v +++ /dev/null @@ -1,1124 +0,0 @@ -(* *********************************************************************) -(* *) -(* The Compcert verified compiler *) -(* *) -(* Xavier Leroy, INRIA Paris-Rocquencourt *) -(* Prashanth Mundkur, SRI International *) -(* *) -(* 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. *) -(* *) -(* The contributions by Prashanth Mundkur are reused and adapted *) -(* under the terms of a Contributor License Agreement between *) -(* SRI International and INRIA. *) -(* *) -(* *********************************************************************) - -(** Correctness of instruction selection for operators *) - -Require Import Coqlib Zbits. -Require Import AST Integers Floats. -Require Import Values Memory Builtins Globalenvs. -Require Import Cminor Op CminorSel. -Require Import SelectOp. -Require Import OpHelpers. -Require Import OpHelpersproof. -Require Import Lia. - -Local Open Scope cminorsel_scope. - -(** * Useful lemmas and tactics *) - -(** The following are trivial lemmas and custom tactics that help - perform backward (inversion) and forward reasoning over the evaluation - of operator applications. *) - -Ltac EvalOp := eapply eval_Eop; eauto with evalexpr. - -Ltac InvEval1 := - match goal with - | [ H: (eval_expr _ _ _ _ _ (Eop _ Enil) _) |- _ ] => - inv H; InvEval1 - | [ H: (eval_expr _ _ _ _ _ (Eop _ (_ ::: Enil)) _) |- _ ] => - inv H; InvEval1 - | [ H: (eval_expr _ _ _ _ _ (Eop _ (_ ::: _ ::: Enil)) _) |- _ ] => - inv H; InvEval1 - | [ H: (eval_exprlist _ _ _ _ _ Enil _) |- _ ] => - inv H; InvEval1 - | [ H: (eval_exprlist _ _ _ _ _ (_ ::: _) _) |- _ ] => - inv H; InvEval1 - | _ => - idtac - end. - -Ltac InvEval2 := - match goal with - | [ H: (eval_operation _ _ _ nil _ = Some _) |- _ ] => - simpl in H; inv H - | [ H: (eval_operation _ _ _ (_ :: nil) _ = Some _) |- _ ] => - simpl in H; FuncInv - | [ H: (eval_operation _ _ _ (_ :: _ :: nil) _ = Some _) |- _ ] => - simpl in H; FuncInv - | [ H: (eval_operation _ _ _ (_ :: _ :: _ :: nil) _ = Some _) |- _ ] => - simpl in H; FuncInv - | _ => - idtac - end. - -Ltac InvEval := InvEval1; InvEval2; InvEval2. - -Ltac TrivialExists := - match goal with - | [ |- exists v, _ /\ Val.lessdef ?a v ] => exists a; split; [EvalOp | auto] - end. - -(** * Correctness of the smart constructors *) - -Section CMCONSTR. -Variable prog: program. -Variable hf: helper_functions. -Hypothesis HELPERS: helper_functions_declared prog hf. -Let ge := Genv.globalenv prog. -Variable sp: val. -Variable e: env. -Variable m: mem. - -(** We now show that the code generated by "smart constructor" functions - such as [Selection.notint] behaves as expected. Continuing the - [notint] example, we show that if the expression [e] - evaluates to some integer value [Vint n], then [Selection.notint e] - evaluates to a value [Vint (Int.not n)] which is indeed the integer - negation of the value of [e]. - - All proofs follow a common pattern: -- Reasoning by case over the result of the classification functions - (such as [add_match] for integer addition), gathering additional - information on the shape of the argument expressions in the non-default - cases. -- Inversion of the evaluations of the arguments, exploiting the additional - information thus gathered. -- Equational reasoning over the arithmetic operations performed, - using the lemmas from the [Int] and [Float] modules. -- Construction of an evaluation derivation for the expression returned - by the smart constructor. -*) - -Definition unary_constructor_sound (cstr: expr -> expr) (sem: val -> val) : Prop := - forall le a x, - eval_expr ge sp e m le a x -> - exists v, eval_expr ge sp e m le (cstr a) v /\ Val.lessdef (sem x) v. - -Definition binary_constructor_sound (cstr: expr -> expr -> expr) (sem: val -> val -> val) : Prop := - forall le a x b y, - eval_expr ge sp e m le a x -> - eval_expr ge sp e m le b y -> - exists v, eval_expr ge sp e m le (cstr a b) v /\ Val.lessdef (sem x y) v. - -Theorem eval_addrsymbol: - forall le id ofs, - exists v, eval_expr ge sp e m le (addrsymbol id ofs) v /\ Val.lessdef (Genv.symbol_address ge id ofs) v. -Proof. - intros. unfold addrsymbol. econstructor; split. - EvalOp. simpl; eauto. - auto. -Qed. - -Theorem eval_addrstack: - forall le ofs, - exists v, eval_expr ge sp e m le (addrstack ofs) v /\ Val.lessdef (Val.offset_ptr sp ofs) v. -Proof. - intros. unfold addrstack. econstructor; split. - EvalOp. simpl; eauto. - auto. -Qed. - -Theorem eval_addimm: - forall n, unary_constructor_sound (addimm n) (fun x => Val.add x (Vint n)). -Proof. - red; unfold addimm; intros until x. - predSpec Int.eq Int.eq_spec n Int.zero. - - subst n. intros. exists x; split; auto. - destruct x; simpl; auto. - rewrite Int.add_zero; auto. - destruct Archi.ptr64; auto. rewrite Ptrofs.add_zero; auto. - - case (addimm_match a); intros; InvEval; simpl. - + TrivialExists; simpl. rewrite Int.add_commut. auto. - + econstructor; split. EvalOp. simpl; eauto. - unfold Genv.symbol_address. destruct (Genv.find_symbol ge s); simpl; auto. - destruct Archi.ptr64; auto. rewrite Ptrofs.add_commut; auto. - + econstructor; split. EvalOp. simpl; eauto. - destruct sp; simpl; auto. destruct Archi.ptr64; auto. - rewrite Ptrofs.add_assoc. rewrite (Ptrofs.add_commut m0). auto. - + TrivialExists; simpl. subst x. rewrite Val.add_assoc. rewrite Int.add_commut. auto. - + TrivialExists. -Qed. - -Theorem eval_add: binary_constructor_sound add Val.add. -Proof. - red; intros until y. - unfold add; case (add_match a b); intros; InvEval. - - rewrite Val.add_commut. apply eval_addimm; auto. - - apply eval_addimm; auto. - - subst. - replace (Val.add (Val.add v1 (Vint n1)) (Val.add v0 (Vint n2))) - with (Val.add (Val.add v1 v0) (Val.add (Vint n1) (Vint n2))). - apply eval_addimm. EvalOp. - repeat rewrite Val.add_assoc. decEq. apply Val.add_permut. - - subst. econstructor; split. - EvalOp. constructor. EvalOp. simpl; eauto. constructor. eauto. constructor. simpl; eauto. - rewrite Val.add_commut. destruct sp; simpl; auto. - destruct v1; simpl; auto. - destruct Archi.ptr64 eqn:SF; auto. - apply Val.lessdef_same. f_equal. rewrite ! Ptrofs.add_assoc. f_equal. - rewrite (Ptrofs.add_commut (Ptrofs.of_int n1)), Ptrofs.add_assoc. f_equal. auto with ptrofs. - destruct Archi.ptr64 eqn:SF; auto. - - subst. econstructor; split. - EvalOp. constructor. EvalOp. simpl; eauto. constructor. eauto. constructor. simpl; eauto. - destruct sp; simpl; auto. - destruct v1; simpl; auto. - destruct Archi.ptr64 eqn:SF; auto. - apply Val.lessdef_same. f_equal. rewrite ! Ptrofs.add_assoc. f_equal. f_equal. - rewrite Ptrofs.add_commut. auto with ptrofs. - destruct Archi.ptr64 eqn:SF; auto. - - subst. - replace (Val.add (Val.add v1 (Vint n1)) y) - with (Val.add (Val.add v1 y) (Vint n1)). - apply eval_addimm. EvalOp. - repeat rewrite Val.add_assoc. decEq. apply Val.add_commut. - - subst. - replace (Val.add x (Val.add v1 (Vint n2))) - with (Val.add (Val.add x v1) (Vint n2)). - apply eval_addimm. EvalOp. - repeat rewrite Val.add_assoc. reflexivity. - - TrivialExists. -Qed. - -Theorem eval_sub: binary_constructor_sound sub Val.sub. -Proof. - red; intros until y. - unfold sub; case (sub_match a b); intros; InvEval. - - rewrite Val.sub_add_opp. apply eval_addimm; auto. - - subst. rewrite Val.sub_add_l. rewrite Val.sub_add_r. - rewrite Val.add_assoc. simpl. rewrite Int.add_commut. rewrite <- Int.sub_add_opp. - apply eval_addimm; EvalOp. - - subst. rewrite Val.sub_add_l. apply eval_addimm; EvalOp. - - subst. rewrite Val.sub_add_r. apply eval_addimm; EvalOp. - - TrivialExists. -Qed. - -Theorem eval_negint: unary_constructor_sound negint (fun v => Val.sub Vzero v). -Proof. - red; intros until x. unfold negint. case (negint_match a); intros; InvEval. - TrivialExists. - TrivialExists. -Qed. - -Theorem eval_shlimm: - forall n, unary_constructor_sound (fun a => shlimm a n) - (fun x => Val.shl x (Vint n)). -Proof. - red; intros until x. unfold shlimm. - - predSpec Int.eq Int.eq_spec n Int.zero. - intros; subst. exists x; split; auto. destruct x; simpl; auto. rewrite Int.shl_zero; auto. - - destruct (Int.ltu n Int.iwordsize) eqn:LT; simpl. - destruct (shlimm_match a); intros; InvEval. - - exists (Vint (Int.shl n1 n)); split. EvalOp. - simpl. rewrite LT. auto. - - destruct (Int.ltu (Int.add n n1) Int.iwordsize) eqn:?. - + exists (Val.shl v1 (Vint (Int.add n n1))); split. EvalOp. - subst. destruct v1; simpl; auto. - rewrite Heqb. - destruct (Int.ltu n1 Int.iwordsize) eqn:?; simpl; auto. - destruct (Int.ltu n Int.iwordsize) eqn:?; simpl; auto. - rewrite Int.add_commut. rewrite Int.shl_shl; auto. rewrite Int.add_commut; auto. - + subst. TrivialExists. econstructor. EvalOp. simpl; eauto. constructor. - simpl. auto. - - TrivialExists. - - intros; TrivialExists. constructor. eauto. constructor. EvalOp. simpl; eauto. constructor. - auto. -Qed. - -Theorem eval_shruimm: - forall n, unary_constructor_sound (fun a => shruimm a n) - (fun x => Val.shru x (Vint n)). -Proof. - red; intros until x. unfold shruimm. - - predSpec Int.eq Int.eq_spec n Int.zero. - intros; subst. exists x; split; auto. destruct x; simpl; auto. rewrite Int.shru_zero; auto. - - destruct (Int.ltu n Int.iwordsize) eqn:LT; simpl. - destruct (shruimm_match a); intros; InvEval. - - exists (Vint (Int.shru n1 n)); split. EvalOp. - simpl. rewrite LT; auto. - - destruct (Int.ltu (Int.add n n1) Int.iwordsize) eqn:?. - exists (Val.shru v1 (Vint (Int.add n n1))); split. EvalOp. - subst. destruct v1; simpl; auto. - rewrite Heqb. - destruct (Int.ltu n1 Int.iwordsize) eqn:?; simpl; auto. - rewrite LT. rewrite Int.add_commut. rewrite Int.shru_shru; auto. rewrite Int.add_commut; auto. - subst. TrivialExists. econstructor. EvalOp. simpl; eauto. constructor. - simpl. auto. - - TrivialExists. - - intros; TrivialExists. constructor. eauto. constructor. EvalOp. simpl; eauto. constructor. - auto. -Qed. - -Theorem eval_shrimm: - forall n, unary_constructor_sound (fun a => shrimm a n) - (fun x => Val.shr x (Vint n)). -Proof. - red; intros until x. unfold shrimm. - - predSpec Int.eq Int.eq_spec n Int.zero. - intros; subst. exists x; split; auto. destruct x; simpl; auto. rewrite Int.shr_zero; auto. - - destruct (Int.ltu n Int.iwordsize) eqn:LT; simpl. - destruct (shrimm_match a); intros; InvEval. - - exists (Vint (Int.shr n1 n)); split. EvalOp. - simpl. rewrite LT; auto. - - destruct (Int.ltu (Int.add n n1) Int.iwordsize) eqn:?. - exists (Val.shr v1 (Vint (Int.add n n1))); split. EvalOp. - subst. destruct v1; simpl; auto. - rewrite Heqb. - destruct (Int.ltu n1 Int.iwordsize) eqn:?; simpl; auto. - rewrite LT. - rewrite Int.add_commut. rewrite Int.shr_shr; auto. rewrite Int.add_commut; auto. - subst. TrivialExists. econstructor. EvalOp. simpl; eauto. constructor. - simpl. auto. - - TrivialExists. - - intros; TrivialExists. constructor. eauto. constructor. EvalOp. simpl; eauto. constructor. - auto. -Qed. - -Lemma eval_mulimm_base: - forall n, unary_constructor_sound (mulimm_base n) (fun x => Val.mul x (Vint n)). -Proof. - intros; red; intros; unfold mulimm_base. - - assert (DFL: exists v, eval_expr ge sp e m le (Eop Omul (Eop (Ointconst n) Enil ::: a ::: Enil)) v /\ Val.lessdef (Val.mul x (Vint n)) v). - TrivialExists. econstructor. EvalOp. simpl; eauto. econstructor. eauto. constructor. - rewrite Val.mul_commut. auto. - - generalize (Int.one_bits_decomp n). - generalize (Int.one_bits_range n). - destruct (Int.one_bits n). - - intros. auto. - - destruct l. - + intros. rewrite H1. simpl. - rewrite Int.add_zero. - replace (Vint (Int.shl Int.one i)) with (Val.shl Vone (Vint i)). rewrite Val.shl_mul. - apply eval_shlimm. auto. simpl. rewrite H0; auto with coqlib. - + destruct l. - intros. rewrite H1. simpl. - exploit (eval_shlimm i (x :: le) (Eletvar 0) x). constructor; auto. intros [v1 [A1 B1]]. - exploit (eval_shlimm i0 (x :: le) (Eletvar 0) x). constructor; auto. intros [v2 [A2 B2]]. - exploit (eval_add (x :: le)). eexact A1. eexact A2. intros [v [A B]]. - exists v; split. econstructor; eauto. - rewrite Int.add_zero. - replace (Vint (Int.add (Int.shl Int.one i) (Int.shl Int.one i0))) - with (Val.add (Val.shl Vone (Vint i)) (Val.shl Vone (Vint i0))). - rewrite Val.mul_add_distr_r. - repeat rewrite Val.shl_mul. eapply Val.lessdef_trans. 2: eauto. apply Val.add_lessdef; auto. - simpl. repeat rewrite H0; auto with coqlib. - intros. auto. -Qed. - -Theorem eval_mulimm: - forall n, unary_constructor_sound (mulimm n) (fun x => Val.mul x (Vint n)). -Proof. - intros; red; intros until x; unfold mulimm. - - predSpec Int.eq Int.eq_spec n Int.zero. - intros. exists (Vint Int.zero); split. EvalOp. - destruct x; simpl; auto. subst n. rewrite Int.mul_zero. auto. - - predSpec Int.eq Int.eq_spec n Int.one. - intros. exists x; split; auto. - destruct x; simpl; auto. subst n. rewrite Int.mul_one. auto. - - case (mulimm_match a); intros; InvEval. - - TrivialExists. simpl. rewrite Int.mul_commut; auto. - - subst. rewrite Val.mul_add_distr_l. - exploit eval_mulimm_base; eauto. instantiate (1 := n). intros [v' [A1 B1]]. - exploit (eval_addimm (Int.mul n n2) le (mulimm_base n t2) v'). auto. intros [v'' [A2 B2]]. - exists v''; split; auto. eapply Val.lessdef_trans. eapply Val.add_lessdef; eauto. - rewrite Val.mul_commut; auto. - - apply eval_mulimm_base; auto. -Qed. - -Theorem eval_mul: binary_constructor_sound mul Val.mul. -Proof. - red; intros until y. - unfold mul; case (mul_match a b); intros; InvEval. - rewrite Val.mul_commut. apply eval_mulimm. auto. - apply eval_mulimm. auto. - TrivialExists. -Qed. - -Theorem eval_mulhs: binary_constructor_sound mulhs Val.mulhs. -Proof. - red; intros. unfold mulhs; destruct Archi.ptr64 eqn:SF. -- econstructor; split. - EvalOp. constructor. EvalOp. constructor. EvalOp. constructor. EvalOp. simpl; eauto. - constructor. EvalOp. simpl; eauto. constructor. - simpl; eauto. constructor. simpl; eauto. constructor. simpl; eauto. - destruct x; simpl; auto. destruct y; simpl; auto. - change (Int.ltu (Int.repr 32) Int64.iwordsize') with true; simpl. - apply Val.lessdef_same. f_equal. - transitivity (Int.repr (Z.shiftr (Int.signed i * Int.signed i0) 32)). - unfold Int.mulhs; f_equal. rewrite Zshiftr_div_two_p by lia. reflexivity. - apply Int.same_bits_eq; intros n N. - change Int.zwordsize with 32 in *. - assert (N1: 0 <= n < 64) by lia. - rewrite Int64.bits_loword by auto. - rewrite Int64.bits_shr' by auto. - change (Int.unsigned (Int.repr 32)) with 32. change Int64.zwordsize with 64. - rewrite zlt_true by lia. - rewrite Int.testbit_repr by auto. - unfold Int64.mul. rewrite Int64.testbit_repr by (change Int64.zwordsize with 64; lia). - transitivity (Z.testbit (Int.signed i * Int.signed i0) (n + 32)). - rewrite Z.shiftr_spec by lia. auto. - apply Int64.same_bits_eqm. apply Int64.eqm_mult; apply Int64.eqm_unsigned_repr. - change Int64.zwordsize with 64; lia. -- TrivialExists. -Qed. - -Theorem eval_mulhu: binary_constructor_sound mulhu Val.mulhu. -Proof. - red; intros. unfold mulhu; destruct Archi.ptr64 eqn:SF. -- econstructor; split. - EvalOp. constructor. EvalOp. constructor. EvalOp. constructor. EvalOp. simpl; eauto. - constructor. EvalOp. simpl; eauto. constructor. - simpl; eauto. constructor. simpl; eauto. constructor. simpl; eauto. - destruct x; simpl; auto. destruct y; simpl; auto. - change (Int.ltu (Int.repr 32) Int64.iwordsize') with true; simpl. - apply Val.lessdef_same. f_equal. - transitivity (Int.repr (Z.shiftr (Int.unsigned i * Int.unsigned i0) 32)). - unfold Int.mulhu; f_equal. rewrite Zshiftr_div_two_p by lia. reflexivity. - apply Int.same_bits_eq; intros n N. - change Int.zwordsize with 32 in *. - assert (N1: 0 <= n < 64) by lia. - rewrite Int64.bits_loword by auto. - rewrite Int64.bits_shru' by auto. - change (Int.unsigned (Int.repr 32)) with 32. change Int64.zwordsize with 64. - rewrite zlt_true by lia. - rewrite Int.testbit_repr by auto. - unfold Int64.mul. rewrite Int64.testbit_repr by (change Int64.zwordsize with 64; lia). - transitivity (Z.testbit (Int.unsigned i * Int.unsigned i0) (n + 32)). - rewrite Z.shiftr_spec by lia. auto. - apply Int64.same_bits_eqm. apply Int64.eqm_mult; apply Int64.eqm_unsigned_repr. - change Int64.zwordsize with 64; lia. -- TrivialExists. -Qed. - -Theorem eval_andimm: - forall n, unary_constructor_sound (andimm n) (fun x => Val.and x (Vint n)). -Proof. - intros; red; intros until x. unfold andimm. - - predSpec Int.eq Int.eq_spec n Int.zero. - intros. exists (Vint Int.zero); split. EvalOp. - destruct x; simpl; auto. subst n. rewrite Int.and_zero. auto. - - predSpec Int.eq Int.eq_spec n Int.mone. - intros. exists x; split; auto. - subst. destruct x; simpl; auto. rewrite Int.and_mone; auto. - - case (andimm_match a); intros. - - InvEval. TrivialExists. simpl. rewrite Int.and_commut; auto. - - InvEval. subst. rewrite Val.and_assoc. simpl. rewrite Int.and_commut. TrivialExists. - - TrivialExists. -Qed. - -Theorem eval_and: binary_constructor_sound and Val.and. -Proof. - red; intros until y; unfold and; case (and_match a b); intros; InvEval. - - rewrite Val.and_commut. apply eval_andimm; auto. - - apply eval_andimm; auto. - - TrivialExists. -Qed. - -Theorem eval_orimm: - forall n, unary_constructor_sound (orimm n) (fun x => Val.or x (Vint n)). -Proof. - intros; red; intros until x. unfold orimm. - - predSpec Int.eq Int.eq_spec n Int.zero. - intros. subst. exists x; split; auto. - destruct x; simpl; auto. rewrite Int.or_zero; auto. - - predSpec Int.eq Int.eq_spec n Int.mone. - intros. exists (Vint Int.mone); split. EvalOp. - destruct x; simpl; auto. subst n. rewrite Int.or_mone. auto. - - destruct (orimm_match a); intros; InvEval. - - TrivialExists. simpl. rewrite Int.or_commut; auto. - - subst. rewrite Val.or_assoc. simpl. rewrite Int.or_commut. TrivialExists. - - TrivialExists. -Qed. - -Theorem eval_or: binary_constructor_sound or Val.or. -Proof. - red; intros until y; unfold or; case (or_match a b); intros; InvEval. - - rewrite Val.or_commut. apply eval_orimm; auto. - - apply eval_orimm; auto. - - TrivialExists. -Qed. - -Theorem eval_xorimm: - forall n, unary_constructor_sound (xorimm n) (fun x => Val.xor x (Vint n)). -Proof. - intros; red; intros until x. unfold xorimm. - - predSpec Int.eq Int.eq_spec n Int.zero. - intros. exists x; split. auto. - destruct x; simpl; auto. subst n. rewrite Int.xor_zero. auto. - - intros. destruct (xorimm_match a); intros; InvEval. - - TrivialExists. simpl. rewrite Int.xor_commut; auto. - - subst. rewrite Val.xor_assoc. simpl. rewrite Int.xor_commut. - predSpec Int.eq Int.eq_spec (Int.xor n2 n) Int.zero. - + exists v1; split; auto. destruct v1; simpl; auto. rewrite H0, Int.xor_zero; auto. - + TrivialExists. - - TrivialExists. -Qed. - -Theorem eval_xor: binary_constructor_sound xor Val.xor. -Proof. - red; intros until y; unfold xor; case (xor_match a b); intros; InvEval. - - rewrite Val.xor_commut. apply eval_xorimm; auto. - - apply eval_xorimm; auto. - - TrivialExists. -Qed. - -Theorem eval_notint: unary_constructor_sound notint Val.notint. -Proof. - unfold notint; red; intros. rewrite Val.not_xor. apply eval_xorimm; auto. -Qed. - -Theorem eval_divs_base: - forall le a b x y z, - eval_expr ge sp e m le a x -> - eval_expr ge sp e m le b y -> - Val.divs x y = Some z -> - exists v, eval_expr ge sp e m le (divs_base a b) v /\ Val.lessdef z v. -Proof. - intros. unfold divs_base. exists z; split. EvalOp. - 2: apply Val.lessdef_refl. - cbn. - rewrite H1. - cbn. - trivial. -Qed. - -Theorem eval_mods_base: - forall le a b x y z, - eval_expr ge sp e m le a x -> - eval_expr ge sp e m le b y -> - Val.mods x y = Some z -> - exists v, eval_expr ge sp e m le (mods_base a b) v /\ Val.lessdef z v. -Proof. - intros. unfold mods_base. exists z; split. EvalOp. - 2: apply Val.lessdef_refl. - cbn. - rewrite H1. - cbn. - trivial. -Qed. - -Theorem eval_divu_base: - forall le a b x y z, - eval_expr ge sp e m le a x -> - eval_expr ge sp e m le b y -> - Val.divu x y = Some z -> - exists v, eval_expr ge sp e m le (divu_base a b) v /\ Val.lessdef z v. -Proof. - intros. unfold divu_base. exists z; split. EvalOp. - 2: apply Val.lessdef_refl. - cbn. - rewrite H1. - cbn. - trivial. -Qed. - -Theorem eval_modu_base: - forall le a b x y z, - eval_expr ge sp e m le a x -> - eval_expr ge sp e m le b y -> - Val.modu x y = Some z -> - exists v, eval_expr ge sp e m le (modu_base a b) v /\ Val.lessdef z v. -Proof. - intros. unfold modu_base. exists z; split. EvalOp. - 2: apply Val.lessdef_refl. - cbn. - rewrite H1. - cbn. - trivial. -Qed. - -Theorem eval_shrximm: - forall le a n x z, - eval_expr ge sp e m le a x -> - Val.shrx x (Vint n) = Some z -> - exists v, eval_expr ge sp e m le (shrximm a n) v /\ Val.lessdef z v. -Proof. - intros. unfold shrximm. - predSpec Int.eq Int.eq_spec n Int.zero. - subst n. exists x; split; auto. - destruct x; simpl in H0; try discriminate. - destruct (Int.ltu Int.zero (Int.repr 31)); inv H0. - replace (Int.shrx i Int.zero) with i. auto. - unfold Int.shrx, Int.divs. rewrite Int.shl_zero. - change (Int.signed Int.one) with 1. rewrite Z.quot_1_r. rewrite Int.repr_signed; auto. -<<<<<<< HEAD - econstructor; split. EvalOp. - cbn. - rewrite H0. - cbn. - reflexivity. - apply Val.lessdef_refl. -======= - econstructor; split. EvalOp. auto. -(* - intros. destruct x; simpl in H0; try discriminate. - destruct (Int.ltu n (Int.repr 31)) eqn:LTU; inv H0. - unfold shrximm. - predSpec Int.eq Int.eq_spec n Int.zero. - - subst n. exists (Vint i); split; auto. - unfold Int.shrx, Int.divs. rewrite Z.quot_1_r. rewrite Int.repr_signed. auto. - - assert (NZ: Int.unsigned n <> 0). - { intro EQ; elim H0. rewrite <- (Int.repr_unsigned n). rewrite EQ; auto. } - assert (LT: 0 <= Int.unsigned n < 31) by (apply Int.ltu_inv in LTU; assumption). - assert (LTU2: Int.ltu (Int.sub Int.iwordsize n) Int.iwordsize = true). - { unfold Int.ltu; apply zlt_true. - unfold Int.sub. change (Int.unsigned Int.iwordsize) with 32. - rewrite Int.unsigned_repr. lia. - assert (32 < Int.max_unsigned) by reflexivity. lia. } - assert (X: eval_expr ge sp e m le - (Eop (Oshrimm (Int.repr (Int.zwordsize - 1))) (a ::: Enil)) - (Vint (Int.shr i (Int.repr (Int.zwordsize - 1))))). - { EvalOp. } - assert (Y: eval_expr ge sp e m le (shrximm_inner a n) - (Vint (Int.shru (Int.shr i (Int.repr (Int.zwordsize - 1))) (Int.sub Int.iwordsize n)))). - { EvalOp. simpl. rewrite LTU2. auto. } - TrivialExists. - constructor. EvalOp. simpl; eauto. constructor. - simpl. unfold Int.ltu; rewrite zlt_true. rewrite Int.shrx_shr_2 by auto. reflexivity. - change (Int.unsigned Int.iwordsize) with 32; lia. -*) ->>>>>>> master -Qed. - -Theorem eval_shl: binary_constructor_sound shl Val.shl. -Proof. - red; intros until y; unfold shl; case (shl_match b); intros. - InvEval. apply eval_shlimm; auto. - TrivialExists. -Qed. - -Theorem eval_shr: binary_constructor_sound shr Val.shr. -Proof. - red; intros until y; unfold shr; case (shr_match b); intros. - InvEval. apply eval_shrimm; auto. - TrivialExists. -Qed. - -Theorem eval_shru: binary_constructor_sound shru Val.shru. -Proof. - red; intros until y; unfold shru; case (shru_match b); intros. - InvEval. apply eval_shruimm; auto. - TrivialExists. -Qed. - -Theorem eval_negf: unary_constructor_sound negf Val.negf. -Proof. - red; intros. TrivialExists. -Qed. - -Theorem eval_absf: unary_constructor_sound absf Val.absf. -Proof. - red; intros. TrivialExists. -Qed. - -Theorem eval_addf: binary_constructor_sound addf Val.addf. -Proof. - red; intros; TrivialExists. -Qed. - -Theorem eval_subf: binary_constructor_sound subf Val.subf. -Proof. - red; intros; TrivialExists. -Qed. - -Theorem eval_mulf: binary_constructor_sound mulf Val.mulf. -Proof. - red; intros; TrivialExists. -Qed. - -Theorem eval_negfs: unary_constructor_sound negfs Val.negfs. -Proof. - red; intros. TrivialExists. -Qed. - -Theorem eval_absfs: unary_constructor_sound absfs Val.absfs. -Proof. - red; intros. TrivialExists. -Qed. - -Theorem eval_addfs: binary_constructor_sound addfs Val.addfs. -Proof. - red; intros; TrivialExists. -Qed. - -Theorem eval_subfs: binary_constructor_sound subfs Val.subfs. -Proof. - red; intros; TrivialExists. -Qed. - -Theorem eval_mulfs: binary_constructor_sound mulfs Val.mulfs. -Proof. - red; intros; TrivialExists. -Qed. - -Section COMP_IMM. - -Variable default: comparison -> int -> condition. -Variable intsem: comparison -> int -> int -> bool. -Variable sem: comparison -> val -> val -> val. - -Hypothesis sem_int: forall c x y, sem c (Vint x) (Vint y) = Val.of_bool (intsem c x y). -Hypothesis sem_undef: forall c v, sem c Vundef v = Vundef. -Hypothesis sem_eq: forall x y, sem Ceq (Vint x) (Vint y) = Val.of_bool (Int.eq x y). -Hypothesis sem_ne: forall x y, sem Cne (Vint x) (Vint y) = Val.of_bool (negb (Int.eq x y)). -Hypothesis sem_default: forall c v n, sem c v (Vint n) = Val.of_optbool (eval_condition (default c n) (v :: nil) m). - -Lemma eval_compimm: - forall le c a n2 x, - eval_expr ge sp e m le a x -> - exists v, eval_expr ge sp e m le (compimm default intsem c a n2) v - /\ Val.lessdef (sem c x (Vint n2)) v. -Proof. - intros until x. - unfold compimm; case (compimm_match c a); intros. -(* constant *) - - InvEval. rewrite sem_int. TrivialExists. simpl. destruct (intsem c0 n1 n2); auto. -(* eq cmp *) - - InvEval. inv H. simpl in H5. inv H5. - destruct (Int.eq_dec n2 Int.zero). - + subst n2. TrivialExists. - simpl. rewrite eval_negate_condition. - destruct (eval_condition c0 vl m); simpl. - unfold Vtrue, Vfalse. destruct b; simpl; rewrite sem_eq; auto. - rewrite sem_undef; auto. - + destruct (Int.eq_dec n2 Int.one). subst n2. TrivialExists. - simpl. destruct (eval_condition c0 vl m); simpl. - unfold Vtrue, Vfalse. destruct b; simpl; rewrite sem_eq; auto. - rewrite sem_undef; auto. - exists (Vint Int.zero); split. EvalOp. - destruct (eval_condition c0 vl m); simpl. - unfold Vtrue, Vfalse. destruct b; rewrite sem_eq; rewrite Int.eq_false; auto. - rewrite sem_undef; auto. -(* ne cmp *) - - InvEval. inv H. simpl in H5. inv H5. - destruct (Int.eq_dec n2 Int.zero). - + subst n2. TrivialExists. - simpl. destruct (eval_condition c0 vl m); simpl. - unfold Vtrue, Vfalse. destruct b; simpl; rewrite sem_ne; auto. - rewrite sem_undef; auto. - + destruct (Int.eq_dec n2 Int.one). subst n2. TrivialExists. - simpl. rewrite eval_negate_condition. destruct (eval_condition c0 vl m); simpl. - unfold Vtrue, Vfalse. destruct b; simpl; rewrite sem_ne; auto. - rewrite sem_undef; auto. - exists (Vint Int.one); split. EvalOp. - destruct (eval_condition c0 vl m); simpl. - unfold Vtrue, Vfalse. destruct b; rewrite sem_ne; rewrite Int.eq_false; auto. - rewrite sem_undef; auto. -(* default *) - - TrivialExists. simpl. rewrite sem_default. auto. -Qed. - -Hypothesis sem_swap: - forall c x y, sem (swap_comparison c) x y = sem c y x. - -Lemma eval_compimm_swap: - forall le c a n2 x, - eval_expr ge sp e m le a x -> - exists v, eval_expr ge sp e m le (compimm default intsem (swap_comparison c) a n2) v - /\ Val.lessdef (sem c (Vint n2) x) v. -Proof. - intros. rewrite <- sem_swap. eapply eval_compimm; eauto. -Qed. - -End COMP_IMM. - -Theorem eval_comp: - forall c, binary_constructor_sound (comp c) (Val.cmp c). -Proof. - intros; red; intros until y. unfold comp; case (comp_match a b); intros; InvEval. - eapply eval_compimm_swap; eauto. - intros. unfold Val.cmp. rewrite Val.swap_cmp_bool; auto. - eapply eval_compimm; eauto. - TrivialExists. -Qed. - -Theorem eval_compu: - forall c, binary_constructor_sound (compu c) (Val.cmpu (Mem.valid_pointer m) c). -Proof. - intros; red; intros until y. unfold compu; case (compu_match a b); intros; InvEval. - eapply eval_compimm_swap; eauto. - intros. unfold Val.cmpu. rewrite Val.swap_cmpu_bool; auto. - eapply eval_compimm; eauto. - TrivialExists. -Qed. - -Theorem eval_compf: - forall c, binary_constructor_sound (compf c) (Val.cmpf c). -Proof. - intros; red; intros. unfold compf. TrivialExists. -Qed. - -Theorem eval_compfs: - forall c, binary_constructor_sound (compfs c) (Val.cmpfs c). -Proof. - intros; red; intros. unfold compfs. TrivialExists. -Qed. - -Theorem eval_cast8signed: unary_constructor_sound cast8signed (Val.sign_ext 8). -Proof. - red; intros until x. unfold cast8signed. case (cast8signed_match a); intros; InvEval. - TrivialExists. - TrivialExists. -Qed. - -Theorem eval_cast8unsigned: unary_constructor_sound cast8unsigned (Val.zero_ext 8). -Proof. - red; intros until x. unfold cast8unsigned. - rewrite Val.zero_ext_and. apply eval_andimm. lia. -Qed. - -Theorem eval_cast16signed: unary_constructor_sound cast16signed (Val.sign_ext 16). -Proof. - red; intros until x. unfold cast16signed. case (cast16signed_match a); intros; InvEval. - TrivialExists. - TrivialExists. -Qed. - -Theorem eval_cast16unsigned: unary_constructor_sound cast16unsigned (Val.zero_ext 16). -Proof. - red; intros until x. unfold cast8unsigned. - rewrite Val.zero_ext_and. apply eval_andimm. lia. -Qed. - -Theorem eval_intoffloat: - forall le a x y, - eval_expr ge sp e m le a x -> - Val.intoffloat x = Some y -> - exists v, eval_expr ge sp e m le (intoffloat a) v /\ Val.lessdef y v. -Proof. - intros; unfold intoffloat. TrivialExists. - cbn. rewrite H0. reflexivity. -Qed. - -Theorem eval_intuoffloat: - forall le a x y, - eval_expr ge sp e m le a x -> - Val.intuoffloat x = Some y -> - exists v, eval_expr ge sp e m le (intuoffloat a) v /\ Val.lessdef y v. -Proof. - intros; unfold intuoffloat. TrivialExists. - cbn. rewrite H0. reflexivity. -Qed. - -Theorem eval_floatofintu: - forall le a x y, - eval_expr ge sp e m le a x -> - Val.floatofintu x = Some y -> - exists v, eval_expr ge sp e m le (floatofintu a) v /\ Val.lessdef y v. -Proof. - intros until y; unfold floatofintu. case (floatofintu_match a); intros. - InvEval. simpl in H0. TrivialExists. - TrivialExists. - cbn. rewrite H0. reflexivity. -Qed. - -Theorem eval_floatofint: - forall le a x y, - eval_expr ge sp e m le a x -> - Val.floatofint x = Some y -> - exists v, eval_expr ge sp e m le (floatofint a) v /\ Val.lessdef y v. -Proof. - intros until y; unfold floatofint. case (floatofint_match a); intros. - InvEval. simpl in H0. TrivialExists. - TrivialExists. - cbn. rewrite H0. reflexivity. -Qed. - -Theorem eval_intofsingle: - forall le a x y, - eval_expr ge sp e m le a x -> - Val.intofsingle x = Some y -> - exists v, eval_expr ge sp e m le (intofsingle a) v /\ Val.lessdef y v. -Proof. - intros; unfold intofsingle. TrivialExists. - cbn. rewrite H0. reflexivity. -Qed. - -Theorem eval_singleofint: - forall le a x y, - eval_expr ge sp e m le a x -> - Val.singleofint x = Some y -> - exists v, eval_expr ge sp e m le (singleofint a) v /\ Val.lessdef y v. -Proof. - intros; unfold singleofint; TrivialExists. - cbn. rewrite H0. reflexivity. -Qed. - -Theorem eval_intuofsingle: - forall le a x y, - eval_expr ge sp e m le a x -> - Val.intuofsingle x = Some y -> - exists v, eval_expr ge sp e m le (intuofsingle a) v /\ Val.lessdef y v. -Proof. - intros; unfold intuofsingle. TrivialExists. - cbn. rewrite H0. reflexivity. -Qed. - -Theorem eval_singleofintu: - forall le a x y, - eval_expr ge sp e m le a x -> - Val.singleofintu x = Some y -> - exists v, eval_expr ge sp e m le (singleofintu a) v /\ Val.lessdef y v. -Proof. - intros; unfold intuofsingle. TrivialExists. - cbn. rewrite H0. reflexivity. -Qed. - -Theorem eval_singleoffloat: unary_constructor_sound singleoffloat Val.singleoffloat. -Proof. - red; intros. unfold singleoffloat. TrivialExists. -Qed. - -Theorem eval_floatofsingle: unary_constructor_sound floatofsingle Val.floatofsingle. -Proof. - red; intros. unfold floatofsingle. TrivialExists. -Qed. - -Lemma mod_small_negative: - forall a modulus, - modulus > 0 -> -modulus < a < 0 -> a mod modulus = a + modulus. -Proof. - intros. - replace (a mod modulus) with ((a + modulus) mod modulus). - apply Z.mod_small. - lia. - rewrite <- Zplus_mod_idemp_r. - rewrite Z.mod_same by lia. - rewrite Z.add_0_r. - reflexivity. -Qed. - -Remark normalize_low_long: forall - (PTR64 : Archi.ptr64 = true) v1, - Val.loword (Val.normalize (Val.longofint v1) Tlong) = Val.normalize v1 Tint. -Proof. - intros. - destruct v1; cbn; try rewrite PTR64; trivial. - f_equal. - unfold Int64.loword. - unfold Int.signed. - destruct zlt. - { rewrite Int64.int_unsigned_repr. - apply Int.repr_unsigned. - } - pose proof (Int.unsigned_range i). - rewrite Int64.unsigned_repr_eq. - replace ((Int.unsigned i - Int.modulus) mod Int64.modulus) - with (Int64.modulus + Int.unsigned i - Int.modulus). - { - rewrite <- (Int.repr_unsigned i) at 2. - apply Int.eqm_samerepr. - unfold Int.eqm, eqmod. - change Int.modulus with 4294967296 in *. - change Int64.modulus with 18446744073709551616 in *. - exists 4294967295. - lia. - } - { rewrite mod_small_negative. - lia. - constructor. - constructor. - change Int.modulus with 4294967296 in *. - change Int.half_modulus with 2147483648 in *. - change Int64.modulus with 18446744073709551616 in *. - lia. - lia. - } -Qed. - -Lemma same_expr_pure_correct: - forall le a1 a2 v1 v2 - (PURE : same_expr_pure a1 a2 = true) - (EVAL1 : eval_expr ge sp e m le a1 v1) - (EVAL2 : eval_expr ge sp e m le a2 v2), - v1 = v2. -Proof. - intros. - destruct a1; destruct a2; cbn in *; try discriminate. - inv EVAL1. inv EVAL2. - destruct (ident_eq i i0); congruence. -Qed. - -Theorem eval_select: - forall le ty cond al vl a1 v1 a2 v2 a b, - select ty cond al a1 a2 = Some a -> - eval_exprlist ge sp e m le al vl -> - eval_expr ge sp e m le a1 v1 -> - eval_expr ge sp e m le a2 v2 -> - eval_condition cond vl m = Some b -> - exists v, - eval_expr ge sp e m le a v - /\ Val.lessdef (Val.select (Some b) v1 v2 ty) v. -Proof. - unfold select; intros. - pose proof (same_expr_pure_correct le a1 a2 v1 v2) as PURE. - destruct (same_expr_pure a1 a2). - { rewrite <- PURE by auto. - inv H. - exists v1. split. assumption. - unfold Val.select. - destruct b; apply Val.lessdef_normalize. - } - clear PURE. - destruct Archi.ptr64 eqn:PTR64. - 2: discriminate. - destruct ty; cbn in *; try discriminate. - - (* Tint *) - inv H. TrivialExists. - + cbn. repeat econstructor; eassumption. - + cbn. f_equal. rewrite ExtValues.normalize_select01. - rewrite H3. destruct b. - * rewrite ExtValues.select01_long_true. apply normalize_low_long; assumption. - * rewrite ExtValues.select01_long_false. apply normalize_low_long; assumption. - - - (* Tfloat *) - inv H. TrivialExists. - + cbn. repeat econstructor; eassumption. - + cbn. f_equal. rewrite ExtValues.normalize_select01. - rewrite H3. destruct b. - * rewrite ExtValues.select01_long_true. - apply ExtValues.float_bits_normalize. - * rewrite ExtValues.select01_long_false. - apply ExtValues.float_bits_normalize. - - - (* Tlong *) - inv H. TrivialExists. - + cbn. repeat econstructor; eassumption. - + cbn. f_equal. rewrite ExtValues.normalize_select01. - rewrite H3. destruct b. - * rewrite ExtValues.select01_long_true. reflexivity. - * rewrite ExtValues.select01_long_false. reflexivity. - - - (* Tsingle *) - inv H. TrivialExists. - + cbn. repeat econstructor; eassumption. - + cbn. f_equal. rewrite ExtValues.normalize_select01. - rewrite H3. destruct b. - * rewrite ExtValues.select01_long_true. - rewrite normalize_low_long by assumption. - apply ExtValues.single_bits_normalize. - * rewrite ExtValues.select01_long_false. - rewrite normalize_low_long by assumption. - apply ExtValues.single_bits_normalize. -Qed. - -Theorem eval_addressing: - forall le chunk a v b ofs, - eval_expr ge sp e m le a v -> - v = Vptr b ofs -> - match addressing chunk a with (mode, args) => - exists vl, - eval_exprlist ge sp e m le args vl /\ - eval_addressing ge sp mode vl = Some v - end. -Proof. - intros until v. unfold addressing; case (addressing_match a); intros; InvEval. - - exists (@nil val); split. eauto with evalexpr. simpl. auto. - - destruct (Archi.pic_code tt). - + exists (Vptr b ofs0 :: nil); split. - constructor. EvalOp. simpl. congruence. constructor. simpl. rewrite Ptrofs.add_zero. congruence. - + exists (@nil val); split. constructor. simpl; auto. - - exists (v1 :: nil); split. eauto with evalexpr. simpl. - destruct v1; simpl in H; try discriminate. destruct Archi.ptr64 eqn:SF; inv H. - simpl. auto. - - exists (v1 :: nil); split. eauto with evalexpr. simpl. - destruct v1; simpl in H; try discriminate. destruct Archi.ptr64 eqn:SF; inv H. - simpl. auto. - - exists (v :: nil); split. eauto with evalexpr. subst. simpl. rewrite Ptrofs.add_zero; auto. -Qed. - -Theorem eval_builtin_arg: - forall a v, - eval_expr ge sp e m nil a v -> - CminorSel.eval_builtin_arg ge sp e m (builtin_arg a) v. -Proof. - intros until v. unfold builtin_arg; case (builtin_arg_match a); intros. -- InvEval. constructor. -- InvEval. constructor. -- InvEval. constructor. -- InvEval. simpl in H5. inv H5. constructor. -- InvEval. subst v. constructor; auto. -- inv H. InvEval. simpl in H6; inv H6. constructor; auto. -- destruct Archi.ptr64 eqn:SF. -+ constructor; auto. -+ InvEval. replace v with (if Archi.ptr64 then Val.addl v1 (Vint n) else Val.add v1 (Vint n)). - repeat constructor; auto. - rewrite SF; auto. -- destruct Archi.ptr64 eqn:SF. -+ InvEval. replace v with (if Archi.ptr64 then Val.addl v1 (Vlong n) else Val.add v1 (Vlong n)). - repeat constructor; auto. - rewrite SF; auto. -+ constructor; auto. -- constructor; auto. -Qed. - -(* floating-point division without HELPERS *) -Theorem eval_divf_base: - forall le a b x y, - eval_expr ge sp e m le a x -> - eval_expr ge sp e m le b y -> - exists v, eval_expr ge sp e m le (divf_base a b) v /\ Val.lessdef (Val.divf x y) v. -Proof. - intros; unfold divf_base. - TrivialExists. -Qed. - -Theorem eval_divfs_base: - forall le a b x y, - eval_expr ge sp e m le a x -> - eval_expr ge sp e m le b y -> - exists v, eval_expr ge sp e m le (divfs_base a b) v /\ Val.lessdef (Val.divfs x y) v. -Proof. - intros; unfold divfs_base. - TrivialExists. -Qed. - -(** Platform-specific known builtins *) - -Theorem eval_platform_builtin: - forall bf al a vl v le, - platform_builtin bf al = Some a -> - eval_exprlist ge sp e m le al vl -> - platform_builtin_sem bf vl = Some v -> - exists v', eval_expr ge sp e m le a v' /\ Val.lessdef v v'. -Proof. - destruct bf; intros until le; intro Heval. - all: try (inversion Heval; subst a; clear Heval; - exists v; split; trivial; - repeat (try econstructor; try eassumption)). -Qed. - -End CMCONSTR. diff --git a/riscV/TO_MERGE/TargetPrinter.ml b/riscV/TO_MERGE/TargetPrinter.ml deleted file mode 100644 index 23fbeb8b..00000000 --- a/riscV/TO_MERGE/TargetPrinter.ml +++ /dev/null @@ -1,677 +0,0 @@ -(* *********************************************************************) -(* *) -(* The Compcert verified compiler *) -(* *) -(* Xavier Leroy, INRIA Paris-Rocquencourt *) -(* Prashanth Mundkur, SRI International *) -(* *) -(* 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. *) -(* *) -(* The contributions by Prashanth Mundkur are reused and adapted *) -(* under the terms of a Contributor License Agreement between *) -(* SRI International and INRIA. *) -(* *) -(* *********************************************************************) - -(* Printing RISC-V assembly code in asm syntax *) - -open Printf -open Camlcoq -open Sections -open AST -open Asm -open AisAnnot -open PrintAsmaux -open Fileinfo - -(* Module containing the printing functions *) - -module Target : TARGET = - struct - -(* Basic printing functions *) - - let comment = "#" - - let symbol = elf_symbol - let symbol_offset = elf_symbol_offset - let label = elf_label - - let print_label oc lbl = label oc (transl_label lbl) - - let use_abi_name = false - - let int_reg_num_name = function - | X1 -> "x1" | X2 -> "x2" | X3 -> "x3" - | X4 -> "x4" | X5 -> "x5" | X6 -> "x6" | X7 -> "x7" - | X8 -> "x8" | X9 -> "x9" | X10 -> "x10" | X11 -> "x11" - | X12 -> "x12" | X13 -> "x13" | X14 -> "x14" | X15 -> "x15" - | X16 -> "x16" | X17 -> "x17" | X18 -> "x18" | X19 -> "x19" - | X20 -> "x20" | X21 -> "x21" | X22 -> "x22" | X23 -> "x23" - | X24 -> "x24" | X25 -> "x25" | X26 -> "x26" | X27 -> "x27" - | X28 -> "x28" | X29 -> "x29" | X30 -> "x30" | X31 -> "x31" - - let int_reg_abi_name = function - | X1 -> "ra" | X2 -> "sp" | X3 -> "gp" - | X4 -> "tp" | X5 -> "t0" | X6 -> "t1" | X7 -> "t2" - | X8 -> "s0" | X9 -> "s1" | X10 -> "a0" | X11 -> "a1" - | X12 -> "a2" | X13 -> "a3" | X14 -> "a4" | X15 -> "a5" - | X16 -> "a6" | X17 -> "a7" | X18 -> "s2" | X19 -> "s3" - | X20 -> "s4" | X21 -> "s5" | X22 -> "s6" | X23 -> "s7" - | X24 -> "s8" | X25 -> "s9" | X26 -> "s10" | X27 -> "s11" - | X28 -> "t3" | X29 -> "t4" | X30 -> "t5" | X31 -> "t6" - - let float_reg_num_name = function - | F0 -> "f0" | F1 -> "f1" | F2 -> "f2" | F3 -> "f3" - | F4 -> "f4" | F5 -> "f5" | F6 -> "f6" | F7 -> "f7" - | F8 -> "f8" | F9 -> "f9" | F10 -> "f10" | F11 -> "f11" - | F12 -> "f12" | F13 -> "f13" | F14 -> "f14" | F15 -> "f15" - | F16 -> "f16" | F17 -> "f17" | F18 -> "f18" | F19 -> "f19" - | F20 -> "f20" | F21 -> "f21" | F22 -> "f22" | F23 -> "f23" - | F24 -> "f24" | F25 -> "f25" | F26 -> "f26" | F27 -> "f27" - | F28 -> "f28" | F29 -> "f29" | F30 -> "f30" | F31 -> "f31" - - let float_reg_abi_name = function - | F0 -> "ft0" | F1 -> "ft1" | F2 -> "ft2" | F3 -> "ft3" - | F4 -> "ft4" | F5 -> "ft5" | F6 -> "ft6" | F7 -> "ft7" - | F8 -> "fs0" | F9 -> "fs1" | F10 -> "fa0" | F11 -> "fa1" - | F12 -> "fa2" | F13 -> "fa3" | F14 -> "fa4" | F15 -> "fa5" - | F16 -> "fa6" | F17 -> "fa7" | F18 -> "fs2" | F19 -> "fs3" - | F20 -> "fs4" | F21 -> "fs5" | F22 -> "fs6" | F23 -> "fs7" - | F24 -> "fs8" | F25 -> "fs9" | F26 ->"fs10" | F27 -> "fs11" - | F28 -> "ft3" | F29 -> "ft4" | F30 -> "ft5" | F31 -> "ft6" - - let int_reg_name = if use_abi_name then int_reg_abi_name else int_reg_num_name - let float_reg_name = if use_abi_name then float_reg_abi_name else float_reg_num_name - - let ireg oc r = output_string oc (int_reg_name r) - let freg oc r = output_string oc (float_reg_name r) - - let ireg0 oc = function - | X0 -> output_string oc "x0" - | X r -> ireg oc r - - let preg_asm oc ty = function - | IR r -> ireg oc r - | FR r -> freg oc r - | _ -> assert false - - let preg_annot = function - | IR r -> int_reg_name r - | FR r -> float_reg_name r - | _ -> assert false - -(* Names of sections *) - - let name_of_section = function - | Section_text -> ".text" -<<<<<<< HEAD - | Section_data(i, true) -> - failwith "_Thread_local unsupported on this platform" - | Section_data(i, false) | Section_small_data i -> - if i then ".data" else common_section () -======= - | Section_data i | Section_small_data i -> - variable_section ~sec:".data" ~bss:".bss" i ->>>>>>> master - | Section_const i | Section_small_const i -> - variable_section ~sec:".section .rodata" i - | Section_string -> ".section .rodata" - | Section_literal -> ".section .rodata" - | Section_jumptable -> ".section .rodata" - | Section_debug_info _ -> ".section .debug_info,\"\",%progbits" - | Section_debug_loc -> ".section .debug_loc,\"\",%progbits" - | Section_debug_abbrev -> ".section .debug_abbrev,\"\",%progbits" - | Section_debug_line _ -> ".section .debug_line,\"\",%progbits" - | Section_debug_ranges -> ".section .debug_ranges,\"\",%progbits" - | Section_debug_str -> ".section .debug_str,\"MS\",%progbits,1" - | Section_user(s, wr, ex) -> - sprintf ".section \"%s\",\"a%s%s\",%%progbits" - s (if wr then "w" else "") (if ex then "x" else "") - | Section_ais_annotation -> sprintf ".section \"__compcert_ais_annotations\",\"\",@note" - - let section oc sec = - fprintf oc " %s\n" (name_of_section sec) - -(* Associate labels to floating-point constants and to symbols. *) - - let emit_constants oc lit = - if exists_constants () then begin - section oc lit; - if Hashtbl.length literal64_labels > 0 then - begin - fprintf oc " .align 3\n"; - Hashtbl.iter - (fun bf lbl -> fprintf oc "%a: .quad 0x%Lx\n" label lbl bf) - literal64_labels - end; - if Hashtbl.length literal32_labels > 0 then - begin - fprintf oc " .align 2\n"; - Hashtbl.iter - (fun bf lbl -> - fprintf oc "%a: .long 0x%lx\n" label lbl bf) - literal32_labels - end; - reset_literals () - end - -(* Generate code to load the address of id + ofs in register r *) - - let loadsymbol oc r id ofs = - if Archi.pic_code () then begin - assert (ofs = Integers.Ptrofs.zero); - fprintf oc " la %a, %s\n" ireg r (extern_atom id) - end else begin - fprintf oc " lui %a, %%hi(%a)\n" - ireg r symbol_offset (id, ofs); - fprintf oc " addi %a, %a, %%lo(%a)\n" - ireg r ireg r symbol_offset (id, ofs) - end - -(* Emit .file / .loc debugging directives *) - - let print_file_line oc file line = - print_file_line oc comment file line - -(* - let print_location oc loc = - if loc <> Cutil.no_loc then print_file_line oc (fst loc) (snd loc) -*) - -(* Add "w" suffix to 32-bit instructions if we are in 64-bit mode *) - - let w oc = - if Archi.ptr64 then output_string oc "w" - -(* Offset part of a load or store *) - - let offset oc = function - | Ofsimm n -> ptrofs oc n - | Ofslow(id, ofs) -> fprintf oc "%%lo(%a)" symbol_offset (id, ofs) - -(* Printing of instructions *) - let print_instruction oc = function - | Pmv(rd, rs) -> - fprintf oc " mv %a, %a\n" ireg rd ireg rs - - (* 32-bit integer register-immediate instructions *) - | Paddiw (rd, rs, imm) -> - fprintf oc " addi%t %a, %a, %a\n" w ireg rd ireg0 rs coqint imm - | Psltiw (rd, rs, imm) -> - fprintf oc " slti %a, %a, %a\n" ireg rd ireg0 rs coqint imm - | Psltiuw (rd, rs, imm) -> - fprintf oc " sltiu %a, %a, %a\n" ireg rd ireg0 rs coqint imm - | Pandiw (rd, rs, imm) -> - fprintf oc " andi %a, %a, %a\n" ireg rd ireg0 rs coqint imm - | Poriw (rd, rs, imm) -> - fprintf oc " ori %a, %a, %a\n" ireg rd ireg0 rs coqint imm - | Pxoriw (rd, rs, imm) -> - fprintf oc " xori %a, %a, %a\n" ireg rd ireg0 rs coqint imm - | Pslliw (rd, rs, imm) -> - fprintf oc " slli%t %a, %a, %a\n" w ireg rd ireg0 rs coqint imm - | Psrliw (rd, rs, imm) -> - fprintf oc " srli%t %a, %a, %a\n" w ireg rd ireg0 rs coqint imm - | Psraiw (rd, rs, imm) -> - fprintf oc " srai%t %a, %a, %a\n" w ireg rd ireg0 rs coqint imm - | Pluiw (rd, imm) -> - fprintf oc " lui %a, %a\n" ireg rd coqint imm - - (* 32-bit integer register-register instructions *) - | Paddw(rd, rs1, rs2) -> - fprintf oc " add%t %a, %a, %a\n" w ireg rd ireg0 rs1 ireg0 rs2 - | Psubw(rd, rs1, rs2) -> - fprintf oc " sub%t %a, %a, %a\n" w ireg rd ireg0 rs1 ireg0 rs2 - - | Pmulw(rd, rs1, rs2) -> - fprintf oc " mul%t %a, %a, %a\n" w ireg rd ireg0 rs1 ireg0 rs2 - | Pmulhw(rd, rs1, rs2) -> assert (not Archi.ptr64); - fprintf oc " mulh %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 - | Pmulhuw(rd, rs1, rs2) -> assert (not Archi.ptr64); - fprintf oc " mulhu %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 - - | Pdivw(rd, rs1, rs2) -> - fprintf oc " div%t %a, %a, %a\n" w ireg rd ireg0 rs1 ireg0 rs2 - | Pdivuw(rd, rs1, rs2) -> - fprintf oc " divu%t %a, %a, %a\n" w ireg rd ireg0 rs1 ireg0 rs2 - | Premw(rd, rs1, rs2) -> - fprintf oc " rem%t %a, %a, %a\n" w ireg rd ireg0 rs1 ireg0 rs2 - | Premuw(rd, rs1, rs2) -> - fprintf oc " remu%t %a, %a, %a\n" w ireg rd ireg0 rs1 ireg0 rs2 - - | Psltw(rd, rs1, rs2) -> - fprintf oc " slt %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 - | Psltuw(rd, rs1, rs2) -> - fprintf oc " sltu %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 - - | Pandw(rd, rs1, rs2) -> - fprintf oc " and %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 - | Porw(rd, rs1, rs2) -> - fprintf oc " or %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 - | Pxorw(rd, rs1, rs2) -> - fprintf oc " xor %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 - | Psllw(rd, rs1, rs2) -> - fprintf oc " sll%t %a, %a, %a\n" w ireg rd ireg0 rs1 ireg0 rs2 - | Psrlw(rd, rs1, rs2) -> - fprintf oc " srl%t %a, %a, %a\n" w ireg rd ireg0 rs1 ireg0 rs2 - | Psraw(rd, rs1, rs2) -> - fprintf oc " sra%t %a, %a, %a\n" w ireg rd ireg0 rs1 ireg0 rs2 - - (* 64-bit integer register-immediate instructions *) - | Paddil (rd, rs, imm) -> assert Archi.ptr64; - fprintf oc " addi %a, %a, %a\n" ireg rd ireg0 rs coqint64 imm - | Psltil (rd, rs, imm) -> assert Archi.ptr64; - fprintf oc " slti %a, %a, %a\n" ireg rd ireg0 rs coqint64 imm - | Psltiul (rd, rs, imm) -> assert Archi.ptr64; - fprintf oc " sltiu %a, %a, %a\n" ireg rd ireg0 rs coqint64 imm - | Pandil (rd, rs, imm) -> assert Archi.ptr64; - fprintf oc " andi %a, %a, %a\n" ireg rd ireg0 rs coqint64 imm - | Poril (rd, rs, imm) -> assert Archi.ptr64; - fprintf oc " ori %a, %a, %a\n" ireg rd ireg0 rs coqint64 imm - | Pxoril (rd, rs, imm) -> assert Archi.ptr64; - fprintf oc " xori %a, %a, %a\n" ireg rd ireg0 rs coqint64 imm - | Psllil (rd, rs, imm) -> assert Archi.ptr64; - fprintf oc " slli %a, %a, %a\n" ireg rd ireg0 rs coqint64 imm - | Psrlil (rd, rs, imm) -> assert Archi.ptr64; - fprintf oc " srli %a, %a, %a\n" ireg rd ireg0 rs coqint64 imm - | Psrail (rd, rs, imm) -> assert Archi.ptr64; - fprintf oc " srai %a, %a, %a\n" ireg rd ireg0 rs coqint64 imm - | Pluil (rd, imm) -> assert Archi.ptr64; - fprintf oc " lui %a, %a\n" ireg rd coqint64 imm - - (* 64-bit integer register-register instructions *) - | Paddl(rd, rs1, rs2) -> assert Archi.ptr64; - fprintf oc " add %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 - | Psubl(rd, rs1, rs2) -> assert Archi.ptr64; - fprintf oc " sub %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 - - | Pmull(rd, rs1, rs2) -> assert Archi.ptr64; - fprintf oc " mul %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 - | Pmulhl(rd, rs1, rs2) -> assert Archi.ptr64; - fprintf oc " mulh %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 - | Pmulhul(rd, rs1, rs2) -> assert Archi.ptr64; - fprintf oc " mulhu %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 - - | Pdivl(rd, rs1, rs2) -> assert Archi.ptr64; - fprintf oc " div %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 - | Pdivul(rd, rs1, rs2) -> assert Archi.ptr64; - fprintf oc " divu %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 - | Preml(rd, rs1, rs2) -> assert Archi.ptr64; - fprintf oc " rem %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 - | Premul(rd, rs1, rs2) -> assert Archi.ptr64; - fprintf oc " remu %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 - - | Psltl(rd, rs1, rs2) -> assert Archi.ptr64; - fprintf oc " slt %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 - | Psltul(rd, rs1, rs2) -> assert Archi.ptr64; - fprintf oc " sltu %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 - - | Pandl(rd, rs1, rs2) -> assert Archi.ptr64; - fprintf oc " and %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 - | Porl(rd, rs1, rs2) -> assert Archi.ptr64; - fprintf oc " or %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 - | Pxorl(rd, rs1, rs2) -> assert Archi.ptr64; - fprintf oc " xor %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 - | Pslll(rd, rs1, rs2) -> assert Archi.ptr64; - fprintf oc " sll %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 - | Psrll(rd, rs1, rs2) -> assert Archi.ptr64; - fprintf oc " srl %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 - | Psral(rd, rs1, rs2) -> assert Archi.ptr64; - fprintf oc " sra %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 - - (* Unconditional jumps. Links are always to X1/RA. *) - (* TODO: fix up arguments for calls to variadics, to move *) - (* floating point arguments to integer registers. How? *) - | Pj_l(l) -> - fprintf oc " j %a\n" print_label l - | Pj_s(s, sg) -> - fprintf oc " j %a\n" symbol s - | Pj_r(r, sg) -> - fprintf oc " jr %a\n" ireg r - | Pjal_s(s, sg) -> - fprintf oc " call %a\n" symbol s - | Pjal_r(r, sg) -> - fprintf oc " jalr %a\n" ireg r - - (* Conditional branches, 32-bit comparisons *) - | Pbeqw(rs1, rs2, l) -> - fprintf oc " beq %a, %a, %a\n" ireg0 rs1 ireg0 rs2 print_label l - | Pbnew(rs1, rs2, l) -> - fprintf oc " bne %a, %a, %a\n" ireg0 rs1 ireg0 rs2 print_label l - | Pbltw(rs1, rs2, l) -> - fprintf oc " blt %a, %a, %a\n" ireg0 rs1 ireg0 rs2 print_label l - | Pbltuw(rs1, rs2, l) -> - fprintf oc " bltu %a, %a, %a\n" ireg0 rs1 ireg0 rs2 print_label l - | Pbgew(rs1, rs2, l) -> - fprintf oc " bge %a, %a, %a\n" ireg0 rs1 ireg0 rs2 print_label l - | Pbgeuw(rs1, rs2, l) -> - fprintf oc " bgeu %a, %a, %a\n" ireg0 rs1 ireg0 rs2 print_label l - - (* Conditional branches, 64-bit comparisons *) - | Pbeql(rs1, rs2, l) -> assert Archi.ptr64; - fprintf oc " beq %a, %a, %a\n" ireg0 rs1 ireg0 rs2 print_label l - | Pbnel(rs1, rs2, l) -> assert Archi.ptr64; - fprintf oc " bne %a, %a, %a\n" ireg0 rs1 ireg0 rs2 print_label l - | Pbltl(rs1, rs2, l) -> assert Archi.ptr64; - fprintf oc " blt %a, %a, %a\n" ireg0 rs1 ireg0 rs2 print_label l - | Pbltul(rs1, rs2, l) -> assert Archi.ptr64; - fprintf oc " bltu %a, %a, %a\n" ireg0 rs1 ireg0 rs2 print_label l - | Pbgel(rs1, rs2, l) -> assert Archi.ptr64; - fprintf oc " bge %a, %a, %a\n" ireg0 rs1 ireg0 rs2 print_label l - | Pbgeul(rs1, rs2, l) -> assert Archi.ptr64; - fprintf oc " bgeu %a, %a, %a\n" ireg0 rs1 ireg0 rs2 print_label l - - (* Loads and stores *) - | Plb(rd, ra, ofs) -> - fprintf oc " lb %a, %a(%a)\n" ireg rd offset ofs ireg ra - | Plbu(rd, ra, ofs) -> - fprintf oc " lbu %a, %a(%a)\n" ireg rd offset ofs ireg ra - | Plh(rd, ra, ofs) -> - fprintf oc " lh %a, %a(%a)\n" ireg rd offset ofs ireg ra - | Plhu(rd, ra, ofs) -> - fprintf oc " lhu %a, %a(%a)\n" ireg rd offset ofs ireg ra - | Plw(rd, ra, ofs) | Plw_a(rd, ra, ofs) -> - fprintf oc " lw %a, %a(%a)\n" ireg rd offset ofs ireg ra - | Pld(rd, ra, ofs) | Pld_a(rd, ra, ofs) -> assert Archi.ptr64; - fprintf oc " ld %a, %a(%a)\n" ireg rd offset ofs ireg ra - - | Psb(rd, ra, ofs) -> - fprintf oc " sb %a, %a(%a)\n" ireg rd offset ofs ireg ra - | Psh(rd, ra, ofs) -> - fprintf oc " sh %a, %a(%a)\n" ireg rd offset ofs ireg ra - | Psw(rd, ra, ofs) | Psw_a(rd, ra, ofs) -> - fprintf oc " sw %a, %a(%a)\n" ireg rd offset ofs ireg ra - | Psd(rd, ra, ofs) | Psd_a(rd, ra, ofs) -> assert Archi.ptr64; - fprintf oc " sd %a, %a(%a)\n" ireg rd offset ofs ireg ra - - - (* Synchronization *) - | Pfence -> - fprintf oc " fence\n" - - (* floating point register move. - fmv.d preserves single-precision register contents, and hence - is applicable to both single- and double-precision moves. - *) - | Pfmv (fd,fs) -> - fprintf oc " fmv.d %a, %a\n" freg fd freg fs - | Pfmvxs (rd,fs) -> - fprintf oc " fmv.x.s %a, %a\n" ireg rd freg fs - | Pfmvsx (fd,rs) -> - fprintf oc " fmv.s.x %a, %a\n" freg fd ireg rs - | Pfmvxd (rd,fs) -> - fprintf oc " fmv.x.d %a, %a\n" ireg rd freg fs -<<<<<<< HEAD - | Pfmvsx (fd,rs) -> - fprintf oc " fmv.s.x %a, %a\n" freg fd ireg rs -======= ->>>>>>> master - | Pfmvdx (fd,rs) -> - fprintf oc " fmv.d.x %a, %a\n" freg fd ireg rs - - (* 32-bit (single-precision) floating point *) - | Pfls (fd, ra, ofs) -> - fprintf oc " flw %a, %a(%a)\n" freg fd offset ofs ireg ra - | Pfss (fs, ra, ofs) -> - fprintf oc " fsw %a, %a(%a)\n" freg fs offset ofs ireg ra - - | Pfnegs (fd, fs) -> - fprintf oc " fneg.s %a, %a\n" freg fd freg fs - | Pfabss (fd, fs) -> - fprintf oc " fabs.s %a, %a\n" freg fd freg fs - - | Pfadds (fd, fs1, fs2) -> - fprintf oc " fadd.s %a, %a, %a\n" freg fd freg fs1 freg fs2 - | Pfsubs (fd, fs1, fs2) -> - fprintf oc " fsub.s %a, %a, %a\n" freg fd freg fs1 freg fs2 - | Pfmuls (fd, fs1, fs2) -> - fprintf oc " fmul.s %a, %a, %a\n" freg fd freg fs1 freg fs2 - | Pfdivs (fd, fs1, fs2) -> - fprintf oc " fdiv.s %a, %a, %a\n" freg fd freg fs1 freg fs2 - | Pfmins (fd, fs1, fs2) -> - fprintf oc " fmin.s %a, %a, %a\n" freg fd freg fs1 freg fs2 - | Pfmaxs (fd, fs1, fs2) -> - fprintf oc " fmax.s %a, %a, %a\n" freg fd freg fs1 freg fs2 - - | Pfeqs (rd, fs1, fs2) -> - fprintf oc " feq.s %a, %a, %a\n" ireg rd freg fs1 freg fs2 - | Pflts (rd, fs1, fs2) -> - fprintf oc " flt.s %a, %a, %a\n" ireg rd freg fs1 freg fs2 - | Pfles (rd, fs1, fs2) -> - fprintf oc " fle.s %a, %a, %a\n" ireg rd freg fs1 freg fs2 - - | Pfsqrts (fd, fs) -> - fprintf oc " fsqrt.s %a, %a\n" freg fd freg fs - - | Pfmadds (fd, fs1, fs2, fs3) -> - fprintf oc " fmadd.s %a, %a, %a, %a\n" freg fd freg fs1 freg fs2 freg fs3 - | Pfmsubs (fd, fs1, fs2, fs3) -> - fprintf oc " fmsub.s %a, %a, %a, %a\n" freg fd freg fs1 freg fs2 freg fs3 - | Pfnmadds (fd, fs1, fs2, fs3) -> - fprintf oc " fnmadd.s %a, %a, %a, %a\n" freg fd freg fs1 freg fs2 freg fs3 - | Pfnmsubs (fd, fs1, fs2, fs3) -> - fprintf oc " fnmsub.s %a, %a, %a, %a\n" freg fd freg fs1 freg fs2 freg fs3 - - | Pfcvtws (rd, fs) -> - fprintf oc " fcvt.w.s %a, %a, rtz\n" ireg rd freg fs - | Pfcvtwus (rd, fs) -> - fprintf oc " fcvt.wu.s %a, %a, rtz\n" ireg rd freg fs - | Pfcvtsw (fd, rs) -> - fprintf oc " fcvt.s.w %a, %a\n" freg fd ireg0 rs - | Pfcvtswu (fd, rs) -> - fprintf oc " fcvt.s.wu %a, %a\n" freg fd ireg0 rs - - | Pfcvtls (rd, fs) -> assert Archi.ptr64; - fprintf oc " fcvt.l.s %a, %a, rtz\n" ireg rd freg fs - | Pfcvtlus (rd, fs) -> assert Archi.ptr64; - fprintf oc " fcvt.lu.s %a, %a, rtz\n" ireg rd freg fs - | Pfcvtsl (fd, rs) -> assert Archi.ptr64; - fprintf oc " fcvt.s.l %a, %a\n" freg fd ireg0 rs - | Pfcvtslu (fd, rs) -> assert Archi.ptr64; - fprintf oc " fcvt.s.lu %a, %a\n" freg fd ireg0 rs - - (* 64-bit (double-precision) floating point *) - | Pfld (fd, ra, ofs) | Pfld_a (fd, ra, ofs) -> - fprintf oc " fld %a, %a(%a)\n" freg fd offset ofs ireg ra - | Pfsd (fs, ra, ofs) | Pfsd_a (fs, ra, ofs) -> - fprintf oc " fsd %a, %a(%a)\n" freg fs offset ofs ireg ra - - | Pfnegd (fd, fs) -> - fprintf oc " fneg.d %a, %a\n" freg fd freg fs - | Pfabsd (fd, fs) -> - fprintf oc " fabs.d %a, %a\n" freg fd freg fs - - | Pfaddd (fd, fs1, fs2) -> - fprintf oc " fadd.d %a, %a, %a\n" freg fd freg fs1 freg fs2 - | Pfsubd (fd, fs1, fs2) -> - fprintf oc " fsub.d %a, %a, %a\n" freg fd freg fs1 freg fs2 - | Pfmuld (fd, fs1, fs2) -> - fprintf oc " fmul.d %a, %a, %a\n" freg fd freg fs1 freg fs2 - | Pfdivd (fd, fs1, fs2) -> - fprintf oc " fdiv.d %a, %a, %a\n" freg fd freg fs1 freg fs2 - | Pfmind (fd, fs1, fs2) -> - fprintf oc " fmin.d %a, %a, %a\n" freg fd freg fs1 freg fs2 - | Pfmaxd (fd, fs1, fs2) -> - fprintf oc " fmax.d %a, %a, %a\n" freg fd freg fs1 freg fs2 - - | Pfeqd (rd, fs1, fs2) -> - fprintf oc " feq.d %a, %a, %a\n" ireg rd freg fs1 freg fs2 - | Pfltd (rd, fs1, fs2) -> - fprintf oc " flt.d %a, %a, %a\n" ireg rd freg fs1 freg fs2 - | Pfled (rd, fs1, fs2) -> - fprintf oc " fle.d %a, %a, %a\n" ireg rd freg fs1 freg fs2 - - | Pfsqrtd (fd, fs) -> - fprintf oc " fsqrt.d %a, %a\n" freg fd freg fs - - | Pfmaddd (fd, fs1, fs2, fs3) -> - fprintf oc " fmadd.d %a, %a, %a, %a\n" freg fd freg fs1 freg fs2 freg fs3 - | Pfmsubd (fd, fs1, fs2, fs3) -> - fprintf oc " fmsub.d %a, %a, %a, %a\n" freg fd freg fs1 freg fs2 freg fs3 - | Pfnmaddd (fd, fs1, fs2, fs3) -> - fprintf oc " fnmadd.d %a, %a, %a, %a\n" freg fd freg fs1 freg fs2 freg fs3 - | Pfnmsubd (fd, fs1, fs2, fs3) -> - fprintf oc " fnmsub.d %a, %a, %a, %a\n" freg fd freg fs1 freg fs2 freg fs3 - - | Pfcvtwd (rd, fs) -> - fprintf oc " fcvt.w.d %a, %a, rtz\n" ireg rd freg fs - | Pfcvtwud (rd, fs) -> - fprintf oc " fcvt.wu.d %a, %a, rtz\n" ireg rd freg fs - | Pfcvtdw (fd, rs) -> - fprintf oc " fcvt.d.w %a, %a\n" freg fd ireg0 rs - | Pfcvtdwu (fd, rs) -> - fprintf oc " fcvt.d.wu %a, %a\n" freg fd ireg0 rs - - | Pfcvtld (rd, fs) -> assert Archi.ptr64; - fprintf oc " fcvt.l.d %a, %a, rtz\n" ireg rd freg fs - | Pfcvtlud (rd, fs) -> assert Archi.ptr64; - fprintf oc " fcvt.lu.d %a, %a, rtz\n" ireg rd freg fs - | Pfcvtdl (fd, rs) -> assert Archi.ptr64; - fprintf oc " fcvt.d.l %a, %a\n" freg fd ireg0 rs - | Pfcvtdlu (fd, rs) -> assert Archi.ptr64; - fprintf oc " fcvt.d.lu %a, %a\n" freg fd ireg0 rs - - | Pfcvtds (fd, fs) -> - fprintf oc " fcvt.d.s %a, %a\n" freg fd freg fs - | Pfcvtsd (fd, fs) -> - fprintf oc " fcvt.s.d %a, %a\n" freg fd freg fs - - (* Pseudo-instructions expanded in Asmexpand *) - | Pselectl(_, _, _, _) -> - assert false - | Pallocframe(sz, ofs) -> - assert false - | Pfreeframe(sz, ofs) -> - assert false - | Pseqw _ | Psnew _ | Pseql _ | Psnel _ | Pcvtl2w _ | Pcvtw2l _ -> - assert false - - (* Pseudo-instructions that remain *) - | Plabel lbl -> - fprintf oc "%a:\n" print_label lbl - | Ploadsymbol(rd, id, ofs) -> - loadsymbol oc rd id ofs - | Ploadsymbol_high(rd, id, ofs) -> - fprintf oc " lui %a, %%hi(%a)\n" ireg rd symbol_offset (id, ofs) - | Ploadli(rd, n) -> - let d = camlint64_of_coqint n in - let lbl = label_literal64 d in - fprintf oc " ld %a, %a %s %Lx\n" ireg rd label lbl comment d - | Ploadfi(rd, f) -> - let d = camlint64_of_coqint(Floats.Float.to_bits f) in - let lbl = label_literal64 d in - fprintf oc " fld %a, %a, x31 %s %.18g\n" - freg rd label lbl comment (camlfloat_of_coqfloat f) - | Ploadsi(rd, f) -> - let s = camlint_of_coqint(Floats.Float32.to_bits f) in - let lbl = label_literal32 s in - fprintf oc " flw %a, %a, x31 %s %.18g\n" - freg rd label lbl comment (camlfloat_of_coqfloat32 f) - | Pbtbl(r, tbl) -> - let lbl = new_label() in - fprintf oc "%s jumptable [ " comment; - List.iter (fun l -> fprintf oc "%a " print_label l) tbl; - fprintf oc "]\n"; - fprintf oc " sll x5, %a, 2\n" ireg r; - fprintf oc " la x31, %a\n" label lbl; - fprintf oc " add x5, x31, x5\n"; - fprintf oc " lw x5, 0(x5)\n"; - fprintf oc " add x5, x31, x5\n"; - fprintf oc " jr x5\n"; - jumptables := (lbl, tbl) :: !jumptables; - fprintf oc "%s end pseudoinstr btbl\n" comment - | Pnop -> - fprintf oc " nop\n" - | Pbuiltin(ef, args, res) -> - begin match ef with - | EF_annot(kind,txt, targs) -> - begin match (P.to_int kind) with - | 1 -> let annot = annot_text preg_annot "x2" (camlstring_of_coqstring txt) args in - fprintf oc "%s annotation: %S\n" comment annot - | 2 -> let lbl = new_label () in - fprintf oc "%a:\n" label lbl; - add_ais_annot lbl preg_annot "x2" (camlstring_of_coqstring txt) args - | _ -> assert false - end - | EF_debug(kind, txt, targs) -> - print_debug_info comment print_file_line preg_annot "sp" oc - (P.to_int kind) (extern_atom txt) args - | EF_inline_asm(txt, sg, clob) -> - fprintf oc "%s begin inline assembly\n\t" comment; - print_inline_asm preg_asm oc (camlstring_of_coqstring txt) sg args res; - fprintf oc "%s end inline assembly\n" comment - | _ -> - assert false - end - - let get_section_names name = - let (text, lit) = - match C2C.atom_sections name with - | t :: l :: _ -> (t, l) - | _ -> (Section_text, Section_literal) in - text,lit,Section_jumptable - - let print_align oc alignment = - fprintf oc " .balign %d\n" alignment - - let print_jumptable oc jmptbl = - let print_tbl oc (lbl, tbl) = - fprintf oc "%a:\n" label lbl; - List.iter - (fun l -> fprintf oc " .long %a - %a\n" - print_label l label lbl) - tbl in - if !jumptables <> [] then - begin - section oc jmptbl; - fprintf oc " .balign 4\n"; - List.iter (print_tbl oc) !jumptables; - jumptables := [] - end - - let print_fun_info = elf_print_fun_info - - let print_optional_fun_info _ = () - - let print_var_info = elf_print_var_info - - let print_comm_symb oc sz name align = - if C2C.atom_is_static name then - fprintf oc " .local %a\n" symbol name; - fprintf oc " .comm %a, %s, %d\n" - symbol name - (Z.to_string sz) - align - - let print_instructions oc fn = - current_function_sig := fn.fn_sig; - List.iter (print_instruction oc) fn.fn_code - - -(* Data *) - - let address = if Archi.ptr64 then ".quad" else ".long" - - let print_prologue oc = - fprintf oc " .option %s\n" (if Archi.pic_code() then "pic" else "nopic"); - if !Clflags.option_g then begin - section oc Section_text; - end - - let print_epilogue oc = - if !Clflags.option_g then begin - Debug.compute_gnu_file_enum (fun f -> ignore (print_file oc f)); - section oc Section_text; - end - - let default_falignment = 2 - - let cfi_startproc oc = () - let cfi_endproc oc = () - - end - -let sel_target () = - (module Target:TARGET) diff --git a/riscV/TargetPrinter.ml b/riscV/TargetPrinter.ml new file mode 100644 index 00000000..aab6b9b8 --- /dev/null +++ b/riscV/TargetPrinter.ml @@ -0,0 +1,667 @@ +(* *********************************************************************) +(* *) +(* The Compcert verified compiler *) +(* *) +(* Xavier Leroy, INRIA Paris-Rocquencourt *) +(* Prashanth Mundkur, SRI International *) +(* *) +(* 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. *) +(* *) +(* The contributions by Prashanth Mundkur are reused and adapted *) +(* under the terms of a Contributor License Agreement between *) +(* SRI International and INRIA. *) +(* *) +(* *********************************************************************) + +(* Printing RISC-V assembly code in asm syntax *) + +open Printf +open Camlcoq +open Sections +open AST +open Asm +open AisAnnot +open PrintAsmaux +open Fileinfo + +(* Module containing the printing functions *) + +module Target : TARGET = + struct + +(* Basic printing functions *) + + let comment = "#" + + let symbol = elf_symbol + let symbol_offset = elf_symbol_offset + let label = elf_label + + let print_label oc lbl = label oc (transl_label lbl) + + let use_abi_name = false + + let int_reg_num_name = function + | X1 -> "x1" | X2 -> "x2" | X3 -> "x3" + | X4 -> "x4" | X5 -> "x5" | X6 -> "x6" | X7 -> "x7" + | X8 -> "x8" | X9 -> "x9" | X10 -> "x10" | X11 -> "x11" + | X12 -> "x12" | X13 -> "x13" | X14 -> "x14" | X15 -> "x15" + | X16 -> "x16" | X17 -> "x17" | X18 -> "x18" | X19 -> "x19" + | X20 -> "x20" | X21 -> "x21" | X22 -> "x22" | X23 -> "x23" + | X24 -> "x24" | X25 -> "x25" | X26 -> "x26" | X27 -> "x27" + | X28 -> "x28" | X29 -> "x29" | X30 -> "x30" | X31 -> "x31" + + let int_reg_abi_name = function + | X1 -> "ra" | X2 -> "sp" | X3 -> "gp" + | X4 -> "tp" | X5 -> "t0" | X6 -> "t1" | X7 -> "t2" + | X8 -> "s0" | X9 -> "s1" | X10 -> "a0" | X11 -> "a1" + | X12 -> "a2" | X13 -> "a3" | X14 -> "a4" | X15 -> "a5" + | X16 -> "a6" | X17 -> "a7" | X18 -> "s2" | X19 -> "s3" + | X20 -> "s4" | X21 -> "s5" | X22 -> "s6" | X23 -> "s7" + | X24 -> "s8" | X25 -> "s9" | X26 -> "s10" | X27 -> "s11" + | X28 -> "t3" | X29 -> "t4" | X30 -> "t5" | X31 -> "t6" + + let float_reg_num_name = function + | F0 -> "f0" | F1 -> "f1" | F2 -> "f2" | F3 -> "f3" + | F4 -> "f4" | F5 -> "f5" | F6 -> "f6" | F7 -> "f7" + | F8 -> "f8" | F9 -> "f9" | F10 -> "f10" | F11 -> "f11" + | F12 -> "f12" | F13 -> "f13" | F14 -> "f14" | F15 -> "f15" + | F16 -> "f16" | F17 -> "f17" | F18 -> "f18" | F19 -> "f19" + | F20 -> "f20" | F21 -> "f21" | F22 -> "f22" | F23 -> "f23" + | F24 -> "f24" | F25 -> "f25" | F26 -> "f26" | F27 -> "f27" + | F28 -> "f28" | F29 -> "f29" | F30 -> "f30" | F31 -> "f31" + + let float_reg_abi_name = function + | F0 -> "ft0" | F1 -> "ft1" | F2 -> "ft2" | F3 -> "ft3" + | F4 -> "ft4" | F5 -> "ft5" | F6 -> "ft6" | F7 -> "ft7" + | F8 -> "fs0" | F9 -> "fs1" | F10 -> "fa0" | F11 -> "fa1" + | F12 -> "fa2" | F13 -> "fa3" | F14 -> "fa4" | F15 -> "fa5" + | F16 -> "fa6" | F17 -> "fa7" | F18 -> "fs2" | F19 -> "fs3" + | F20 -> "fs4" | F21 -> "fs5" | F22 -> "fs6" | F23 -> "fs7" + | F24 -> "fs8" | F25 -> "fs9" | F26 ->"fs10" | F27 -> "fs11" + | F28 -> "ft3" | F29 -> "ft4" | F30 -> "ft5" | F31 -> "ft6" + + let int_reg_name = if use_abi_name then int_reg_abi_name else int_reg_num_name + let float_reg_name = if use_abi_name then float_reg_abi_name else float_reg_num_name + + let ireg oc r = output_string oc (int_reg_name r) + let freg oc r = output_string oc (float_reg_name r) + + let ireg0 oc = function + | X0 -> output_string oc "x0" + | X r -> ireg oc r + + let preg_asm oc ty = function + | IR r -> ireg oc r + | FR r -> freg oc r + | _ -> assert false + + let preg_annot = function + | IR r -> int_reg_name r + | FR r -> float_reg_name r + | _ -> assert false + +(* Names of sections *) + + let name_of_section = function + | Section_text -> ".text" + | Section_data(i, true) -> + failwith "_Thread_local unsupported on this platform" + | Section_data(i, false) | Section_small_data i -> + variable_section ~sec:".data" ~bss:".bss" i + | Section_const i | Section_small_const i -> + variable_section ~sec:".section .rodata" i + | Section_string -> ".section .rodata" + | Section_literal -> ".section .rodata" + | Section_jumptable -> ".section .rodata" + | Section_debug_info _ -> ".section .debug_info,\"\",%progbits" + | Section_debug_loc -> ".section .debug_loc,\"\",%progbits" + | Section_debug_abbrev -> ".section .debug_abbrev,\"\",%progbits" + | Section_debug_line _ -> ".section .debug_line,\"\",%progbits" + | Section_debug_ranges -> ".section .debug_ranges,\"\",%progbits" + | Section_debug_str -> ".section .debug_str,\"MS\",%progbits,1" + | Section_user(s, wr, ex) -> + sprintf ".section \"%s\",\"a%s%s\",%%progbits" + s (if wr then "w" else "") (if ex then "x" else "") + | Section_ais_annotation -> sprintf ".section \"__compcert_ais_annotations\",\"\",@note" + + let section oc sec = + fprintf oc " %s\n" (name_of_section sec) + +(* Associate labels to floating-point constants and to symbols. *) + + let emit_constants oc lit = + if exists_constants () then begin + section oc lit; + if Hashtbl.length literal64_labels > 0 then + begin + fprintf oc " .align 3\n"; + Hashtbl.iter + (fun bf lbl -> fprintf oc "%a: .quad 0x%Lx\n" label lbl bf) + literal64_labels + end; + if Hashtbl.length literal32_labels > 0 then + begin + fprintf oc " .align 2\n"; + Hashtbl.iter + (fun bf lbl -> + fprintf oc "%a: .long 0x%lx\n" label lbl bf) + literal32_labels + end; + reset_literals () + end + +(* Generate code to load the address of id + ofs in register r *) + + let loadsymbol oc r id ofs = + if Archi.pic_code () then begin + assert (ofs = Integers.Ptrofs.zero); + fprintf oc " la %a, %s\n" ireg r (extern_atom id) + end else begin + fprintf oc " lui %a, %%hi(%a)\n" + ireg r symbol_offset (id, ofs); + fprintf oc " addi %a, %a, %%lo(%a)\n" + ireg r ireg r symbol_offset (id, ofs) + end + +(* Emit .file / .loc debugging directives *) + + let print_file_line oc file line = + print_file_line oc comment file line + +(* + let print_location oc loc = + if loc <> Cutil.no_loc then print_file_line oc (fst loc) (snd loc) +*) + +(* Add "w" suffix to 32-bit instructions if we are in 64-bit mode *) + + let w oc = + if Archi.ptr64 then output_string oc "w" + +(* Offset part of a load or store *) + + let offset oc = function + | Ofsimm n -> ptrofs oc n + | Ofslow(id, ofs) -> fprintf oc "%%lo(%a)" symbol_offset (id, ofs) + +(* Printing of instructions *) + let print_instruction oc = function + | Pmv(rd, rs) -> + fprintf oc " mv %a, %a\n" ireg rd ireg rs + + (* 32-bit integer register-immediate instructions *) + | Paddiw (rd, rs, imm) -> + fprintf oc " addi%t %a, %a, %a\n" w ireg rd ireg0 rs coqint imm + | Psltiw (rd, rs, imm) -> + fprintf oc " slti %a, %a, %a\n" ireg rd ireg0 rs coqint imm + | Psltiuw (rd, rs, imm) -> + fprintf oc " sltiu %a, %a, %a\n" ireg rd ireg0 rs coqint imm + | Pandiw (rd, rs, imm) -> + fprintf oc " andi %a, %a, %a\n" ireg rd ireg0 rs coqint imm + | Poriw (rd, rs, imm) -> + fprintf oc " ori %a, %a, %a\n" ireg rd ireg0 rs coqint imm + | Pxoriw (rd, rs, imm) -> + fprintf oc " xori %a, %a, %a\n" ireg rd ireg0 rs coqint imm + | Pslliw (rd, rs, imm) -> + fprintf oc " slli%t %a, %a, %a\n" w ireg rd ireg0 rs coqint imm + | Psrliw (rd, rs, imm) -> + fprintf oc " srli%t %a, %a, %a\n" w ireg rd ireg0 rs coqint imm + | Psraiw (rd, rs, imm) -> + fprintf oc " srai%t %a, %a, %a\n" w ireg rd ireg0 rs coqint imm + | Pluiw (rd, imm) -> + fprintf oc " lui %a, %a\n" ireg rd coqint imm + + (* 32-bit integer register-register instructions *) + | Paddw(rd, rs1, rs2) -> + fprintf oc " add%t %a, %a, %a\n" w ireg rd ireg0 rs1 ireg0 rs2 + | Psubw(rd, rs1, rs2) -> + fprintf oc " sub%t %a, %a, %a\n" w ireg rd ireg0 rs1 ireg0 rs2 + + | Pmulw(rd, rs1, rs2) -> + fprintf oc " mul%t %a, %a, %a\n" w ireg rd ireg0 rs1 ireg0 rs2 + | Pmulhw(rd, rs1, rs2) -> assert (not Archi.ptr64); + fprintf oc " mulh %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 + | Pmulhuw(rd, rs1, rs2) -> assert (not Archi.ptr64); + fprintf oc " mulhu %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 + + | Pdivw(rd, rs1, rs2) -> + fprintf oc " div%t %a, %a, %a\n" w ireg rd ireg0 rs1 ireg0 rs2 + | Pdivuw(rd, rs1, rs2) -> + fprintf oc " divu%t %a, %a, %a\n" w ireg rd ireg0 rs1 ireg0 rs2 + | Premw(rd, rs1, rs2) -> + fprintf oc " rem%t %a, %a, %a\n" w ireg rd ireg0 rs1 ireg0 rs2 + | Premuw(rd, rs1, rs2) -> + fprintf oc " remu%t %a, %a, %a\n" w ireg rd ireg0 rs1 ireg0 rs2 + + | Psltw(rd, rs1, rs2) -> + fprintf oc " slt %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 + | Psltuw(rd, rs1, rs2) -> + fprintf oc " sltu %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 + + | Pandw(rd, rs1, rs2) -> + fprintf oc " and %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 + | Porw(rd, rs1, rs2) -> + fprintf oc " or %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 + | Pxorw(rd, rs1, rs2) -> + fprintf oc " xor %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 + | Psllw(rd, rs1, rs2) -> + fprintf oc " sll%t %a, %a, %a\n" w ireg rd ireg0 rs1 ireg0 rs2 + | Psrlw(rd, rs1, rs2) -> + fprintf oc " srl%t %a, %a, %a\n" w ireg rd ireg0 rs1 ireg0 rs2 + | Psraw(rd, rs1, rs2) -> + fprintf oc " sra%t %a, %a, %a\n" w ireg rd ireg0 rs1 ireg0 rs2 + + (* 64-bit integer register-immediate instructions *) + | Paddil (rd, rs, imm) -> assert Archi.ptr64; + fprintf oc " addi %a, %a, %a\n" ireg rd ireg0 rs coqint64 imm + | Psltil (rd, rs, imm) -> assert Archi.ptr64; + fprintf oc " slti %a, %a, %a\n" ireg rd ireg0 rs coqint64 imm + | Psltiul (rd, rs, imm) -> assert Archi.ptr64; + fprintf oc " sltiu %a, %a, %a\n" ireg rd ireg0 rs coqint64 imm + | Pandil (rd, rs, imm) -> assert Archi.ptr64; + fprintf oc " andi %a, %a, %a\n" ireg rd ireg0 rs coqint64 imm + | Poril (rd, rs, imm) -> assert Archi.ptr64; + fprintf oc " ori %a, %a, %a\n" ireg rd ireg0 rs coqint64 imm + | Pxoril (rd, rs, imm) -> assert Archi.ptr64; + fprintf oc " xori %a, %a, %a\n" ireg rd ireg0 rs coqint64 imm + | Psllil (rd, rs, imm) -> assert Archi.ptr64; + fprintf oc " slli %a, %a, %a\n" ireg rd ireg0 rs coqint64 imm + | Psrlil (rd, rs, imm) -> assert Archi.ptr64; + fprintf oc " srli %a, %a, %a\n" ireg rd ireg0 rs coqint64 imm + | Psrail (rd, rs, imm) -> assert Archi.ptr64; + fprintf oc " srai %a, %a, %a\n" ireg rd ireg0 rs coqint64 imm + | Pluil (rd, imm) -> assert Archi.ptr64; + fprintf oc " lui %a, %a\n" ireg rd coqint64 imm + + (* 64-bit integer register-register instructions *) + | Paddl(rd, rs1, rs2) -> assert Archi.ptr64; + fprintf oc " add %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 + | Psubl(rd, rs1, rs2) -> assert Archi.ptr64; + fprintf oc " sub %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 + + | Pmull(rd, rs1, rs2) -> assert Archi.ptr64; + fprintf oc " mul %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 + | Pmulhl(rd, rs1, rs2) -> assert Archi.ptr64; + fprintf oc " mulh %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 + | Pmulhul(rd, rs1, rs2) -> assert Archi.ptr64; + fprintf oc " mulhu %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 + + | Pdivl(rd, rs1, rs2) -> assert Archi.ptr64; + fprintf oc " div %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 + | Pdivul(rd, rs1, rs2) -> assert Archi.ptr64; + fprintf oc " divu %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 + | Preml(rd, rs1, rs2) -> assert Archi.ptr64; + fprintf oc " rem %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 + | Premul(rd, rs1, rs2) -> assert Archi.ptr64; + fprintf oc " remu %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 + + | Psltl(rd, rs1, rs2) -> assert Archi.ptr64; + fprintf oc " slt %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 + | Psltul(rd, rs1, rs2) -> assert Archi.ptr64; + fprintf oc " sltu %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 + + | Pandl(rd, rs1, rs2) -> assert Archi.ptr64; + fprintf oc " and %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 + | Porl(rd, rs1, rs2) -> assert Archi.ptr64; + fprintf oc " or %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 + | Pxorl(rd, rs1, rs2) -> assert Archi.ptr64; + fprintf oc " xor %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 + | Pslll(rd, rs1, rs2) -> assert Archi.ptr64; + fprintf oc " sll %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 + | Psrll(rd, rs1, rs2) -> assert Archi.ptr64; + fprintf oc " srl %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 + | Psral(rd, rs1, rs2) -> assert Archi.ptr64; + fprintf oc " sra %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2 + + (* Unconditional jumps. Links are always to X1/RA. *) + (* TODO: fix up arguments for calls to variadics, to move *) + (* floating point arguments to integer registers. How? *) + | Pj_l(l) -> + fprintf oc " j %a\n" print_label l + | Pj_s(s, sg) -> + fprintf oc " j %a\n" symbol s + | Pj_r(r, sg) -> + fprintf oc " jr %a\n" ireg r + | Pjal_s(s, sg) -> + fprintf oc " call %a\n" symbol s + | Pjal_r(r, sg) -> + fprintf oc " jalr %a\n" ireg r + + (* Conditional branches, 32-bit comparisons *) + | Pbeqw(rs1, rs2, l) -> + fprintf oc " beq %a, %a, %a\n" ireg0 rs1 ireg0 rs2 print_label l + | Pbnew(rs1, rs2, l) -> + fprintf oc " bne %a, %a, %a\n" ireg0 rs1 ireg0 rs2 print_label l + | Pbltw(rs1, rs2, l) -> + fprintf oc " blt %a, %a, %a\n" ireg0 rs1 ireg0 rs2 print_label l + | Pbltuw(rs1, rs2, l) -> + fprintf oc " bltu %a, %a, %a\n" ireg0 rs1 ireg0 rs2 print_label l + | Pbgew(rs1, rs2, l) -> + fprintf oc " bge %a, %a, %a\n" ireg0 rs1 ireg0 rs2 print_label l + | Pbgeuw(rs1, rs2, l) -> + fprintf oc " bgeu %a, %a, %a\n" ireg0 rs1 ireg0 rs2 print_label l + + (* Conditional branches, 64-bit comparisons *) + | Pbeql(rs1, rs2, l) -> assert Archi.ptr64; + fprintf oc " beq %a, %a, %a\n" ireg0 rs1 ireg0 rs2 print_label l + | Pbnel(rs1, rs2, l) -> assert Archi.ptr64; + fprintf oc " bne %a, %a, %a\n" ireg0 rs1 ireg0 rs2 print_label l + | Pbltl(rs1, rs2, l) -> assert Archi.ptr64; + fprintf oc " blt %a, %a, %a\n" ireg0 rs1 ireg0 rs2 print_label l + | Pbltul(rs1, rs2, l) -> assert Archi.ptr64; + fprintf oc " bltu %a, %a, %a\n" ireg0 rs1 ireg0 rs2 print_label l + | Pbgel(rs1, rs2, l) -> assert Archi.ptr64; + fprintf oc " bge %a, %a, %a\n" ireg0 rs1 ireg0 rs2 print_label l + | Pbgeul(rs1, rs2, l) -> assert Archi.ptr64; + fprintf oc " bgeu %a, %a, %a\n" ireg0 rs1 ireg0 rs2 print_label l + + (* Loads and stores *) + | Plb(rd, ra, ofs) -> + fprintf oc " lb %a, %a(%a)\n" ireg rd offset ofs ireg ra + | Plbu(rd, ra, ofs) -> + fprintf oc " lbu %a, %a(%a)\n" ireg rd offset ofs ireg ra + | Plh(rd, ra, ofs) -> + fprintf oc " lh %a, %a(%a)\n" ireg rd offset ofs ireg ra + | Plhu(rd, ra, ofs) -> + fprintf oc " lhu %a, %a(%a)\n" ireg rd offset ofs ireg ra + | Plw(rd, ra, ofs) | Plw_a(rd, ra, ofs) -> + fprintf oc " lw %a, %a(%a)\n" ireg rd offset ofs ireg ra + | Pld(rd, ra, ofs) | Pld_a(rd, ra, ofs) -> assert Archi.ptr64; + fprintf oc " ld %a, %a(%a)\n" ireg rd offset ofs ireg ra + + | Psb(rd, ra, ofs) -> + fprintf oc " sb %a, %a(%a)\n" ireg rd offset ofs ireg ra + | Psh(rd, ra, ofs) -> + fprintf oc " sh %a, %a(%a)\n" ireg rd offset ofs ireg ra + | Psw(rd, ra, ofs) | Psw_a(rd, ra, ofs) -> + fprintf oc " sw %a, %a(%a)\n" ireg rd offset ofs ireg ra + | Psd(rd, ra, ofs) | Psd_a(rd, ra, ofs) -> assert Archi.ptr64; + fprintf oc " sd %a, %a(%a)\n" ireg rd offset ofs ireg ra + + + (* Synchronization *) + | Pfence -> + fprintf oc " fence\n" + + (* floating point register move. + fmv.d preserves single-precision register contents, and hence + is applicable to both single- and double-precision moves. + *) + | Pfmv (fd,fs) -> + fprintf oc " fmv.d %a, %a\n" freg fd freg fs + | Pfmvxs (rd,fs) -> + fprintf oc " fmv.x.s %a, %a\n" ireg rd freg fs + | Pfmvsx (fd,rs) -> + fprintf oc " fmv.s.x %a, %a\n" freg fd ireg rs + | Pfmvxd (rd,fs) -> + fprintf oc " fmv.x.d %a, %a\n" ireg rd freg fs + | Pfmvdx (fd,rs) -> + fprintf oc " fmv.d.x %a, %a\n" freg fd ireg rs + + (* 32-bit (single-precision) floating point *) + | Pfls (fd, ra, ofs) -> + fprintf oc " flw %a, %a(%a)\n" freg fd offset ofs ireg ra + | Pfss (fs, ra, ofs) -> + fprintf oc " fsw %a, %a(%a)\n" freg fs offset ofs ireg ra + + | Pfnegs (fd, fs) -> + fprintf oc " fneg.s %a, %a\n" freg fd freg fs + | Pfabss (fd, fs) -> + fprintf oc " fabs.s %a, %a\n" freg fd freg fs + + | Pfadds (fd, fs1, fs2) -> + fprintf oc " fadd.s %a, %a, %a\n" freg fd freg fs1 freg fs2 + | Pfsubs (fd, fs1, fs2) -> + fprintf oc " fsub.s %a, %a, %a\n" freg fd freg fs1 freg fs2 + | Pfmuls (fd, fs1, fs2) -> + fprintf oc " fmul.s %a, %a, %a\n" freg fd freg fs1 freg fs2 + | Pfdivs (fd, fs1, fs2) -> + fprintf oc " fdiv.s %a, %a, %a\n" freg fd freg fs1 freg fs2 + | Pfmins (fd, fs1, fs2) -> + fprintf oc " fmin.s %a, %a, %a\n" freg fd freg fs1 freg fs2 + | Pfmaxs (fd, fs1, fs2) -> + fprintf oc " fmax.s %a, %a, %a\n" freg fd freg fs1 freg fs2 + + | Pfeqs (rd, fs1, fs2) -> + fprintf oc " feq.s %a, %a, %a\n" ireg rd freg fs1 freg fs2 + | Pflts (rd, fs1, fs2) -> + fprintf oc " flt.s %a, %a, %a\n" ireg rd freg fs1 freg fs2 + | Pfles (rd, fs1, fs2) -> + fprintf oc " fle.s %a, %a, %a\n" ireg rd freg fs1 freg fs2 + + | Pfsqrts (fd, fs) -> + fprintf oc " fsqrt.s %a, %a\n" freg fd freg fs + + | Pfmadds (fd, fs1, fs2, fs3) -> + fprintf oc " fmadd.s %a, %a, %a, %a\n" freg fd freg fs1 freg fs2 freg fs3 + | Pfmsubs (fd, fs1, fs2, fs3) -> + fprintf oc " fmsub.s %a, %a, %a, %a\n" freg fd freg fs1 freg fs2 freg fs3 + | Pfnmadds (fd, fs1, fs2, fs3) -> + fprintf oc " fnmadd.s %a, %a, %a, %a\n" freg fd freg fs1 freg fs2 freg fs3 + | Pfnmsubs (fd, fs1, fs2, fs3) -> + fprintf oc " fnmsub.s %a, %a, %a, %a\n" freg fd freg fs1 freg fs2 freg fs3 + + | Pfcvtws (rd, fs) -> + fprintf oc " fcvt.w.s %a, %a, rtz\n" ireg rd freg fs + | Pfcvtwus (rd, fs) -> + fprintf oc " fcvt.wu.s %a, %a, rtz\n" ireg rd freg fs + | Pfcvtsw (fd, rs) -> + fprintf oc " fcvt.s.w %a, %a\n" freg fd ireg0 rs + | Pfcvtswu (fd, rs) -> + fprintf oc " fcvt.s.wu %a, %a\n" freg fd ireg0 rs + + | Pfcvtls (rd, fs) -> assert Archi.ptr64; + fprintf oc " fcvt.l.s %a, %a, rtz\n" ireg rd freg fs + | Pfcvtlus (rd, fs) -> assert Archi.ptr64; + fprintf oc " fcvt.lu.s %a, %a, rtz\n" ireg rd freg fs + | Pfcvtsl (fd, rs) -> assert Archi.ptr64; + fprintf oc " fcvt.s.l %a, %a\n" freg fd ireg0 rs + | Pfcvtslu (fd, rs) -> assert Archi.ptr64; + fprintf oc " fcvt.s.lu %a, %a\n" freg fd ireg0 rs + + (* 64-bit (double-precision) floating point *) + | Pfld (fd, ra, ofs) | Pfld_a (fd, ra, ofs) -> + fprintf oc " fld %a, %a(%a)\n" freg fd offset ofs ireg ra + | Pfsd (fs, ra, ofs) | Pfsd_a (fs, ra, ofs) -> + fprintf oc " fsd %a, %a(%a)\n" freg fs offset ofs ireg ra + + | Pfnegd (fd, fs) -> + fprintf oc " fneg.d %a, %a\n" freg fd freg fs + | Pfabsd (fd, fs) -> + fprintf oc " fabs.d %a, %a\n" freg fd freg fs + + | Pfaddd (fd, fs1, fs2) -> + fprintf oc " fadd.d %a, %a, %a\n" freg fd freg fs1 freg fs2 + | Pfsubd (fd, fs1, fs2) -> + fprintf oc " fsub.d %a, %a, %a\n" freg fd freg fs1 freg fs2 + | Pfmuld (fd, fs1, fs2) -> + fprintf oc " fmul.d %a, %a, %a\n" freg fd freg fs1 freg fs2 + | Pfdivd (fd, fs1, fs2) -> + fprintf oc " fdiv.d %a, %a, %a\n" freg fd freg fs1 freg fs2 + | Pfmind (fd, fs1, fs2) -> + fprintf oc " fmin.d %a, %a, %a\n" freg fd freg fs1 freg fs2 + | Pfmaxd (fd, fs1, fs2) -> + fprintf oc " fmax.d %a, %a, %a\n" freg fd freg fs1 freg fs2 + + | Pfeqd (rd, fs1, fs2) -> + fprintf oc " feq.d %a, %a, %a\n" ireg rd freg fs1 freg fs2 + | Pfltd (rd, fs1, fs2) -> + fprintf oc " flt.d %a, %a, %a\n" ireg rd freg fs1 freg fs2 + | Pfled (rd, fs1, fs2) -> + fprintf oc " fle.d %a, %a, %a\n" ireg rd freg fs1 freg fs2 + + | Pfsqrtd (fd, fs) -> + fprintf oc " fsqrt.d %a, %a\n" freg fd freg fs + + | Pfmaddd (fd, fs1, fs2, fs3) -> + fprintf oc " fmadd.d %a, %a, %a, %a\n" freg fd freg fs1 freg fs2 freg fs3 + | Pfmsubd (fd, fs1, fs2, fs3) -> + fprintf oc " fmsub.d %a, %a, %a, %a\n" freg fd freg fs1 freg fs2 freg fs3 + | Pfnmaddd (fd, fs1, fs2, fs3) -> + fprintf oc " fnmadd.d %a, %a, %a, %a\n" freg fd freg fs1 freg fs2 freg fs3 + | Pfnmsubd (fd, fs1, fs2, fs3) -> + fprintf oc " fnmsub.d %a, %a, %a, %a\n" freg fd freg fs1 freg fs2 freg fs3 + + | Pfcvtwd (rd, fs) -> + fprintf oc " fcvt.w.d %a, %a, rtz\n" ireg rd freg fs + | Pfcvtwud (rd, fs) -> + fprintf oc " fcvt.wu.d %a, %a, rtz\n" ireg rd freg fs + | Pfcvtdw (fd, rs) -> + fprintf oc " fcvt.d.w %a, %a\n" freg fd ireg0 rs + | Pfcvtdwu (fd, rs) -> + fprintf oc " fcvt.d.wu %a, %a\n" freg fd ireg0 rs + + | Pfcvtld (rd, fs) -> assert Archi.ptr64; + fprintf oc " fcvt.l.d %a, %a, rtz\n" ireg rd freg fs + | Pfcvtlud (rd, fs) -> assert Archi.ptr64; + fprintf oc " fcvt.lu.d %a, %a, rtz\n" ireg rd freg fs + | Pfcvtdl (fd, rs) -> assert Archi.ptr64; + fprintf oc " fcvt.d.l %a, %a\n" freg fd ireg0 rs + | Pfcvtdlu (fd, rs) -> assert Archi.ptr64; + fprintf oc " fcvt.d.lu %a, %a\n" freg fd ireg0 rs + + | Pfcvtds (fd, fs) -> + fprintf oc " fcvt.d.s %a, %a\n" freg fd freg fs + | Pfcvtsd (fd, fs) -> + fprintf oc " fcvt.s.d %a, %a\n" freg fd freg fs + + (* Pseudo-instructions expanded in Asmexpand *) + | Pselectl(_, _, _, _) -> + assert false + | Pallocframe(sz, ofs) -> + assert false + | Pfreeframe(sz, ofs) -> + assert false + | Pseqw _ | Psnew _ | Pseql _ | Psnel _ | Pcvtl2w _ | Pcvtw2l _ -> + assert false + + (* Pseudo-instructions that remain *) + | Plabel lbl -> + fprintf oc "%a:\n" print_label lbl + | Ploadsymbol(rd, id, ofs) -> + loadsymbol oc rd id ofs + | Ploadsymbol_high(rd, id, ofs) -> + fprintf oc " lui %a, %%hi(%a)\n" ireg rd symbol_offset (id, ofs) + | Ploadli(rd, n) -> + let d = camlint64_of_coqint n in + let lbl = label_literal64 d in + fprintf oc " ld %a, %a %s %Lx\n" ireg rd label lbl comment d + | Ploadfi(rd, f) -> + let d = camlint64_of_coqint(Floats.Float.to_bits f) in + let lbl = label_literal64 d in + fprintf oc " fld %a, %a, x31 %s %.18g\n" + freg rd label lbl comment (camlfloat_of_coqfloat f) + | Ploadsi(rd, f) -> + let s = camlint_of_coqint(Floats.Float32.to_bits f) in + let lbl = label_literal32 s in + fprintf oc " flw %a, %a, x31 %s %.18g\n" + freg rd label lbl comment (camlfloat_of_coqfloat32 f) + | Pbtbl(r, tbl) -> + let lbl = new_label() in + fprintf oc "%s jumptable [ " comment; + List.iter (fun l -> fprintf oc "%a " print_label l) tbl; + fprintf oc "]\n"; + fprintf oc " sll x5, %a, 2\n" ireg r; + fprintf oc " la x31, %a\n" label lbl; + fprintf oc " add x5, x31, x5\n"; + fprintf oc " lw x5, 0(x5)\n"; + fprintf oc " add x5, x31, x5\n"; + fprintf oc " jr x5\n"; + jumptables := (lbl, tbl) :: !jumptables; + fprintf oc "%s end pseudoinstr btbl\n" comment + | Pnop -> + fprintf oc " nop\n" + | Pbuiltin(ef, args, res) -> + begin match ef with + | EF_annot(kind,txt, targs) -> + begin match (P.to_int kind) with + | 1 -> let annot = annot_text preg_annot "x2" (camlstring_of_coqstring txt) args in + fprintf oc "%s annotation: %S\n" comment annot + | 2 -> let lbl = new_label () in + fprintf oc "%a:\n" label lbl; + add_ais_annot lbl preg_annot "x2" (camlstring_of_coqstring txt) args + | _ -> assert false + end + | EF_debug(kind, txt, targs) -> + print_debug_info comment print_file_line preg_annot "sp" oc + (P.to_int kind) (extern_atom txt) args + | EF_inline_asm(txt, sg, clob) -> + fprintf oc "%s begin inline assembly\n\t" comment; + print_inline_asm preg_asm oc (camlstring_of_coqstring txt) sg args res; + fprintf oc "%s end inline assembly\n" comment + | _ -> + assert false + end + + let get_section_names name = + let (text, lit) = + match C2C.atom_sections name with + | t :: l :: _ -> (t, l) + | _ -> (Section_text, Section_literal) in + text,lit,Section_jumptable + + let print_align oc alignment = + fprintf oc " .balign %d\n" alignment + + let print_jumptable oc jmptbl = + let print_tbl oc (lbl, tbl) = + fprintf oc "%a:\n" label lbl; + List.iter + (fun l -> fprintf oc " .long %a - %a\n" + print_label l label lbl) + tbl in + if !jumptables <> [] then + begin + section oc jmptbl; + fprintf oc " .balign 4\n"; + List.iter (print_tbl oc) !jumptables; + jumptables := [] + end + + let print_fun_info = elf_print_fun_info + + let print_optional_fun_info _ = () + + let print_var_info = elf_print_var_info + + let print_comm_symb oc sz name align = + if C2C.atom_is_static name then + fprintf oc " .local %a\n" symbol name; + fprintf oc " .comm %a, %s, %d\n" + symbol name + (Z.to_string sz) + align + + let print_instructions oc fn = + current_function_sig := fn.fn_sig; + List.iter (print_instruction oc) fn.fn_code + + +(* Data *) + + let address = if Archi.ptr64 then ".quad" else ".long" + + let print_prologue oc = + fprintf oc " .option %s\n" (if Archi.pic_code() then "pic" else "nopic"); + if !Clflags.option_g then begin + section oc Section_text; + end + + let print_epilogue oc = + if !Clflags.option_g then begin + Debug.compute_gnu_file_enum (fun f -> ignore (print_file oc f)); + section oc Section_text; + end + + let default_falignment = 2 + + let cfi_startproc oc = () + let cfi_endproc oc = () + + end + +let sel_target () = + (module Target:TARGET) -- cgit From 7cc2810b4b1ea92a8f8a8739f69a94d5578e7b9d Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Mon, 29 Mar 2021 11:12:07 +0200 Subject: replacing omega with lia in some file --- riscV/RTLpathSE_simplify.v | 14 +++++++------- riscV/ValueAOp.v | 4 ++-- 2 files changed, 9 insertions(+), 9 deletions(-) (limited to 'riscV') diff --git a/riscV/RTLpathSE_simplify.v b/riscV/RTLpathSE_simplify.v index 6a0297e9..2669d4bc 100644 --- a/riscV/RTLpathSE_simplify.v +++ b/riscV/RTLpathSE_simplify.v @@ -524,7 +524,7 @@ Remark ltu_12_wordsize: Proof. unfold Int.iwordsize, Int.zwordsize. simpl. unfold Int.ltu. apply zlt_true. - rewrite !Int.unsigned_repr; try cbn; try omega. + rewrite !Int.unsigned_repr; try cbn; try lia. Qed. (* Signed longs *) @@ -575,14 +575,14 @@ Proof. intros v n EQMAX. unfold Val.cmp_bool; destruct v; simpl; auto. unfold Int.lt. replace (Int.signed (Int.add n Int.one)) with (Int.signed n + 1). destruct (zlt (Int.signed n) (Int.signed i)). - rewrite zlt_false by omega. auto. - rewrite zlt_true by omega. auto. + rewrite zlt_false by lia. auto. + rewrite zlt_true by lia. auto. rewrite Int.add_signed. symmetry; apply Int.signed_repr. specialize (Int.eq_spec n (Int.repr Int.max_signed)). rewrite EQMAX; simpl; intros. assert (Int.signed n <> Int.max_signed). { red; intros E. elim H. rewrite <- (Int.repr_signed n). rewrite E. auto. } - generalize (Int.signed_range n); omega. + generalize (Int.signed_range n); lia. Qed. Lemma cmpl_ltle_add_one: forall v n, @@ -593,14 +593,14 @@ Proof. intros v n EQMAX. unfold Val.cmpl_bool; destruct v; simpl; auto. unfold Int64.lt. replace (Int64.signed (Int64.add n Int64.one)) with (Int64.signed n + 1). destruct (zlt (Int64.signed n) (Int64.signed i)). - rewrite zlt_false by omega. auto. - rewrite zlt_true by omega. auto. + rewrite zlt_false by lia. auto. + rewrite zlt_true by lia. auto. rewrite Int64.add_signed. symmetry; apply Int64.signed_repr. specialize (Int64.eq_spec n (Int64.repr Int64.max_signed)). rewrite EQMAX; simpl; intros. assert (Int64.signed n <> Int64.max_signed). { red; intros E. elim H. rewrite <- (Int64.repr_signed n). rewrite E. auto. } - generalize (Int64.signed_range n); omega. + generalize (Int64.signed_range n); lia. Qed. Remark lt_maxsgn_false_int: forall i, diff --git a/riscV/ValueAOp.v b/riscV/ValueAOp.v index 97f3ff61..ca0834db 100644 --- a/riscV/ValueAOp.v +++ b/riscV/ValueAOp.v @@ -13,7 +13,7 @@ Require Import Coqlib Compopts. Require Import AST Integers Floats Values Memory Globalenvs. Require Import Op RTL ValueDomain. -Require Import Zbits. +Require Import Zbits Lia. (** Value analysis for RISC V operators *) @@ -390,7 +390,7 @@ Proof. assert (DEFAULT: vmatch bc (Val.maketotal (option_map Val.of_bool ob)) (Uns Pbot 1)). { destruct ob; simpl; auto with va. - destruct b; constructor; try omega. + destruct b; constructor; try lia. change 1 with (usize Int.one). apply is_uns_usize. red; intros. apply Int.bits_zero. } -- cgit From 320c55590cc30d4ef5b2c1a226f0f940a6bdb445 Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Sun, 25 Apr 2021 13:57:47 +0200 Subject: Support __builtin_unreachable Not yet used for optimizations. --- riscV/Asmexpand.ml | 4 ++++ 1 file changed, 4 insertions(+) (limited to 'riscV') diff --git a/riscV/Asmexpand.ml b/riscV/Asmexpand.ml index e8c142e9..dc0ec184 100644 --- a/riscV/Asmexpand.ml +++ b/riscV/Asmexpand.ml @@ -646,8 +646,12 @@ let expand_builtin_inline name args res = (fun rl -> emit (Pmulw (rl, X a, X b)); emit (Pmulhuw (rh, X a, X b))) + (* No operation *) | "__builtin_nop", [], _ -> emit Pnop + (* Optimization hint *) + | "__builtin_unreachable", [], _ -> + () (* Catch-all *) | _ -> raise (Error ("unrecognized builtin " ^ name)) -- cgit From 04f499c632a76e460560fc9ec4e14d8216e7fc18 Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Tue, 4 May 2021 10:46:17 +0200 Subject: Use the LGPL instead of the GPL for dual-licensed files The GPL makes sense for whole applications, but the dual-licensed Coq and OCaml files are more like libraries to be combined with other code, so the LGPL is more appropriate. --- riscV/Archi.v | 9 +++++---- riscV/Builtins1.v | 9 +++++---- riscV/CBuiltins.ml | 9 +++++---- riscV/extractionMachdep.v | 9 +++++---- 4 files changed, 20 insertions(+), 16 deletions(-) (limited to 'riscV') diff --git a/riscV/Archi.v b/riscV/Archi.v index 1b24e732..9e561ca8 100644 --- a/riscV/Archi.v +++ b/riscV/Archi.v @@ -7,10 +7,11 @@ (* *) (* Copyright Institut National de Recherche en Informatique et en *) (* Automatique. All rights reserved. This file is distributed *) -(* under the terms of the GNU General Public License as published by *) -(* the Free Software Foundation, either version 2 of the License, or *) -(* (at your option) any later version. This file is also distributed *) -(* under the terms of the INRIA Non-Commercial License Agreement. *) +(* under the terms of the GNU Lesser General Public License as *) +(* published by the Free Software Foundation, either version 2.1 of *) +(* the License, or (at your option) any later version. *) +(* This file is also distributed under the terms of the *) +(* INRIA Non-Commercial License Agreement. *) (* *) (* *********************************************************************) diff --git a/riscV/Builtins1.v b/riscV/Builtins1.v index 53c83d7e..cd6f8cc4 100644 --- a/riscV/Builtins1.v +++ b/riscV/Builtins1.v @@ -6,10 +6,11 @@ (* *) (* Copyright Institut National de Recherche en Informatique et en *) (* Automatique. All rights reserved. This file is distributed *) -(* under the terms of the GNU General Public License as published by *) -(* the Free Software Foundation, either version 2 of the License, or *) -(* (at your option) any later version. This file is also distributed *) -(* under the terms of the INRIA Non-Commercial License Agreement. *) +(* under the terms of the GNU Lesser General Public License as *) +(* published by the Free Software Foundation, either version 2.1 of *) +(* the License, or (at your option) any later version. *) +(* This file is also distributed under the terms of the *) +(* INRIA Non-Commercial License Agreement. *) (* *) (* *********************************************************************) diff --git a/riscV/CBuiltins.ml b/riscV/CBuiltins.ml index a2087cb7..9ff4e029 100644 --- a/riscV/CBuiltins.ml +++ b/riscV/CBuiltins.ml @@ -6,10 +6,11 @@ (* *) (* Copyright Institut National de Recherche en Informatique et en *) (* Automatique. All rights reserved. This file is distributed *) -(* under the terms of the GNU General Public License as published by *) -(* the Free Software Foundation, either version 2 of the License, or *) -(* (at your option) any later version. This file is also distributed *) -(* under the terms of the INRIA Non-Commercial License Agreement. *) +(* under the terms of the GNU Lesser General Public License as *) +(* published by the Free Software Foundation, either version 2.1 of *) +(* the License, or (at your option) any later version. *) +(* This file is also distributed under the terms of the *) +(* INRIA Non-Commercial License Agreement. *) (* *) (* *********************************************************************) diff --git a/riscV/extractionMachdep.v b/riscV/extractionMachdep.v index c9a1040a..890735ba 100644 --- a/riscV/extractionMachdep.v +++ b/riscV/extractionMachdep.v @@ -6,10 +6,11 @@ (* *) (* Copyright Institut National de Recherche en Informatique et en *) (* Automatique. All rights reserved. This file is distributed *) -(* under the terms of the GNU General Public License as published by *) -(* the Free Software Foundation, either version 2 of the License, or *) -(* (at your option) any later version. This file is also distributed *) -(* under the terms of the INRIA Non-Commercial License Agreement. *) +(* under the terms of the GNU Lesser General Public License as *) +(* published by the Free Software Foundation, either version 2.1 of *) +(* the License, or (at your option) any later version. *) +(* This file is also distributed under the terms of the *) +(* INRIA Non-Commercial License Agreement. *) (* *) (* *********************************************************************) -- cgit From 39710f78062a4a999c079b58181a58e62b78c30b Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Thu, 13 May 2021 17:26:05 +0200 Subject: Register X1 is destroyed by some built-in functions E.g. __builtin_bswap. Update Asm modeling of builtins accordingly. --- riscV/Asm.v | 2 +- riscV/Asmgenproof.v | 6 ++++-- 2 files changed, 5 insertions(+), 3 deletions(-) (limited to 'riscV') diff --git a/riscV/Asm.v b/riscV/Asm.v index 7e1b1fc8..a47573a2 100644 --- a/riscV/Asm.v +++ b/riscV/Asm.v @@ -1080,7 +1080,7 @@ Inductive step: state -> trace -> state -> Prop := rs' = nextinstr (set_res res vres (undef_regs (map preg_of (destroyed_by_builtin ef)) - (rs#X31 <- Vundef))) -> + (rs #X1 <- Vundef #X31 <- Vundef))) -> step (State rs m) t (State rs' m') | exec_step_external: forall b ef args res rs m t rs' m', diff --git a/riscV/Asmgenproof.v b/riscV/Asmgenproof.v index ab07d071..798dad9f 100644 --- a/riscV/Asmgenproof.v +++ b/riscV/Asmgenproof.v @@ -835,13 +835,15 @@ Local Transparent destroyed_by_op. econstructor; eauto. instantiate (2 := tf); instantiate (1 := x). unfold nextinstr. rewrite Pregmap.gss. - rewrite set_res_other. rewrite undef_regs_other_2. rewrite Pregmap.gso by congruence. + rewrite set_res_other. rewrite undef_regs_other_2. + rewrite ! Pregmap.gso by congruence. rewrite <- H1. simpl. econstructor; eauto. eapply code_tail_next_int; eauto. rewrite preg_notin_charact. intros. auto with asmgen. auto with asmgen. apply agree_nextinstr. eapply agree_set_res; auto. - eapply agree_undef_regs; eauto. intros. rewrite undef_regs_other_2; auto. apply Pregmap.gso; auto with asmgen. + eapply agree_undef_regs; eauto. intros. rewrite undef_regs_other_2; auto. + rewrite ! Pregmap.gso; auto with asmgen. congruence. - (* Mgoto *) -- cgit From 25a3de1e003f02012eefe427ea24c875972c5c22 Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Sat, 29 May 2021 23:09:08 +0200 Subject: just remove a debug print --- riscV/ExpansionOracle.ml | 1 - 1 file changed, 1 deletion(-) (limited to 'riscV') diff --git a/riscV/ExpansionOracle.ml b/riscV/ExpansionOracle.ml index 1384b9b3..b3f1f8ce 100644 --- a/riscV/ExpansionOracle.ml +++ b/riscV/ExpansionOracle.ml @@ -696,7 +696,6 @@ let expanse (sb : superblock) code pm = was_exp := false; let inst = get_some @@ PTree.get n code in (if !Clflags.option_fexpanse_rtlcond then - debug "We are checking node %d\n" (p2i n); match inst with (* Expansion of conditions - Ocmp *) | Iop (Ocmp (Ccomp c), a1 :: a2 :: nil, dest, succ) -> -- cgit From c44fc24eb6111c177d1d6fc973a366ebf646202b Mon Sep 17 00:00:00 2001 From: Léo Gourdin Date: Tue, 1 Jun 2021 11:18:59 +0200 Subject: removing some Expansion when loading float/single constants --- riscV/ExpansionOracle.ml | 38 ++++++++++++++++++++++---------------- 1 file changed, 22 insertions(+), 16 deletions(-) (limited to 'riscV') diff --git a/riscV/ExpansionOracle.ml b/riscV/ExpansionOracle.ml index b3f1f8ce..68d4e4d2 100644 --- a/riscV/ExpansionOracle.ml +++ b/riscV/ExpansionOracle.ml @@ -825,22 +825,28 @@ let expanse (sb : superblock) code pm = | _ -> ()); (if !Clflags.option_fexpanse_others && not !was_exp then match inst with - | Iop (Ofloatconst f, nil, dest, succ) -> - debug "Iop/Ofloatconst\n"; - let r = r2pi () in - let l = loadimm64 vn r (Floats.Float.to_bits f) in - let r', l' = extract_arg l in - exp := addinst vn Ofloat_of_bits [ r' ] dest :: l'; - exp := extract_final vn !exp dest succ; - was_exp := true - | Iop (Osingleconst f, nil, dest, succ) -> - debug "Iop/Osingleconst\n"; - let r = r2pi () in - let l = loadimm32 vn r (Floats.Float32.to_bits f) in - let r', l' = extract_arg l in - exp := addinst vn Osingle_of_bits [ r' ] dest :: l'; - exp := extract_final vn !exp dest succ; - was_exp := true + | Iop (Ofloatconst f, nil, dest, succ) -> ( + match make_immed64 (Floats.Float.to_bits f) with + | Imm64_single _ | Imm64_large _ -> () + | Imm64_pair (hi, lo) -> + debug "Iop/Ofloatconst\n"; + let r = r2pi () in + let l = load_hilo64 vn r hi lo in + let r', l' = extract_arg l in + exp := addinst vn Ofloat_of_bits [ r' ] dest :: l'; + exp := extract_final vn !exp dest succ; + was_exp := true) + | Iop (Osingleconst f, nil, dest, succ) -> ( + match make_immed32 (Floats.Float32.to_bits f) with + | Imm32_single imm -> () + | Imm32_pair (hi, lo) -> + debug "Iop/Osingleconst\n"; + let r = r2pi () in + let l = load_hilo32 vn r hi lo in + let r', l' = extract_arg l in + exp := addinst vn Osingle_of_bits [ r' ] dest :: l'; + exp := extract_final vn !exp dest succ; + was_exp := true) | Iop (Ointconst n, nil, dest, succ) -> debug "Iop/Ointconst\n"; exp := loadimm32 vn dest n; -- cgit From 5a632954c85e8b2b5afea124e4fc83f39c5d3598 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Tue, 1 Jun 2021 14:37:07 +0200 Subject: [BROKEN] Merge with v3.9 : something broken for __builtin_expect in cfrontend/C2C.ml --- riscV/Archi.v | 9 +++++---- riscV/Asmexpand.ml | 4 ++++ riscV/Builtins1.v | 9 +++++---- riscV/CBuiltins.ml | 9 +++++---- riscV/extractionMachdep.v | 9 +++++---- 5 files changed, 24 insertions(+), 16 deletions(-) (limited to 'riscV') diff --git a/riscV/Archi.v b/riscV/Archi.v index 1bb80e89..96f34276 100644 --- a/riscV/Archi.v +++ b/riscV/Archi.v @@ -7,10 +7,11 @@ (* *) (* Copyright Institut National de Recherche en Informatique et en *) (* Automatique. All rights reserved. This file is distributed *) -(* under the terms of the GNU General Public License as published by *) -(* the Free Software Foundation, either version 2 of the License, or *) -(* (at your option) any later version. This file is also distributed *) -(* under the terms of the INRIA Non-Commercial License Agreement. *) +(* under the terms of the GNU Lesser General Public License as *) +(* published by the Free Software Foundation, either version 2.1 of *) +(* the License, or (at your option) any later version. *) +(* This file is also distributed under the terms of the *) +(* INRIA Non-Commercial License Agreement. *) (* *) (* *********************************************************************) diff --git a/riscV/Asmexpand.ml b/riscV/Asmexpand.ml index a49efce8..50dc20be 100644 --- a/riscV/Asmexpand.ml +++ b/riscV/Asmexpand.ml @@ -646,8 +646,12 @@ let expand_builtin_inline name args res = (fun rl -> emit (Pmulw (rl, X a, X b)); emit (Pmulhuw (rh, X a, X b))) + (* No operation *) | "__builtin_nop", [], _ -> emit Pnop + (* Optimization hint *) + | "__builtin_unreachable", [], _ -> + () (* Catch-all *) | _ -> raise (Error ("unrecognized builtin " ^ name)) diff --git a/riscV/Builtins1.v b/riscV/Builtins1.v index 47bacffa..6691d15c 100644 --- a/riscV/Builtins1.v +++ b/riscV/Builtins1.v @@ -6,10 +6,11 @@ (* *) (* Copyright Institut National de Recherche en Informatique et en *) (* Automatique. All rights reserved. This file is distributed *) -(* under the terms of the GNU General Public License as published by *) -(* the Free Software Foundation, either version 2 of the License, or *) -(* (at your option) any later version. This file is also distributed *) -(* under the terms of the INRIA Non-Commercial License Agreement. *) +(* under the terms of the GNU Lesser General Public License as *) +(* published by the Free Software Foundation, either version 2.1 of *) +(* the License, or (at your option) any later version. *) +(* This file is also distributed under the terms of the *) +(* INRIA Non-Commercial License Agreement. *) (* *) (* *********************************************************************) diff --git a/riscV/CBuiltins.ml b/riscV/CBuiltins.ml index 00b44fd5..ca0dbc6d 100644 --- a/riscV/CBuiltins.ml +++ b/riscV/CBuiltins.ml @@ -6,10 +6,11 @@ (* *) (* Copyright Institut National de Recherche en Informatique et en *) (* Automatique. All rights reserved. This file is distributed *) -(* under the terms of the GNU General Public License as published by *) -(* the Free Software Foundation, either version 2 of the License, or *) -(* (at your option) any later version. This file is also distributed *) -(* under the terms of the INRIA Non-Commercial License Agreement. *) +(* under the terms of the GNU Lesser General Public License as *) +(* published by the Free Software Foundation, either version 2.1 of *) +(* the License, or (at your option) any later version. *) +(* This file is also distributed under the terms of the *) +(* INRIA Non-Commercial License Agreement. *) (* *) (* *********************************************************************) diff --git a/riscV/extractionMachdep.v b/riscV/extractionMachdep.v index c9a1040a..890735ba 100644 --- a/riscV/extractionMachdep.v +++ b/riscV/extractionMachdep.v @@ -6,10 +6,11 @@ (* *) (* Copyright Institut National de Recherche en Informatique et en *) (* Automatique. All rights reserved. This file is distributed *) -(* under the terms of the GNU General Public License as published by *) -(* the Free Software Foundation, either version 2 of the License, or *) -(* (at your option) any later version. This file is also distributed *) -(* under the terms of the INRIA Non-Commercial License Agreement. *) +(* under the terms of the GNU Lesser General Public License as *) +(* published by the Free Software Foundation, either version 2.1 of *) +(* the License, or (at your option) any later version. *) +(* This file is also distributed under the terms of the *) +(* INRIA Non-Commercial License Agreement. *) (* *) (* *********************************************************************) -- cgit From dbff5b8a016fe9f6667ea007be3de764a50b620a Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 8 Jun 2021 00:30:31 +0200 Subject: omega -> lia --- riscV/Asmgenproof1.v | 16 ++++++++-------- 1 file changed, 8 insertions(+), 8 deletions(-) (limited to 'riscV') diff --git a/riscV/Asmgenproof1.v b/riscV/Asmgenproof1.v index 6abde89f..42ab8375 100644 --- a/riscV/Asmgenproof1.v +++ b/riscV/Asmgenproof1.v @@ -821,18 +821,18 @@ Proof. unfold Val.cmp; destruct (rs#r1); simpl; auto. rewrite B1. unfold Int.lt. rewrite zlt_false. auto. change (Int.signed (Int.repr Int.max_signed)) with Int.max_signed. - generalize (Int.signed_range i); omega. + generalize (Int.signed_range i); lia. * exploit (opimm32_correct Psltw Psltiw (Val.cmp Clt)); eauto. intros (rs1 & A1 & B1 & C1). exists rs1; split. eexact A1. split; auto. rewrite B1. unfold Val.cmp; simpl; destruct (rs#r1); simpl; auto. unfold Int.lt. replace (Int.signed (Int.add n Int.one)) with (Int.signed n + 1). destruct (zlt (Int.signed n) (Int.signed i)). - rewrite zlt_false by omega. auto. - rewrite zlt_true by omega. auto. + rewrite zlt_false by lia. auto. + rewrite zlt_true by lia. auto. rewrite Int.add_signed. symmetry; apply Int.signed_repr. assert (Int.signed n <> Int.max_signed). { red; intros E. elim H1. rewrite <- (Int.repr_signed n). rewrite E. auto. } - generalize (Int.signed_range n); omega. + generalize (Int.signed_range n); lia. + apply DFL. + apply DFL. Qed. @@ -919,18 +919,18 @@ Proof. unfold Val.cmpl; destruct (rs#r1); simpl; auto. rewrite B1. unfold Int64.lt. rewrite zlt_false. auto. change (Int64.signed (Int64.repr Int64.max_signed)) with Int64.max_signed. - generalize (Int64.signed_range i); omega. + generalize (Int64.signed_range i); lia. * exploit (opimm64_correct Psltl Psltil (fun v1 v2 => Val.maketotal (Val.cmpl Clt v1 v2))); eauto. intros (rs1 & A1 & B1 & C1). exists rs1; split. eexact A1. split; auto. rewrite B1. unfold Val.cmpl; simpl; destruct (rs#r1); simpl; auto. unfold Int64.lt. replace (Int64.signed (Int64.add n Int64.one)) with (Int64.signed n + 1). destruct (zlt (Int64.signed n) (Int64.signed i)). - rewrite zlt_false by omega. auto. - rewrite zlt_true by omega. auto. + rewrite zlt_false by lia. auto. + rewrite zlt_true by lia. auto. rewrite Int64.add_signed. symmetry; apply Int64.signed_repr. assert (Int64.signed n <> Int64.max_signed). { red; intros E. elim H1. rewrite <- (Int64.repr_signed n). rewrite E. auto. } - generalize (Int64.signed_range n); omega. + generalize (Int64.signed_range n); lia. + apply DFL. + apply DFL. Qed. -- cgit