From 23fa2a18e015b9d330ad6f1f08cf50adf90bd80b Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Thu, 21 Mar 2019 22:39:27 +0100 Subject: try to be portable across archs --- powerpc/Machregsaux.ml | 5 +++++ powerpc/Machregsaux.mli | 2 ++ 2 files changed, 7 insertions(+) (limited to 'powerpc') diff --git a/powerpc/Machregsaux.ml b/powerpc/Machregsaux.ml index 664f71a0..0b0d4548 100644 --- a/powerpc/Machregsaux.ml +++ b/powerpc/Machregsaux.ml @@ -33,3 +33,8 @@ let register_by_name s = let can_reserve_register r = List.mem r Conventions1.int_callee_save_regs || List.mem r Conventions1.float_callee_save_regs + +let class_of_type = function + | AST.Tint | AST.Tlong -> 0 + | AST.Tfloat | AST.Tsingle -> 1 + | AST.Tany32 | AST.Tany64 -> assert false diff --git a/powerpc/Machregsaux.mli b/powerpc/Machregsaux.mli index 9404568d..d7117c21 100644 --- a/powerpc/Machregsaux.mli +++ b/powerpc/Machregsaux.mli @@ -16,3 +16,5 @@ val name_of_register: Machregs.mreg -> string option val register_by_name: string -> Machregs.mreg option val is_scratch_register: string -> bool val can_reserve_register: Machregs.mreg -> bool + +val class_of_type: AST.typ -> int -- cgit From 88c46de03f37dd9edb78e68734a9976ab2ccc056 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 22 Mar 2019 10:28:54 +0100 Subject: seems like powerpc runs but the result segfaults --- powerpc/SelectLong.vp | 2 +- powerpc/SelectLongproof.v | 1 + powerpc/SelectOp.vp | 8 ++++++++ powerpc/SelectOpproof.v | 28 ++++++++++++++++++++++++++-- 4 files changed, 36 insertions(+), 3 deletions(-) (limited to 'powerpc') diff --git a/powerpc/SelectLong.vp b/powerpc/SelectLong.vp index 5f13774b..e4274ba5 100644 --- a/powerpc/SelectLong.vp +++ b/powerpc/SelectLong.vp @@ -16,7 +16,7 @@ Require Import Coqlib. Require Import Compopts. Require Import AST Integers Floats. Require Import Op CminorSel. -Require Import SelectOp SplitLong. +Require Import OpHelpers SelectOp SplitLong. Local Open Scope cminorsel_scope. Local Open Scope string_scope. diff --git a/powerpc/SelectLongproof.v b/powerpc/SelectLongproof.v index a214d131..b4e48596 100644 --- a/powerpc/SelectLongproof.v +++ b/powerpc/SelectLongproof.v @@ -16,6 +16,7 @@ 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. diff --git a/powerpc/SelectOp.vp b/powerpc/SelectOp.vp index d2ca408a..478ce251 100644 --- a/powerpc/SelectOp.vp +++ b/powerpc/SelectOp.vp @@ -43,6 +43,7 @@ Require Import Integers. Require Import Floats. Require Import Op. Require Import CminorSel. +Require Import OpHelpers. Local Open Scope cminorsel_scope. @@ -552,3 +553,10 @@ Nondetfunction builtin_arg (e: expr) := | Eop Oadd (e1:::e2:::Enil) => BA_addptr (BA e1) (BA e2) | _ => BA e end. + +(* floats *) +Definition divf_base (e1: expr) (e2: expr) := + Eop Odivf (e1 ::: e2 ::: Enil). + +Definition divfs_base (e1: expr) (e2: expr) := + Eop Odivfs (e1 ::: e2 ::: Enil). diff --git a/powerpc/SelectOpproof.v b/powerpc/SelectOpproof.v index 5f87d9b9..00b91e70 100644 --- a/powerpc/SelectOpproof.v +++ b/powerpc/SelectOpproof.v @@ -25,6 +25,8 @@ Require Import Cminor. Require Import Op. Require Import CminorSel. Require Import SelectOp. +Require Import OpHelpers. +Require Import OpHelpersproof. Local Open Scope cminorsel_scope. Local Transparent Archi.ptr64. @@ -77,8 +79,10 @@ Ltac TrivialExists := (** * Correctness of the smart constructors *) Section CMCONSTR. - -Variable ge: genv. +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. @@ -1044,4 +1048,24 @@ Proof. - 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. End CMCONSTR. -- cgit From 4c379d48b35e7c8156f3953fede31d5e47faf8ca Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 19 Jul 2019 18:59:44 +0200 Subject: helpers broke compilation --- powerpc/SelectOp.vp | 15 ++------------- powerpc/SelectOpproof.v | 3 +-- 2 files changed, 3 insertions(+), 15 deletions(-) (limited to 'powerpc') diff --git a/powerpc/SelectOp.vp b/powerpc/SelectOp.vp index 07c325c1..50b1bdd6 100644 --- a/powerpc/SelectOp.vp +++ b/powerpc/SelectOp.vp @@ -38,17 +38,8 @@ Require Import Coqlib. Require Import Compopts. -<<<<<<< HEAD -Require Import AST. -Require Import Integers. -Require Import Floats. -Require Import Op. -Require Import CminorSel. -Require Import OpHelpers. -======= Require Import AST Integers Floats Builtins. -Require Import Op CminorSel. ->>>>>>> 91381b65f5aa76e5195caae9ef331b3f5f95afaf +Require Import Op OpHelpers CminorSel. Require Archi. Local Open Scope cminorsel_scope. @@ -573,16 +564,14 @@ Nondetfunction builtin_arg (e: expr) := | _ => BA e end. -<<<<<<< HEAD (* floats *) Definition divf_base (e1: expr) (e2: expr) := Eop Odivf (e1 ::: e2 ::: Enil). Definition divfs_base (e1: expr) (e2: expr) := Eop Odivfs (e1 ::: e2 ::: Enil). -======= + (** Platform-specific known builtins *) Definition platform_builtin (b: platform_builtin) (args: exprlist) : option expr := None. ->>>>>>> 91381b65f5aa76e5195caae9ef331b3f5f95afaf diff --git a/powerpc/SelectOpproof.v b/powerpc/SelectOpproof.v index 3e34244d..63066135 100644 --- a/powerpc/SelectOpproof.v +++ b/powerpc/SelectOpproof.v @@ -1062,7 +1062,6 @@ Proof. - constructor; auto. Qed. -<<<<<<< HEAD (* floating-point division without HELPERS *) Theorem eval_divf_base: forall le a b x y, @@ -1083,7 +1082,7 @@ Proof. intros; unfold divfs_base. TrivialExists. Qed. -======= + (** Platform-specific known builtins *) Theorem eval_platform_builtin: -- cgit From 780ad9d001af651a49d7470e963ed9a49ee11a4c Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 19 Jul 2019 19:49:46 +0200 Subject: various fixes --- powerpc/SelectOpproof.v | 1 - 1 file changed, 1 deletion(-) (limited to 'powerpc') diff --git a/powerpc/SelectOpproof.v b/powerpc/SelectOpproof.v index 63066135..b0cd70a4 100644 --- a/powerpc/SelectOpproof.v +++ b/powerpc/SelectOpproof.v @@ -1095,5 +1095,4 @@ Proof. intros. discriminate. Qed. ->>>>>>> 91381b65f5aa76e5195caae9ef331b3f5f95afaf End CMCONSTR. -- cgit From d84a003dc41c1ce572e86f399f5a610a78eda15f Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sat, 7 Sep 2019 13:48:11 +0200 Subject: PowerPC compiles --- powerpc/Asmexpand.ml | 31 +++++++++++++++---------------- powerpc/Asmgen.v | 14 ++++++++++---- powerpc/Asmgenproof.v | 8 ++++++++ powerpc/Asmgenproof1.v | 5 +++-- powerpc/Op.v | 42 ++++++++++++++++++++++++++++++++++++++++++ 5 files changed, 78 insertions(+), 22 deletions(-) (limited to 'powerpc') diff --git a/powerpc/Asmexpand.ml b/powerpc/Asmexpand.ml index 5ca4c611..0ef0a21e 100644 --- a/powerpc/Asmexpand.ml +++ b/powerpc/Asmexpand.ml @@ -14,7 +14,6 @@ of the PPC assembly code. *) open Camlcoq -open Integers open AST open Asm open Asmexpandaux @@ -89,7 +88,7 @@ let expand_annot_val kind txt targ args res = Note that lfd and stfd cannot trap on ill-formed floats. *) let offset_in_range ofs = - Int.eq (Asmgen.high_s ofs) _0 + Integers.Int.eq (Asmgen.high_s ofs) _0 let memcpy_small_arg sz arg tmp = match arg with @@ -97,7 +96,7 @@ let memcpy_small_arg sz arg tmp = (r, _0) | BA_addrstack ofs -> if offset_in_range ofs - && offset_in_range (Int.add ofs (Int.repr (Z.of_uint sz))) + && offset_in_range (Integers.Int.add ofs (Integers.Int.repr (Z.of_uint sz))) then (GPR1, ofs) else begin emit_addimm tmp GPR1 ofs; (tmp, _0) end | _ -> @@ -112,19 +111,19 @@ let expand_builtin_memcpy_small sz al src dst = if sz >= 8 && al >= 4 && !Clflags.option_ffpu then begin emit (Plfd(FPR13, Cint osrc, rsrc)); emit (Pstfd(FPR13, Cint odst, rdst)); - copy (Int.add osrc _8) (Int.add odst _8) (sz - 8) + copy (Integers.Int.add osrc _8) (Integers.Int.add odst _8) (sz - 8) end else if sz >= 4 then begin emit (Plwz(GPR0, Cint osrc, rsrc)); emit (Pstw(GPR0, Cint odst, rdst)); - copy (Int.add osrc _4) (Int.add odst _4) (sz - 4) + copy (Integers.Int.add osrc _4) (Integers.Int.add odst _4) (sz - 4) end else if sz >= 2 then begin emit (Plhz(GPR0, Cint osrc, rsrc)); emit (Psth(GPR0, Cint odst, rdst)); - copy (Int.add osrc _2) (Int.add odst _2) (sz - 2) + copy (Integers.Int.add osrc _2) (Integers.Int.add odst _2) (sz - 2) end else if sz >= 1 then begin emit (Plbz(GPR0, Cint osrc, rsrc)); emit (Pstb(GPR0, Cint odst, rdst)); - copy (Int.add osrc _1) (Int.add odst _1) (sz - 1) + copy (Integers.Int.add osrc _1) (Integers.Int.add odst _1) (sz - 1) end in copy osrc odst sz @@ -134,7 +133,7 @@ let memcpy_big_arg arg tmp = | BA (IR r) -> emit (Paddi(tmp, r, Cint _m4)) | BA_addrstack ofs -> - emit_addimm tmp GPR1 (Int.add ofs _m4) + emit_addimm tmp GPR1 (Integers.Int.add ofs _m4) | _ -> assert false @@ -227,10 +226,10 @@ let expand_volatile_access let offset_constant cst delta = match cst with | Cint n -> - let n' = Int.add n delta in + let n' = Integers.Int.add n delta in if offset_in_range n' then Some (Cint n') else None | Csymbol_sda(id, ofs) -> - Some (Csymbol_sda(id, Int.add ofs delta)) + Some (Csymbol_sda(id, Integers.Int.add ofs delta)) | _ -> None let expand_load_int64 hi lo base ofs_hi ofs_lo = @@ -438,7 +437,7 @@ let expand_integer_cond_move a1 a2 a3 res = if a2 = a3 then emit (Pmr (res, a2)) else if eref then begin - emit (Pcmpwi (a1,Cint (Int.zero))); + emit (Pcmpwi (a1,Cint (Integers.Int.zero))); emit (Pisel (res,a3,a2,CRbit_2)) end else begin (* a1 has type _Bool, hence it is 0 or 1 *) @@ -683,23 +682,23 @@ let expand_builtin_inline name args res = | "__builtin_icbi", [BA(IR a1)],_ -> emit (Picbi(GPR0,a1)) | "__builtin_dcbtls", [BA (IR a1); BA_int loc],_ -> - if not ((Int.eq loc _0) || (Int.eq loc _2)) then + if not ((Integers.Int.eq loc _0) || (Integers.Int.eq loc _2)) then raise (Error "the second argument of __builtin_dcbtls must be 0 or 2"); emit (Pdcbtls (loc,GPR0,a1)) | "__builtin_dcbtls",_,_ -> raise (Error "the second argument of __builtin_dcbtls must be a constant") | "__builtin_icbtls", [BA (IR a1); BA_int loc],_ -> - if not ((Int.eq loc _0) || (Int.eq loc _2)) then + if not ((Integers.Int.eq loc _0) || (Integers.Int.eq loc _2)) then raise (Error "the second argument of __builtin_icbtls must be 0 or 2"); emit (Picbtls (loc,GPR0,a1)) | "__builtin_icbtls",_,_ -> raise (Error "the second argument of __builtin_icbtls must be a constant") | "__builtin_prefetch" , [BA (IR a1) ;BA_int rw; BA_int loc],_ -> - if not (Int.ltu loc _4) then + if not (Integers.Int.ltu loc _4) then raise (Error "the last argument of __builtin_prefetch must be 0, 1 or 2"); - if Int.eq rw _0 then begin + if Integers.Int.eq rw _0 then begin emit (Pdcbt (loc,GPR0,a1)); - end else if Int.eq rw _1 then begin + end else if Integers.Int.eq rw _1 then begin emit (Pdcbtst (loc,GPR0,a1)); end else raise (Error "the second argument of __builtin_prefetch must be 0 or 1") diff --git a/powerpc/Asmgen.v b/powerpc/Asmgen.v index a686414a..29e2c028 100644 --- a/powerpc/Asmgen.v +++ b/powerpc/Asmgen.v @@ -783,8 +783,13 @@ Definition transl_memory_access Error(msg "Asmgen.transl_memory_access") end. -Definition transl_load (chunk: memory_chunk) (addr: addressing) - (args: list mreg) (dst: mreg) (k: code) := +Definition transl_load + (trap : trapping_mode) + (chunk: memory_chunk) (addr: addressing) + (args: list mreg) (dst: mreg) (k: code) := + match trap with + | NOTRAP => Error (msg "Asmgen.transl_load non-trapping loads unsupported on PPC") + | TRAP => match chunk with | Mint8signed => do r <- ireg_of dst; @@ -812,6 +817,7 @@ Definition transl_load (chunk: memory_chunk) (addr: addressing) transl_memory_access (Plfd r) (Plfdx r) addr args GPR12 k | _ => Error (msg "Asmgen.transl_load") + end end. Definition transl_store (chunk: memory_chunk) (addr: addressing) @@ -869,8 +875,8 @@ Definition transl_instr (f: Mach.function) (i: Mach.instruction) loadind GPR1 f.(fn_link_ofs) Tint R11 k1) | Mop op args res => transl_op op args res k - | Mload chunk addr args dst => - transl_load chunk addr args dst k + | Mload trap chunk addr args dst => + transl_load trap chunk addr args dst k | Mstore chunk addr args src => transl_store chunk addr args src k | Mcall sig (inl r) => diff --git a/powerpc/Asmgenproof.v b/powerpc/Asmgenproof.v index d653633c..21d5ce48 100644 --- a/powerpc/Asmgenproof.v +++ b/powerpc/Asmgenproof.v @@ -328,6 +328,7 @@ Proof. eapply loadind_label; eauto. eapply tail_nolabel_trans; eapply loadind_label; eauto. eapply transl_op_label; eauto. + destruct t; try discriminate. destruct m; monadInv H; (eapply tail_nolabel_trans; [eapply transl_memory_access_label; TailNoLabel|TailNoLabel]). destruct m; monadInv H; eapply transl_memory_access_label; TailNoLabel. destruct s0; monadInv H; TailNoLabel. @@ -657,6 +658,13 @@ Opaque loadind. split. simpl; congruence. apply R; auto with asmgen. + +- (* Mload notrap *) (* isn't there a nicer way? *) + inv AT. simpl in *. unfold bind in *. destruct (transl_code _ _ _) in *; discriminate. + +- (* Mload notrap *) + inv AT. simpl in *. unfold bind in *. destruct (transl_code _ _ _) in *; discriminate. + - (* Mstore *) assert (eval_addressing tge sp addr rs##args = Some a). rewrite <- H. apply eval_addressing_preserved. exact symbols_preserved. diff --git a/powerpc/Asmgenproof1.v b/powerpc/Asmgenproof1.v index 884d5366..6ceb7f85 100644 --- a/powerpc/Asmgenproof1.v +++ b/powerpc/Asmgenproof1.v @@ -1677,8 +1677,8 @@ Qed. (** Translation of loads *) Lemma transl_load_correct: - forall chunk addr args dst k c (rs: regset) m a v, - transl_load chunk addr args dst k = OK c -> + 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#GPR1) addr (map rs (map preg_of args)) = Some a -> Mem.loadv chunk m a = Some v -> exists rs', @@ -1687,6 +1687,7 @@ Lemma transl_load_correct: /\ forall r, r <> PC -> r <> GPR12 -> r <> GPR0 -> r <> preg_of dst -> rs' r = rs r. Proof. intros. + destruct trap; try discriminate. assert (LD: forall v, Val.lessdef a v -> v = a). { intros. inv H2; auto. discriminate H1. } assert (BASE: forall mk1 mk2 k' chunk' v', diff --git a/powerpc/Op.v b/powerpc/Op.v index 0f082c1f..cbd0291b 100644 --- a/powerpc/Op.v +++ b/powerpc/Op.v @@ -1032,6 +1032,21 @@ Proof. apply Val.add_inject; auto. apply H; simpl; auto. Qed. + +Lemma eval_addressing_inj_none: + forall addr sp1 vl1 sp2 vl2, + (forall id ofs, + In id (globals_addressing addr) -> + Val.inject f (Genv.symbol_address ge1 id ofs) (Genv.symbol_address ge2 id ofs)) -> + Val.inject f sp1 sp2 -> + Val.inject_list f vl1 vl2 -> + eval_addressing ge1 sp1 addr vl1 = None -> + eval_addressing ge2 sp2 addr vl2 = None. +Proof. + intros until vl2. intros Hglobal Hinjsp Hinjvl. + destruct addr; simpl in *; + inv Hinjvl; trivial; try discriminate; inv H0; trivial; try discriminate; inv H2; trivial; try discriminate. +Qed. End EVAL_COMPAT. (** Compatibility of the evaluation functions with the ``is less defined'' relation over values. *) @@ -1098,6 +1113,20 @@ Proof. rewrite <- val_inject_list_lessdef. eauto. auto. Qed. + +Lemma eval_addressing_lessdef_none: + forall sp addr vl1 vl2, + Val.lessdef_list vl1 vl2 -> + eval_addressing genv sp addr vl1 = None -> + eval_addressing genv sp addr vl2 = None. +Proof. + intros until vl2. intros Hlessdef Heval1. + destruct addr; simpl in *; + inv Hlessdef; trivial; try discriminate; + inv H0; trivial; try discriminate; + inv H2; trivial; try discriminate. +Qed. + Lemma eval_operation_lessdef: forall sp op vl1 vl2 v1 m1 m2, Val.lessdef_list vl1 vl2 -> @@ -1189,6 +1218,19 @@ Proof. econstructor; eauto. rewrite Ptrofs.add_zero_l; auto. Qed. +Lemma eval_addressing_inject_none: + forall addr vl1 vl2, + Val.inject_list f vl1 vl2 -> + eval_addressing genv (Vptr sp1 Ptrofs.zero) addr vl1 = None -> + eval_addressing genv (Vptr sp2 Ptrofs.zero) (shift_stack_addressing delta addr) vl2 = None. +Proof. + intros. + rewrite eval_shift_stack_addressing. + eapply eval_addressing_inj_none with (sp1 := Vptr sp1 Ptrofs.zero); eauto. + intros. apply symbol_address_inject. + econstructor; eauto. rewrite Ptrofs.add_zero_l; auto. +Qed. + Lemma eval_operation_inject: forall op vl1 vl2 v1 m1 m2, Val.inject_list f vl1 vl2 -> -- cgit From d315994f2d3dbec0bf66a430284eab00dd1d4a18 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 24 Sep 2019 16:51:11 +0200 Subject: trapping ops --- powerpc/Op.v | 24 ++++++++++++++++++++++++ 1 file changed, 24 insertions(+) (limited to 'powerpc') diff --git a/powerpc/Op.v b/powerpc/Op.v index cbd0291b..b73cb14b 100644 --- a/powerpc/Op.v +++ b/powerpc/Op.v @@ -581,6 +581,30 @@ Proof with (try exact I; try reflexivity). unfold Val.select. destruct (eval_condition c vl m). apply Val.normalize_type. exact I. Qed. +Definition is_trapping_op (op : operation) := + match op with + | Odiv | Odivl | Odivu | Odivlu + | Oshrximm _ | Oshrxlimm _ + | Ointoffloat | Ointuoffloat + | Ofloatofint | Ofloatofintu + | Olongoffloat + | Ofloatoflong => true + | _ => false + end. + +Lemma is_trapping_op_sound: + forall op vl sp m, + op <> Omove -> + is_trapping_op op = false -> + (List.length vl) = (List.length (fst (type_of_operation op))) -> + eval_operation genv sp op vl m <> None. +Proof. + destruct op; intros; simpl in *; try congruence. + all: try (destruct vl as [ | vh1 vl1]; try discriminate). + all: try (destruct vl1 as [ | vh2 vl2]; try discriminate). + all: try (destruct vl2 as [ | vh3 vl3]; try discriminate). + all: try (destruct vl3 as [ | vh4 vl4]; try discriminate). +Qed. End SOUNDNESS. (** * Manipulating and transforming operations *) -- cgit From 1b6667cf268189104bc3320e83fa23fe0d053717 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 7 Feb 2020 14:29:32 +0100 Subject: stubs to keep compiling on architectures not K1c --- powerpc/DuplicateOpcodeHeuristic.ml | 3 +++ 1 file changed, 3 insertions(+) create mode 100644 powerpc/DuplicateOpcodeHeuristic.ml (limited to 'powerpc') diff --git a/powerpc/DuplicateOpcodeHeuristic.ml b/powerpc/DuplicateOpcodeHeuristic.ml new file mode 100644 index 00000000..85505245 --- /dev/null +++ b/powerpc/DuplicateOpcodeHeuristic.ml @@ -0,0 +1,3 @@ +exception HeuristicSucceeded + +let opcode_heuristic code cond ifso ifnot preferred = () -- cgit