From 3fde34d48925db4153c5c288fa37da35725502ce Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 4 Feb 2020 09:14:25 +0100 Subject: stuff information into a record --- backend/CSE2.v | 59 ++++++++++++++++++++++++++++++----------------------- backend/CSE2proof.v | 50 +++++++++++++++++++++++++++------------------ 2 files changed, 63 insertions(+), 46 deletions(-) diff --git a/backend/CSE2.v b/backend/CSE2.v index b7665097..0479dad9 100644 --- a/backend/CSE2.v +++ b/backend/CSE2.v @@ -29,13 +29,18 @@ Proof. decide equality. Defined. +Record relmap := mkrel { + var_to_sym : (PTree.t sym_val) + }. + Module RELATION. - -Definition t := (PTree.t sym_val). + +Definition t := relmap. + Definition eq (r1 r2 : t) := - forall x, (PTree.get x r1) = (PTree.get x r2). + forall x, (PTree.get x (var_to_sym r1)) = (PTree.get x (var_to_sym r2)). -Definition top : t := PTree.empty sym_val. +Definition top : t := mkrel (PTree.empty sym_val). Lemma eq_refl: forall x, eq x x. Proof. @@ -58,27 +63,27 @@ Qed. Definition sym_val_beq (x y : sym_val) := if eq_sym_val x y then true else false. -Definition beq (r1 r2 : t) := PTree.beq sym_val_beq r1 r2. +Definition beq (r1 r2 : t) := PTree.beq sym_val_beq (var_to_sym r1) (var_to_sym r2). Lemma beq_correct: forall r1 r2, beq r1 r2 = true -> eq r1 r2. Proof. unfold beq, eq. intros r1 r2 EQ x. - pose proof (PTree.beq_correct sym_val_beq r1 r2) as CORRECT. + pose proof (PTree.beq_correct sym_val_beq (var_to_sym r1) (var_to_sym r2)) as CORRECT. destruct CORRECT as [CORRECTF CORRECTB]. pose proof (CORRECTF EQ x) as EQx. clear CORRECTF CORRECTB EQ. unfold sym_val_beq in *. - destruct (r1 ! x) as [R1x | ] in *; - destruct (r2 ! x) as [R2x | ] in *; + destruct ((var_to_sym r1) ! x) as [R1x | ] in *; + destruct ((var_to_sym r2) ! x) as [R2x | ] in *; trivial; try contradiction. destruct (eq_sym_val R1x R2x) in *; congruence. Qed. Definition ge (r1 r2 : t) := forall x, - match PTree.get x r1 with + match PTree.get x (var_to_sym r1) with | None => True - | Some v => (PTree.get x r2) = Some v + | Some v => (PTree.get x (var_to_sym r2)) = Some v end. Lemma ge_refl: forall r1 r2, eq r1 r2 -> ge r1 r2. @@ -87,7 +92,7 @@ Proof. intros r1 r2 EQ x. pose proof (EQ x) as EQx. clear EQ. - destruct (r1 ! x). + destruct ((var_to_sym r1) ! x). - congruence. - trivial. Qed. @@ -98,12 +103,13 @@ Proof. intros r1 r2 r3 GE12 GE23 x. pose proof (GE12 x) as GE12x; clear GE12. pose proof (GE23 x) as GE23x; clear GE23. - destruct (r1 ! x); trivial. - destruct (r2 ! x); congruence. + destruct ((var_to_sym r1) ! x); trivial. + destruct ((var_to_sym r2) ! x); congruence. Qed. Definition lub (r1 r2 : t) := - PTree.combine + mkrel + (PTree.combine (fun ov1 ov2 => match ov1, ov2 with | (Some v1), (Some v2) => @@ -113,12 +119,12 @@ Definition lub (r1 r2 : t) := | None, _ | _, None => None end) - r1 r2. + (var_to_sym r1) (var_to_sym r2)). Lemma ge_lub_left: forall x y, ge (lub x y) x. Proof. unfold ge, lub. - intros r1 r2 x. + intros r1 r2 x. simpl. rewrite PTree.gcombine by reflexivity. destruct (_ ! _); trivial. destruct (_ ! _); trivial. @@ -128,7 +134,7 @@ Qed. Lemma ge_lub_right: forall x y, ge (lub x y) y. Proof. unfold ge, lub. - intros r1 r2 x. + intros r1 r2 x. simpl. rewrite PTree.gcombine by reflexivity. destruct (_ ! _); trivial. destruct (_ ! _); trivial. @@ -263,8 +269,8 @@ Definition kill_sym_val (dst : reg) (sv : sym_val) := end. Definition kill_reg (dst : reg) (rel : RELATION.t) := - PTree.filter1 (fun x => negb (kill_sym_val dst x)) - (PTree.remove dst rel). + mkrel (PTree.filter1 (fun x => negb (kill_sym_val dst x)) + (PTree.remove dst (var_to_sym rel))). Definition kill_sym_val_mem (sv: sym_val) := match sv with @@ -274,17 +280,18 @@ Definition kill_sym_val_mem (sv: sym_val) := end. Definition kill_mem (rel : RELATION.t) := - PTree.filter1 (fun x => negb (kill_sym_val_mem x)) rel. + mkrel + (PTree.filter1 (fun x => negb (kill_sym_val_mem x)) (var_to_sym rel)). Definition forward_move (rel : RELATION.t) (x : reg) : reg := - match rel ! x with + match (var_to_sym rel) ! x with | Some (SMove org) => org | _ => x end. Definition move (src dst : reg) (rel : RELATION.t) := - PTree.set dst (SMove (forward_move rel src)) (kill_reg dst rel). + mkrel (PTree.set dst (SMove (forward_move rel src)) (var_to_sym (kill_reg dst rel))). Definition find_op_fold op args (already : option reg) x sv := match already with @@ -300,7 +307,7 @@ Definition find_op_fold op args (already : option reg) x sv := end. Definition find_op (rel : RELATION.t) (op : operation) (args : list reg) := - PTree.fold (find_op_fold op args) rel None. + PTree.fold (find_op_fold op args) (var_to_sym rel) None. Definition find_load_fold chunk addr args (already : option reg) x sv := match already with @@ -318,12 +325,12 @@ Definition find_load_fold chunk addr args (already : option reg) x sv := end. Definition find_load (rel : RELATION.t) (chunk : memory_chunk) (addr : addressing) (args : list reg) := - PTree.fold (find_load_fold chunk addr args) rel None. + PTree.fold (find_load_fold chunk addr args) (var_to_sym rel) None. Definition oper2 (op: operation) (dst : reg) (args : list reg) (rel : RELATION.t) := let rel' := kill_reg dst rel in - PTree.set dst (SOp op (List.map (forward_move rel') args)) rel'. + mkrel (PTree.set dst (SOp op (List.map (forward_move rel') args)) (var_to_sym rel')). Definition oper1 (op: operation) (dst : reg) (args : list reg) (rel : RELATION.t) := @@ -348,7 +355,7 @@ Definition gen_oper (op: operation) (dst : reg) (args : list reg) Definition load2 (chunk: memory_chunk) (addr : addressing) (dst : reg) (args : list reg) (rel : RELATION.t) := let rel' := kill_reg dst rel in - PTree.set dst (SLoad chunk addr (List.map (forward_move rel') args)) rel'. + mkrel (PTree.set dst (SLoad chunk addr (List.map (forward_move rel') args)) (var_to_sym rel')). Definition load1 (chunk: memory_chunk) (addr : addressing) (dst : reg) (args : list reg) (rel : RELATION.t) := diff --git a/backend/CSE2proof.v b/backend/CSE2proof.v index 82fa8a28..a29bf202 100644 --- a/backend/CSE2proof.v +++ b/backend/CSE2proof.v @@ -58,7 +58,7 @@ Definition sem_sym_val sym rs (v : option val) : Prop := end. Definition sem_reg (rel : RELATION.t) (x : reg) (rs : regset) (v : val) : Prop := - match rel ! x with + match (var_to_sym rel) ! x with | None => True | Some sym => sem_sym_val sym rs (Some (rs # x)) end. @@ -98,7 +98,7 @@ Proof. pose proof (SEM arg) as SEMarg. simpl. unfold forward_move. unfold sem_sym_val in *. - destruct (t ! arg); trivial. + destruct (_ ! arg); trivial. destruct s; congruence. Qed. @@ -127,7 +127,7 @@ Lemma kill_reg_sound : Proof. unfold sem_rel, kill_reg, sem_reg, sem_sym_val. intros until v. - intros REL x. + intros REL x. simpl. rewrite PTree.gfilter1. destruct (Pos.eq_dec dst x). { @@ -137,7 +137,7 @@ Proof. } rewrite PTree.gro by congruence. rewrite Regmap.gso by congruence. - destruct (rel ! x) as [relx | ] eqn:RELx; trivial. + destruct (_ ! x) as [relx | ] eqn:RELx; trivial. unfold kill_sym_val. pose proof (REL x) as RELinstx. rewrite RELx in RELinstx. @@ -190,12 +190,13 @@ Proof. { subst x. unfold sem_reg. + simpl. rewrite PTree.gss. rewrite Regmap.gss. unfold sem_reg in *. simpl. unfold forward_move. - destruct (rel ! src) as [ sv |]; simpl. + destruct (_ ! src) as [ sv |]; simpl. destruct sv eqn:SV; simpl in *. { destruct (peq dst src0). @@ -211,6 +212,7 @@ Proof. } rewrite Regmap.gso by congruence. unfold sem_reg. + simpl. rewrite PTree.gso by congruence. rewrite Regmap.gso in KILL by congruence. exact KILL. @@ -223,9 +225,10 @@ Lemma move_cases_neq: Proof. intros until a. intro NEQ. unfold kill_reg, forward_move. + simpl. rewrite PTree.gfilter1. rewrite PTree.gro by congruence. - destruct (rel ! a); simpl. + destruct (_ ! a); simpl. 2: congruence. destruct s. { @@ -259,9 +262,10 @@ Proof. unfold kill_reg. unfold sem_reg in RELa. unfold forward_move. + simpl. rewrite PTree.gfilter1. rewrite PTree.gro by auto. - destruct (rel ! a); simpl; trivial. + destruct (_ ! a); simpl; trivial. destruct s; simpl in *; destruct negb; simpl; congruence. Qed. @@ -285,6 +289,7 @@ Proof. { subst x. unfold sem_reg. + simpl. rewrite PTree.gss. rewrite Regmap.gss. simpl. @@ -294,6 +299,7 @@ Proof. } rewrite Regmap.gso by congruence. unfold sem_reg. + simpl. rewrite PTree.gso by congruence. rewrite Regmap.gso in KILL by congruence. exact KILL. @@ -341,13 +347,13 @@ Proof. | Some src => eval_operation genv sp op rs ## args m = Some rs # src end -> fold_left (fun (a : option reg) (p : positive * sym_val) => - find_op_fold op args a (fst p) (snd p)) (PTree.elements rel) start = + find_op_fold op args a (fst p) (snd p)) (PTree.elements (var_to_sym rel)) start = Some src -> eval_operation genv sp op rs ## args m = Some rs # src) as REC. { unfold sem_rel, sem_reg in REL. - generalize (PTree.elements_complete rel). - generalize (PTree.elements rel). + generalize (PTree.elements_complete (var_to_sym rel)). + generalize (PTree.elements (var_to_sym rel)). induction l; simpl. { intros. @@ -372,7 +378,7 @@ Proof. destruct eq_args; trivial. subst args0. simpl. - assert ((rel ! r) = Some (SOp op args)) as RELatr. + assert (((var_to_sym rel) ! r) = Some (SOp op args)) as RELatr. { apply COMPLETE. left. @@ -417,7 +423,7 @@ Proof. end -> fold_left (fun (a : option reg) (p : positive * sym_val) => - find_load_fold chunk addr args a (fst p) (snd p)) (PTree.elements rel) start = + find_load_fold chunk addr args a (fst p) (snd p)) (PTree.elements (var_to_sym rel)) start = Some src -> match eval_addressing genv sp addr rs##args with | Some a => (Mem.loadv chunk m a) = Some (rs # src) @@ -426,8 +432,8 @@ Proof. { unfold sem_rel, sem_reg in REL. - generalize (PTree.elements_complete rel). - generalize (PTree.elements rel). + generalize (PTree.elements_complete (var_to_sym rel)). + generalize (PTree.elements (var_to_sym rel)). induction l; simpl. { intros. @@ -454,7 +460,7 @@ Proof. destruct eq_args; trivial. subst args0. simpl. - assert ((rel ! r) = Some (SLoad chunk addr args)) as RELatr. + assert (((var_to_sym rel) ! r) = Some (SLoad chunk addr args)) as RELatr. { apply COMPLETE. left. @@ -497,7 +503,7 @@ Proof. 2: (apply IHargs; assumption). unfold forward_move, sem_rel, sem_reg, sem_sym_val in *. pose proof (REL a) as RELa. - destruct (rel ! a); trivial. + destruct (_ ! a); trivial. destruct s; congruence. Qed. @@ -580,6 +586,7 @@ Proof. { subst x. unfold sem_reg. + simpl. rewrite PTree.gss. rewrite Regmap.gss. simpl. @@ -588,6 +595,7 @@ Proof. } rewrite Regmap.gso by congruence. unfold sem_reg. + simpl. rewrite PTree.gso by congruence. rewrite Regmap.gso in KILL by congruence. exact KILL. @@ -661,6 +669,7 @@ Proof. intros REL x. pose proof (REL x) as RELx. unfold kill_reg, sem_reg in *. + simpl. rewrite PTree.gfilter1. destruct (peq res x). { subst x. @@ -668,7 +677,7 @@ Proof. reflexivity. } rewrite PTree.gro by congruence. - destruct (mpc ! x) as [sv | ]; trivial. + destruct (_ ! x) as [sv | ]; trivial. destruct negb; trivial. Qed. @@ -691,8 +700,8 @@ Proof. pose proof (RE x) as REx. pose proof (GE x) as GEx. unfold sem_reg in *. - destruct (r1 ! x) as [r1x | ] in *; - destruct (r2 ! x) as [r2x | ] in *; + destruct ((var_to_sym r1) ! x) as [r1x | ] in *; + destruct ((var_to_sym r2) ! x) as [r2x | ] in *; congruence. Qed. End SAME_MEMORY. @@ -708,9 +717,10 @@ Proof. intros SEM x. pose proof (SEM x) as SEMx. unfold kill_mem. + simpl. rewrite PTree.gfilter1. unfold kill_sym_val_mem. - destruct (rel ! x) as [ sv | ]. + destruct ((var_to_sym rel) ! x) as [ sv | ]. 2: reflexivity. destruct sv; simpl in *; trivial. { -- cgit From 04dc160b4962fedd1ef1b322809377a2fa1434a2 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 4 Feb 2020 14:03:35 +0100 Subject: gcombine_null --- lib/Maps.v | 40 +++++++++++++++++++++++++++++++++++++++- 1 file changed, 39 insertions(+), 1 deletion(-) diff --git a/lib/Maps.v b/lib/Maps.v index 9e44a7fe..e938f205 100644 --- a/lib/Maps.v +++ b/lib/Maps.v @@ -625,7 +625,45 @@ Module PTree <: TREE. auto. Qed. - Fixpoint xelements (A : Type) (m : t A) (i : positive) + Section COMBINE_NULL. + + Variables A B C: Type. + Variable f: A -> B -> option C. + + + Fixpoint combine_null (m1: t A) (m2: t B) {struct m1} : t C := + match m1, m2 with + | (Node l1 o1 r1), (Node l2 o2 r2) => + Node' (combine_null l1 l2) + (match o1, o2 with + | (Some x1), (Some x2) => f x1 x2 + | _, _ => None + end) + (combine_null r1 r2) + | _, _ => Leaf + end. + + Theorem gcombine_null: + forall (m1: t A) (m2: t B) (i: positive), + get i (combine_null m1 m2) = + match (get i m1), (get i m2) with + | (Some x1), (Some x2) => f x1 x2 + | _, _ => None + end. + Proof. + induction m1; intros; simpl. + - rewrite gleaf. rewrite gleaf. + reflexivity. + - destruct m2; simpl. + + rewrite gleaf. rewrite gleaf. + destruct get; reflexivity. + + rewrite gnode'. + destruct i; simpl; try rewrite IHm1_1; try rewrite IHm1; trivial. + Qed. + + End COMBINE_NULL. + + Fixpoint xelements (A : Type) (m : t A) (i : positive) (k: list (positive * A)) {struct m} : list (positive * A) := match m with -- cgit From ed0ad804bd09b2b76ec2535367ab9dd57b2600b0 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 4 Feb 2020 14:06:47 +0100 Subject: gcombine_null --- lib/Maps.v | 13 +++++++++++++ 1 file changed, 13 insertions(+) diff --git a/lib/Maps.v b/lib/Maps.v index e938f205..0beb11b4 100644 --- a/lib/Maps.v +++ b/lib/Maps.v @@ -116,6 +116,19 @@ Module Type TREE. forall (m1: t A) (m2: t B) (i: elt), get i (combine f m1 m2) = f (get i m1) (get i m2). + Parameter combine_null : + forall (A B C: Type) (f: A -> B -> option C), + t A -> t B -> t C. + + Axiom gcombine_null: + forall (A B C: Type) (f: A -> B -> option C), + forall (m1: t A) (m2: t B) (i: elt), + get i (combine_null f m1 m2) = + match (get i m1), (get i m2) with + | (Some x1), (Some x2) => f x1 x2 + | _, _ => None + end. + (** Enumerating the bindings of a tree. *) Parameter elements: forall (A: Type), t A -> list (elt * A). -- cgit From 4f5ea8b8373dc994714aa563182bad9c9ed21526 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 4 Feb 2020 14:53:25 +0100 Subject: gremove_t --- lib/Maps.v | 36 ++++++++++++++++++++++++++++++++++++ 1 file changed, 36 insertions(+) diff --git a/lib/Maps.v b/lib/Maps.v index 0beb11b4..5a0e6d5a 100644 --- a/lib/Maps.v +++ b/lib/Maps.v @@ -676,6 +676,42 @@ Module PTree <: TREE. End COMBINE_NULL. + Section REMOVE_TREE. + + Variables A B: Type. + + Fixpoint remove_t (m1: t A) (m2: t B) {struct m1} : t A := + match m1, m2 with + | Leaf, _ | _, Leaf => m1 + | (Node l1 o1 r1), (Node l2 o2 r2) => + Node' (remove_t l1 l2) + (match o2 with + | Some _ => None + | None => o1 + end) + (remove_t r1 r2) + end. + + Theorem gremove_t: + forall m1 : t A, + forall m2 : t B, + forall i : positive, + get i (remove_t m1 m2) = match get i m2 with + | None => get i m1 + | Some _ => None + end. + Proof. + induction m1; intros; simpl. + - rewrite gleaf. + destruct get; reflexivity. + - destruct m2; simpl. + + rewrite gleaf. + reflexivity. + + rewrite gnode'. + destruct i; simpl; try rewrite IHm1_1; try rewrite IHm1; trivial. + Qed. + End REMOVE_TREE. + Fixpoint xelements (A : Type) (m : t A) (i : positive) (k: list (positive * A)) {struct m} : list (positive * A) := -- cgit From 0c07f0f560547ae83f3398adcd53be31e7707a62 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 4 Feb 2020 16:01:59 +0100 Subject: begin well formedness --- backend/CSE2.v | 47 +++++++++++++++++-------- backend/CSE2proof.v | 98 +++++++++++++++++++++++++++++++++++++++++++++++++++++ 2 files changed, 130 insertions(+), 15 deletions(-) diff --git a/backend/CSE2.v b/backend/CSE2.v index 0479dad9..358fade4 100644 --- a/backend/CSE2.v +++ b/backend/CSE2.v @@ -29,9 +29,20 @@ Proof. decide equality. Defined. +Definition pset := PTree.t unit. + +Definition pset_inter : pset -> pset -> pset := + PTree.combine_null + (fun ov1 ov2 => Some tt). + +Definition pset_empty : pset := PTree.empty unit. +Definition pset_add (i : positive) (s : pset) := PTree.set i tt s. +Definition pset_remove (i : positive) (s : pset) := PTree.remove i s. + Record relmap := mkrel { - var_to_sym : (PTree.t sym_val) - }. + var_to_sym : PTree.t sym_val ; + mem_used : pset + }. Module RELATION. @@ -40,7 +51,7 @@ Definition t := relmap. Definition eq (r1 r2 : t) := forall x, (PTree.get x (var_to_sym r1)) = (PTree.get x (var_to_sym r2)). -Definition top : t := mkrel (PTree.empty sym_val). +Definition top : t := mkrel (PTree.empty sym_val) pset_empty. Lemma eq_refl: forall x, eq x x. Proof. @@ -119,7 +130,8 @@ Definition lub (r1 r2 : t) := | None, _ | _, None => None end) - (var_to_sym r1) (var_to_sym r2)). + (var_to_sym r1) (var_to_sym r2)) + (pset_inter (mem_used r1) (mem_used r2)). Lemma ge_lub_left: forall x y, ge (lub x y) x. Proof. @@ -268,9 +280,10 @@ Definition kill_sym_val (dst : reg) (sv : sym_val) := | SLoad chunk addr args => List.existsb (peq dst) args end. -Definition kill_reg (dst : reg) (rel : RELATION.t) := +Definition kill_reg (dst : reg) (rel : RELATION.t) : RELATION.t := mkrel (PTree.filter1 (fun x => negb (kill_sym_val dst x)) - (PTree.remove dst (var_to_sym rel))). + (PTree.remove dst (var_to_sym rel))) + (pset_remove dst (mem_used rel)). Definition kill_sym_val_mem (sv: sym_val) := match sv with @@ -281,7 +294,8 @@ Definition kill_sym_val_mem (sv: sym_val) := Definition kill_mem (rel : RELATION.t) := mkrel - (PTree.filter1 (fun x => negb (kill_sym_val_mem x)) (var_to_sym rel)). + (PTree.filter1 (fun x => negb (kill_sym_val_mem x)) (var_to_sym rel)) + pset_empty. Definition forward_move (rel : RELATION.t) (x : reg) : reg := @@ -291,7 +305,8 @@ Definition forward_move (rel : RELATION.t) (x : reg) : reg := end. Definition move (src dst : reg) (rel : RELATION.t) := - mkrel (PTree.set dst (SMove (forward_move rel src)) (var_to_sym (kill_reg dst rel))). + mkrel (PTree.set dst (SMove (forward_move rel src)) (var_to_sym (kill_reg dst rel))) + (mem_used rel). Definition find_op_fold op args (already : option reg) x sv := match already with @@ -328,34 +343,36 @@ Definition find_load (rel : RELATION.t) (chunk : memory_chunk) (addr : addressin PTree.fold (find_load_fold chunk addr args) (var_to_sym rel) None. Definition oper2 (op: operation) (dst : reg) (args : list reg) - (rel : RELATION.t) := + (rel : RELATION.t) : RELATION.t := let rel' := kill_reg dst rel in - mkrel (PTree.set dst (SOp op (List.map (forward_move rel') args)) (var_to_sym rel')). + mkrel (PTree.set dst (SOp op (List.map (forward_move rel') args)) (var_to_sym rel')) + (mem_used rel). Definition oper1 (op: operation) (dst : reg) (args : list reg) - (rel : RELATION.t) := + (rel : RELATION.t) : RELATION.t := if List.in_dec peq dst args then kill_reg dst rel else oper2 op dst args rel. Definition oper (op: operation) (dst : reg) (args : list reg) - (rel : RELATION.t) := + (rel : RELATION.t) : RELATION.t := match find_op rel op (List.map (forward_move rel) args) with | Some r => move r dst rel | None => oper1 op dst args rel end. Definition gen_oper (op: operation) (dst : reg) (args : list reg) - (rel : RELATION.t) := + (rel : RELATION.t) : RELATION.t := match op, args with | Omove, src::nil => move src dst rel | _, _ => oper op dst args rel end. Definition load2 (chunk: memory_chunk) (addr : addressing) - (dst : reg) (args : list reg) (rel : RELATION.t) := + (dst : reg) (args : list reg) (rel : RELATION.t) : RELATION.t := let rel' := kill_reg dst rel in - mkrel (PTree.set dst (SLoad chunk addr (List.map (forward_move rel') args)) (var_to_sym rel')). + mkrel (PTree.set dst (SLoad chunk addr (List.map (forward_move rel') args)) (var_to_sym rel')) + (pset_add dst (mem_used rel)). Definition load1 (chunk: memory_chunk) (addr : addressing) (dst : reg) (args : list reg) (rel : RELATION.t) := diff --git a/backend/CSE2proof.v b/backend/CSE2proof.v index a29bf202..bd917163 100644 --- a/backend/CSE2proof.v +++ b/backend/CSE2proof.v @@ -37,6 +37,104 @@ Proof. assumption. Qed. +Lemma gpset_inter_none: forall a b i, + (pset_inter a b) ! i = None <-> + (a ! i = None) \/ (b ! i = None). +Proof. + intros. unfold pset_inter. + rewrite PTree.gcombine_null. + destruct (a ! i); destruct (b ! i); intuition discriminate. +Qed. + +Lemma some_some: + forall x : option unit, + x <> None <-> x = Some tt. +Proof. + destruct x as [[] | ]; split; congruence. +Qed. + +Lemma gpset_inter_some: forall a b i, + (pset_inter a b) ! i = Some tt <-> + (a ! i = Some tt) /\ (b ! i = Some tt). +Proof. + intros. unfold pset_inter. + rewrite PTree.gcombine_null. + destruct (a ! i) as [[] | ]; destruct (b ! i) as [[] | ]; intuition congruence. +Qed. + +Definition wellformed (rel : RELATION.t) : Prop := + forall i sv, + (var_to_sym rel) ! i = Some sv -> + kill_sym_val_mem sv = true -> + (mem_used rel) ! i = Some tt. + +Lemma wellformed_top : wellformed RELATION.top. +Proof. + unfold wellformed, RELATION.top, pset_empty. + simpl. + intros. + rewrite PTree.gempty in *. + discriminate. +Qed. + +Lemma wellformed_lub : forall a b, + (wellformed a) -> (wellformed b) -> (wellformed (RELATION.lub a b)). +Proof. + unfold wellformed, RELATION.lub. + simpl. + intros a b Ha Hb. + intros i sv COMBINE KILLABLE. + rewrite PTree.gcombine in *. + 2: reflexivity. + destruct (var_to_sym a) ! i as [syma | ] eqn:SYMA; + destruct (var_to_sym b) ! i as [symb | ] eqn:SYMB; + try discriminate. + destruct (eq_sym_val syma symb); try discriminate. + subst syma. + inv COMBINE. + apply gpset_inter_some. + split; eauto. +Qed. + +Lemma wellformed_kill_reg: + forall dst rel, + (wellformed rel) -> (wellformed (kill_reg dst rel)). +Proof. + unfold wellformed, kill_reg, pset_remove. + simpl. + intros dst rel Hrel. + intros i sv KILLREG KILLABLE. + rewrite PTree.gfilter1 in KILLREG. + destruct (peq dst i). + { subst i. + rewrite PTree.grs in *. + discriminate. + } + rewrite PTree.gro in * by congruence. + specialize Hrel with (i := i). + eapply Hrel. + 2: eassumption. + destruct (_ ! i); try discriminate. + destruct negb; congruence. +Qed. + +Lemma wellformed_kill_mem: + forall rel, + (wellformed rel) -> (wellformed (kill_mem rel)). +Proof. + unfold wellformed, kill_mem. + simpl. + intros rel Hrel. + intros i sv KILLMEM KILLABLE. + rewrite PTree.gfilter1 in KILLMEM. + exfalso. + specialize Hrel with (i := i). + destruct ((var_to_sym rel) ! i) eqn:RELi. + 2: discriminate. + specialize Hrel with (sv := s). + destruct (kill_sym_val_mem s) eqn:KILLs; simpl in KILLMEM; congruence. +Qed. + Section SOUNDNESS. Variable F V : Type. Variable genv: Genv.t F V. -- cgit From e88c7fa00a5174ecf897b3cb59b7adee818a1788 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 4 Feb 2020 16:59:18 +0100 Subject: wellformedness for memory --- backend/CSE2.v | 10 ++-- backend/CSE2proof.v | 165 +++++++++++++++++++++++++++++++++++++++++++++++----- 2 files changed, 158 insertions(+), 17 deletions(-) diff --git a/backend/CSE2.v b/backend/CSE2.v index 358fade4..2fc0c323 100644 --- a/backend/CSE2.v +++ b/backend/CSE2.v @@ -305,8 +305,9 @@ Definition forward_move (rel : RELATION.t) (x : reg) : reg := end. Definition move (src dst : reg) (rel : RELATION.t) := - mkrel (PTree.set dst (SMove (forward_move rel src)) (var_to_sym (kill_reg dst rel))) - (mem_used rel). + let rel0 := kill_reg dst rel in + mkrel (PTree.set dst (SMove (forward_move rel src)) (var_to_sym rel0)) + (mem_used rel0). Definition find_op_fold op args (already : option reg) x sv := match already with @@ -346,7 +347,8 @@ Definition oper2 (op: operation) (dst : reg) (args : list reg) (rel : RELATION.t) : RELATION.t := let rel' := kill_reg dst rel in mkrel (PTree.set dst (SOp op (List.map (forward_move rel') args)) (var_to_sym rel')) - (mem_used rel). + (if op_depends_on_memory op then (pset_add dst (mem_used rel')) + else mem_used rel'). Definition oper1 (op: operation) (dst : reg) (args : list reg) (rel : RELATION.t) : RELATION.t := @@ -372,7 +374,7 @@ Definition load2 (chunk: memory_chunk) (addr : addressing) (dst : reg) (args : list reg) (rel : RELATION.t) : RELATION.t := let rel' := kill_reg dst rel in mkrel (PTree.set dst (SLoad chunk addr (List.map (forward_move rel') args)) (var_to_sym rel')) - (pset_add dst (mem_used rel)). + (pset_add dst (mem_used rel')). Definition load1 (chunk: memory_chunk) (addr : addressing) (dst : reg) (args : list reg) (rel : RELATION.t) := diff --git a/backend/CSE2proof.v b/backend/CSE2proof.v index bd917163..15ce3ffb 100644 --- a/backend/CSE2proof.v +++ b/backend/CSE2proof.v @@ -62,25 +62,26 @@ Proof. destruct (a ! i) as [[] | ]; destruct (b ! i) as [[] | ]; intuition congruence. Qed. -Definition wellformed (rel : RELATION.t) : Prop := +Definition wellformed_mem (rel : RELATION.t) : Prop := forall i sv, (var_to_sym rel) ! i = Some sv -> kill_sym_val_mem sv = true -> (mem_used rel) ! i = Some tt. -Lemma wellformed_top : wellformed RELATION.top. +Lemma wellformed_mem_top : wellformed_mem RELATION.top. Proof. - unfold wellformed, RELATION.top, pset_empty. + unfold wellformed_mem, RELATION.top, pset_empty. simpl. intros. rewrite PTree.gempty in *. discriminate. Qed. +Local Hint Resolve wellformed_mem_top : wellformed. -Lemma wellformed_lub : forall a b, - (wellformed a) -> (wellformed b) -> (wellformed (RELATION.lub a b)). +Lemma wellformed_mem_lub : forall a b, + (wellformed_mem a) -> (wellformed_mem b) -> (wellformed_mem (RELATION.lub a b)). Proof. - unfold wellformed, RELATION.lub. + unfold wellformed_mem, RELATION.lub. simpl. intros a b Ha Hb. intros i sv COMBINE KILLABLE. @@ -95,12 +96,13 @@ Proof. apply gpset_inter_some. split; eauto. Qed. +Local Hint Resolve wellformed_mem_lub : wellformed. -Lemma wellformed_kill_reg: +Lemma wellformed_mem_kill_reg: forall dst rel, - (wellformed rel) -> (wellformed (kill_reg dst rel)). + (wellformed_mem rel) -> (wellformed_mem (kill_reg dst rel)). Proof. - unfold wellformed, kill_reg, pset_remove. + unfold wellformed_mem, kill_reg, pset_remove. simpl. intros dst rel Hrel. intros i sv KILLREG KILLABLE. @@ -117,12 +119,13 @@ Proof. destruct (_ ! i); try discriminate. destruct negb; congruence. Qed. +Local Hint Resolve wellformed_mem_kill_reg : wellformed. -Lemma wellformed_kill_mem: +Lemma wellformed_mem_kill_mem: forall rel, - (wellformed rel) -> (wellformed (kill_mem rel)). + (wellformed_mem rel) -> (wellformed_mem (kill_mem rel)). Proof. - unfold wellformed, kill_mem. + unfold wellformed_mem, kill_mem. simpl. intros rel Hrel. intros i sv KILLMEM KILLABLE. @@ -134,7 +137,143 @@ Proof. specialize Hrel with (sv := s). destruct (kill_sym_val_mem s) eqn:KILLs; simpl in KILLMEM; congruence. Qed. - +Local Hint Resolve wellformed_mem_kill_mem : wellformed. + +Lemma wellformed_mem_move: + forall r dst rel, + (wellformed_mem rel) -> (wellformed_mem (move r dst rel)). +Proof. + Local Opaque kill_reg. + intros; unfold move, wellformed_mem; simpl. + intros i sv MOVE KILL. + destruct (peq i dst). + { subst i. + rewrite PTree.gss in MOVE. + inv MOVE. + simpl in KILL. + discriminate. + } + rewrite PTree.gso in MOVE by congruence. + eapply wellformed_mem_kill_reg; eauto. +Qed. +Local Hint Resolve wellformed_mem_move : wellformed. + +Lemma wellformed_mem_oper2: + forall op dst args rel, + (wellformed_mem rel) -> + (wellformed_mem (oper2 op dst args rel)). +Proof. + intros until rel. intro WELLFORMED. + unfold oper2. + intros i sv MOVE OPER. + simpl in *. + destruct (peq i dst). + { subst i. + rewrite PTree.gss in MOVE. + inv MOVE. + simpl in OPER. + rewrite OPER. + apply PTree.gss. + } + unfold pset_add. + rewrite PTree.gso in MOVE by congruence. + destruct op_depends_on_memory. + - rewrite PTree.gso by congruence. + eapply wellformed_mem_kill_reg; eauto. + - eapply wellformed_mem_kill_reg; eauto. +Qed. +Local Hint Resolve wellformed_mem_oper2 : wellformed. + +Lemma wellformed_mem_oper1: + forall op dst args rel, + (wellformed_mem rel) -> + (wellformed_mem (oper1 op dst args rel)). +Proof. + unfold oper1. intros. + destruct in_dec ; auto with wellformed. +Qed. +Local Hint Resolve wellformed_mem_oper1 : wellformed. + +Lemma wellformed_mem_oper: + forall op dst args rel, + (wellformed_mem rel) -> + (wellformed_mem (oper op dst args rel)). +Proof. + unfold oper. intros. + destruct find_op ; auto with wellformed. +Qed. +Local Hint Resolve wellformed_mem_oper : wellformed. + +Lemma wellformed_mem_gen_oper: + forall op dst args rel, + (wellformed_mem rel) -> + (wellformed_mem (gen_oper op dst args rel)). +Proof. + intros. + unfold gen_oper. + destruct op; auto with wellformed. + destruct args; auto with wellformed. + destruct args; auto with wellformed. +Qed. +Local Hint Resolve wellformed_mem_gen_oper : wellformed. + +Lemma wellformed_mem_load2 : + forall chunk addr dst args rel, + (wellformed_mem rel) -> + (wellformed_mem (load2 chunk addr dst args rel)). +Proof. + intros. + unfold load2, wellformed_mem. + simpl. + intros i sv LOAD KILL. + destruct (peq i dst). + { subst i. + apply PTree.gss. + } + unfold pset_add. + rewrite PTree.gso in * by congruence. + eapply wellformed_mem_kill_reg; eauto. +Qed. +Local Hint Resolve wellformed_mem_load2 : wellformed. + +Lemma wellformed_mem_load1 : + forall chunk addr dst args rel, + (wellformed_mem rel) -> + (wellformed_mem (load1 chunk addr dst args rel)). +Proof. + intros. + unfold load1. + destruct in_dec; eauto with wellformed. +Qed. +Local Hint Resolve wellformed_mem_load1 : wellformed. + +Lemma wellformed_mem_load : + forall chunk addr dst args rel, + (wellformed_mem rel) -> + (wellformed_mem (load chunk addr dst args rel)). +Proof. + intros. + unfold load. + destruct find_load; eauto with wellformed. +Qed. +Local Hint Resolve wellformed_mem_load : wellformed. + +Definition wellformed_mem' (rb : RB.t) := + match rb with + | None => True + | Some r => wellformed_mem r + end. + +Definition wellformed_apply_instr: + forall instr rel, + (wellformed_mem rel) -> + (wellformed_mem' (apply_instr instr rel)). +Proof. + destruct instr; simpl; auto with wellformed. +Qed. + +Local Transparent kill_reg. + Section SOUNDNESS. Variable F V : Type. Variable genv: Genv.t F V. -- cgit From b55e522ca286bbe96be3ce5fc05c984b2a4a130a Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 4 Feb 2020 17:23:41 +0100 Subject: invariant guaranteed --- backend/CSE2.v | 4 +++- backend/CSE2proof.v | 56 ++++++++++++++++++++++++++++++++++++++++++++++++++--- 2 files changed, 56 insertions(+), 4 deletions(-) diff --git a/backend/CSE2.v b/backend/CSE2.v index 2fc0c323..a818996b 100644 --- a/backend/CSE2.v +++ b/backend/CSE2.v @@ -268,6 +268,8 @@ Module ADD_BOTTOM(L : SEMILATTICE_WITHOUT_BOTTOM). - apply L.ge_refl. apply L.eq_refl. Qed. + + Definition top := Some RELATION.top. End ADD_BOTTOM. Module RB := ADD_BOTTOM(RELATION). @@ -441,7 +443,7 @@ Definition apply_instr' code (pc : node) (ro : RB.t) : RB.t := Definition forward_map (f : RTL.function) := DS.fixpoint (RTL.fn_code f) RTL.successors_instr - (apply_instr' (RTL.fn_code f)) (RTL.fn_entrypoint f) (Some RELATION.top). + (apply_instr' (RTL.fn_code f)) (RTL.fn_entrypoint f) RB.top. Definition forward_move_b (rb : RB.t) (x : reg) := match rb with diff --git a/backend/CSE2proof.v b/backend/CSE2proof.v index 15ce3ffb..351a4219 100644 --- a/backend/CSE2proof.v +++ b/backend/CSE2proof.v @@ -258,19 +258,69 @@ Proof. Qed. Local Hint Resolve wellformed_mem_load : wellformed. -Definition wellformed_mem' (rb : RB.t) := +Definition wellformed_mem_rb (rb : RB.t) := match rb with | None => True | Some r => wellformed_mem r end. -Definition wellformed_apply_instr: +Lemma wellformed_mem_apply_instr: forall instr rel, (wellformed_mem rel) -> - (wellformed_mem' (apply_instr instr rel)). + (wellformed_mem_rb (apply_instr instr rel)). Proof. destruct instr; simpl; auto with wellformed. Qed. +Local Hint Resolve wellformed_mem_apply_instr : wellformed. + +Lemma wellformed_mem_rb_bot : + wellformed_mem_rb RB.bot. +Proof. + simpl. + trivial. +Qed. +Local Hint Resolve wellformed_mem_rb_bot: wellformed. + +Lemma wellformed_mem_rb_top : + wellformed_mem_rb RB.top. +Proof. + simpl. + auto with wellformed. +Qed. +Local Hint Resolve wellformed_mem_rb_top: wellformed. + +Lemma wellformed_mem_rb_apply_instr': + forall code pc rel, + (wellformed_mem_rb rel) -> + (wellformed_mem_rb (apply_instr' code pc rel)). +Proof. + destruct rel; simpl; trivial. + intro. + destruct (code ! pc); + eauto with wellformed. +Qed. +Local Hint Resolve wellformed_mem_rb_apply_instr' : wellformed. + +Lemma wellformed_mem_rb_lub : forall a b, + (wellformed_mem_rb a) -> (wellformed_mem_rb b) -> (wellformed_mem_rb (RB.lub a b)). +Proof. + destruct a; destruct b; simpl; auto with wellformed. +Qed. +Local Hint Resolve wellformed_mem_rb_lub: wellformed. + +Definition wellformed_mem_fmap fmap := + forall i, wellformed_mem_rb (fmap !! i). + +Theorem wellformed_mem_forward_map : forall f, + forall fmap, (forward_map f) = Some fmap -> + wellformed_mem_fmap fmap. +Proof. + intros. + unfold forward_map in *. + unfold wellformed_mem_fmap. + intro. + eapply DS.fixpoint_invariant with (ev := Some RELATION.top); eauto with wellformed. +Qed. Local Transparent kill_reg. -- cgit From e0257f612a1358ad9927bd198cb11798cd8ccae4 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 4 Feb 2020 17:57:05 +0100 Subject: kill memory focused --- backend/CSE2.v | 2 +- backend/CSE2proof.v | 64 +++++++++++++++++++++++++++++++++++++---------------- 2 files changed, 46 insertions(+), 20 deletions(-) diff --git a/backend/CSE2.v b/backend/CSE2.v index a818996b..a76104af 100644 --- a/backend/CSE2.v +++ b/backend/CSE2.v @@ -296,7 +296,7 @@ Definition kill_sym_val_mem (sv: sym_val) := Definition kill_mem (rel : RELATION.t) := mkrel - (PTree.filter1 (fun x => negb (kill_sym_val_mem x)) (var_to_sym rel)) + (PTree.remove_t (var_to_sym rel) (mem_used rel)) pset_empty. diff --git a/backend/CSE2proof.v b/backend/CSE2proof.v index 351a4219..825cfbcd 100644 --- a/backend/CSE2proof.v +++ b/backend/CSE2proof.v @@ -129,13 +129,16 @@ Proof. simpl. intros rel Hrel. intros i sv KILLMEM KILLABLE. - rewrite PTree.gfilter1 in KILLMEM. + rewrite PTree.gremove_t in KILLMEM. exfalso. specialize Hrel with (i := i). destruct ((var_to_sym rel) ! i) eqn:RELi. - 2: discriminate. - specialize Hrel with (sv := s). - destruct (kill_sym_val_mem s) eqn:KILLs; simpl in KILLMEM; congruence. + - specialize Hrel with (sv := s). + destruct ((mem_used rel) ! i) eqn:USEDi. + discriminate. + inv KILLMEM. + intuition congruence. + - destruct ((mem_used rel) ! i); discriminate. Qed. Local Hint Resolve wellformed_mem_kill_mem : wellformed. @@ -321,6 +324,7 @@ Proof. intro. eapply DS.fixpoint_invariant with (ev := Some RELATION.top); eauto with wellformed. Qed. +Local Hint Resolve wellformed_mem_forward_map: wellformed. Local Transparent kill_reg. @@ -997,24 +1001,33 @@ Lemma kill_mem_sound : forall m m' : mem, forall rel : RELATION.t, forall rs, + wellformed_mem rel -> sem_rel m rel rs -> sem_rel m' (kill_mem rel) rs. Proof. unfold sem_rel, sem_reg. intros until rs. - intros SEM x. - pose proof (SEM x) as SEMx. + intros WELLFORMED SEM x. + unfold wellformed_mem in *. + specialize SEM with (x := x). + specialize WELLFORMED with (i := x). unfold kill_mem. simpl. - rewrite PTree.gfilter1. + rewrite PTree.gremove_t. unfold kill_sym_val_mem. - destruct ((var_to_sym rel) ! x) as [ sv | ]. - 2: reflexivity. - destruct sv; simpl in *; trivial. - { - destruct op_depends_on_memory eqn:DEPENDS; simpl; trivial. - rewrite SEMx. - apply op_depends_on_memory_correct; auto. - } + destruct ((var_to_sym rel) ! x) as [ sv | ] eqn:VAR. + - specialize WELLFORMED with (sv0 := sv). + destruct ((mem_used rel) ! x) eqn:USED; trivial. + unfold sem_sym_val in *. + destruct sv; simpl in *; trivial. + + rewrite op_depends_on_memory_correct with (m2 := m). + assumption. + destruct op_depends_on_memory; intuition congruence. + + intuition discriminate. + - replace (match (mem_used rel) ! x with + | Some _ | _ => None + end) with (@None sym_val) by + (destruct ((mem_used rel) ! x); trivial). + trivial. Qed. End SOUNDNESS. @@ -1155,6 +1168,21 @@ Ltac TR_AT := generalize (transf_function_at _ _ _ A); intros end. +Lemma wellformed_mem_mpc: + forall f map pc mpc, + (forward_map f) = Some map -> + map # pc = Some mpc -> + wellformed_mem mpc. +Proof. + intros. + assert (wellformed_mem_fmap map) by eauto with wellformed. + unfold wellformed_mem_fmap, wellformed_mem_rb in *. + specialize H1 with (i := pc). + rewrite H0 in H1. + assumption. +Qed. +Hint Resolve wellformed_mem_mpc : wellformed. + Lemma step_simulation: forall S1 t S2, RTL.step ge S1 t S2 -> forall S1', match_states S1 S1' -> @@ -1413,8 +1441,7 @@ Proof. rewrite H. reflexivity. } - apply (kill_mem_sound' sp m). - assumption. + apply (kill_mem_sound' sp m); eauto with wellformed. (* call *) - econstructor; split. @@ -1444,8 +1471,7 @@ Proof. reflexivity. } apply kill_reg_sound. - apply (kill_mem_sound' sp m). - assumption. + apply (kill_mem_sound' sp m); eauto with wellformed. } (* tailcall *) -- cgit From c877fef46e86bbd088e3d3937ce96572bee0101d Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 4 Feb 2020 20:31:27 +0100 Subject: wellformedness for reg begins --- backend/CSE2proof.v | 53 +++++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 53 insertions(+) diff --git a/backend/CSE2proof.v b/backend/CSE2proof.v index 825cfbcd..86bc123d 100644 --- a/backend/CSE2proof.v +++ b/backend/CSE2proof.v @@ -68,6 +68,13 @@ Definition wellformed_mem (rel : RELATION.t) : Prop := kill_sym_val_mem sv = true -> (mem_used rel) ! i = Some tt. +Definition wellformed_reg (rel : RELATION.t) : Prop := + forall i j sv, + (var_to_sym rel) ! i = Some sv -> + kill_sym_val j sv = true -> + exists uses, (reg_used rel) ! j = Some uses /\ + uses ! i = Some tt. + Lemma wellformed_mem_top : wellformed_mem RELATION.top. Proof. unfold wellformed_mem, RELATION.top, pset_empty. @@ -78,6 +85,16 @@ Proof. Qed. Local Hint Resolve wellformed_mem_top : wellformed. +Lemma wellformed_reg_top : wellformed_reg RELATION.top. +Proof. + unfold wellformed_reg, RELATION.top, pset_empty. + simpl. + intros. + rewrite PTree.gempty in *. + discriminate. +Qed. +Local Hint Resolve wellformed_reg_top : wellformed. + Lemma wellformed_mem_lub : forall a b, (wellformed_mem a) -> (wellformed_mem b) -> (wellformed_mem (RELATION.lub a b)). Proof. @@ -98,6 +115,42 @@ Proof. Qed. Local Hint Resolve wellformed_mem_lub : wellformed. +Lemma wellformed_reg_lub : forall a b, + (wellformed_reg a) -> (wellformed_reg b) -> (wellformed_reg (RELATION.lub a b)). +Proof. + unfold wellformed_reg, RELATION.lub. + simpl. + intros a b Ha Hb. + intros i j sv COMBINE KILLABLE. + specialize Ha with (i := i) (j := j). + specialize Hb with (i := i) (j := j). + rewrite PTree.gcombine in *. + 2: reflexivity. + destruct (var_to_sym a) ! i as [syma | ] eqn:SYMA; + destruct (var_to_sym b) ! i as [symb | ] eqn:SYMB; + try discriminate. + specialize Ha with (sv := syma). + specialize Hb with (sv := symb). + destruct (eq_sym_val syma symb); try discriminate. + subst syma. + inv COMBINE. + assert (exists usesA : pset, (reg_used a) ! j = Some usesA /\ usesA ! i = Some tt) as USESA by auto. + assert (exists usesB : pset, (reg_used b) ! j = Some usesB /\ usesB ! i = Some tt) as USESB by auto. + destruct USESA as [uA [uA1 uA2]]. + destruct USESB as [uB [uB1 uB2]]. + rewrite PTree.gcombine_null. + rewrite uA1. + rewrite uB1. + pose proof (PTree.bempty_canon_correct (pset_inter uA uB) i) as EMPTY. + destruct PTree.bempty_canon. + - rewrite gpset_inter_none in EMPTY. + intuition congruence. + - econstructor; split; eauto. + rewrite gpset_inter_some. + auto. +Qed. +Local Hint Resolve wellformed_reg_lub : wellformed. + Lemma wellformed_mem_kill_reg: forall dst rel, (wellformed_mem rel) -> (wellformed_mem (kill_reg dst rel)). -- cgit From dabfcfe9aeffe4bcb3d2a27a46e2b10b5d725154 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 4 Feb 2020 21:08:04 +0100 Subject: wellformedness for reg begins; simplified --- backend/CSE2proof.v | 21 +++++++++------------ 1 file changed, 9 insertions(+), 12 deletions(-) diff --git a/backend/CSE2proof.v b/backend/CSE2proof.v index 86bc123d..350cdb24 100644 --- a/backend/CSE2proof.v +++ b/backend/CSE2proof.v @@ -72,8 +72,10 @@ Definition wellformed_reg (rel : RELATION.t) : Prop := forall i j sv, (var_to_sym rel) ! i = Some sv -> kill_sym_val j sv = true -> - exists uses, (reg_used rel) ! j = Some uses /\ - uses ! i = Some tt. + match (reg_used rel) ! j with + | Some uses => uses ! i = Some tt + | None => False + end. Lemma wellformed_mem_top : wellformed_mem RELATION.top. Proof. @@ -134,20 +136,15 @@ Proof. destruct (eq_sym_val syma symb); try discriminate. subst syma. inv COMBINE. - assert (exists usesA : pset, (reg_used a) ! j = Some usesA /\ usesA ! i = Some tt) as USESA by auto. - assert (exists usesB : pset, (reg_used b) ! j = Some usesB /\ usesB ! i = Some tt) as USESB by auto. - destruct USESA as [uA [uA1 uA2]]. - destruct USESB as [uB [uB1 uB2]]. rewrite PTree.gcombine_null. - rewrite uA1. - rewrite uB1. - pose proof (PTree.bempty_canon_correct (pset_inter uA uB) i) as EMPTY. - destruct PTree.bempty_canon. + destruct ((reg_used a) ! j) as [uA| ]; destruct ((reg_used b) ! j) as [uB | ]; auto. + { pose proof (PTree.bempty_canon_correct (pset_inter uA uB) i) as EMPTY. + destruct PTree.bempty_canon. - rewrite gpset_inter_none in EMPTY. intuition congruence. - - econstructor; split; eauto. - rewrite gpset_inter_some. + - rewrite gpset_inter_some. auto. + } Qed. Local Hint Resolve wellformed_reg_lub : wellformed. -- cgit From ad16517ee345e53398c69c62d975474475880799 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 5 Feb 2020 15:37:58 +0100 Subject: progress on wellformed reg --- backend/CSE2.v | 64 +++++++++++++++++++++------ backend/CSE2proof.v | 32 ++++++++++++++ lib/Maps.v | 121 +++++++++++++++++++++++++++++++++++++++++++++++++++- 3 files changed, 204 insertions(+), 13 deletions(-) diff --git a/backend/CSE2.v b/backend/CSE2.v index a76104af..1e3bc3b7 100644 --- a/backend/CSE2.v +++ b/backend/CSE2.v @@ -41,7 +41,8 @@ Definition pset_remove (i : positive) (s : pset) := PTree.remove i s. Record relmap := mkrel { var_to_sym : PTree.t sym_val ; - mem_used : pset + mem_used : pset ; + reg_used : PTree.t pset }. Module RELATION. @@ -51,7 +52,7 @@ Definition t := relmap. Definition eq (r1 r2 : t) := forall x, (PTree.get x (var_to_sym r1)) = (PTree.get x (var_to_sym r2)). -Definition top : t := mkrel (PTree.empty sym_val) pset_empty. +Definition top : t := mkrel (PTree.empty sym_val) pset_empty (PTree.empty pset). Lemma eq_refl: forall x, eq x x. Proof. @@ -131,7 +132,13 @@ Definition lub (r1 r2 : t) := | _, None => None end) (var_to_sym r1) (var_to_sym r2)) - (pset_inter (mem_used r1) (mem_used r2)). + (pset_inter (mem_used r1) (mem_used r2)) + (PTree.combine_null + (fun x y => let r := pset_inter x y in + if PTree.bempty_canon r + then None + else Some r) + (reg_used r1) (reg_used r2)). Lemma ge_lub_left: forall x y, ge (lub x y) x. Proof. @@ -275,17 +282,46 @@ End ADD_BOTTOM. Module RB := ADD_BOTTOM(RELATION). Module DS := Dataflow_Solver(RB)(NodeSetForward). -Definition kill_sym_val (dst : reg) (sv : sym_val) := +Definition kill_sym_val (dst : reg) (sv : sym_val) : bool := match sv with | SMove src => if peq dst src then true else false | SOp op args => List.existsb (peq dst) args | SLoad chunk addr args => List.existsb (peq dst) args end. - + +Definition referenced_by sv := + match sv with + | SMove src => src :: nil + | SOp op args => args + | SLoad chunk addr args => args + end. + +Definition referenced_by' osv := + match osv with + | None => nil + | Some sv => referenced_by sv + end. + +Definition remove_from_sets + (to_remove : reg) : + list reg -> PTree.t pset -> PTree.t pset := + List.fold_left + (fun sets reg => + match PTree.get reg sets with + | None => sets + | Some xset => + let xset' := PTree.remove to_remove xset in + if PTree.bempty_canon xset' + then PTree.remove reg sets + else PTree.set reg xset' sets + end). + Definition kill_reg (dst : reg) (rel : RELATION.t) : RELATION.t := - mkrel (PTree.filter1 (fun x => negb (kill_sym_val dst x)) - (PTree.remove dst (var_to_sym rel))) - (pset_remove dst (mem_used rel)). + let rel' := PTree.remove dst (var_to_sym rel) in + mkrel (PTree.filter1 (fun x => negb (kill_sym_val dst x)) rel') + (pset_remove dst (mem_used rel)) + (remove_from_sets dst (referenced_by' (PTree.get dst (var_to_sym rel))) + (PTree.remove dst (reg_used rel))). Definition kill_sym_val_mem (sv: sym_val) := match sv with @@ -297,7 +333,8 @@ Definition kill_sym_val_mem (sv: sym_val) := Definition kill_mem (rel : RELATION.t) := mkrel (PTree.remove_t (var_to_sym rel) (mem_used rel)) - pset_empty. + pset_empty + (reg_used rel). (* FIXME *) Definition forward_move (rel : RELATION.t) (x : reg) : reg := @@ -309,7 +346,8 @@ Definition forward_move (rel : RELATION.t) (x : reg) : reg := Definition move (src dst : reg) (rel : RELATION.t) := let rel0 := kill_reg dst rel in mkrel (PTree.set dst (SMove (forward_move rel src)) (var_to_sym rel0)) - (mem_used rel0). + (mem_used rel0) + (reg_used rel0). (* FIXME *) Definition find_op_fold op args (already : option reg) x sv := match already with @@ -350,7 +388,8 @@ Definition oper2 (op: operation) (dst : reg) (args : list reg) let rel' := kill_reg dst rel in mkrel (PTree.set dst (SOp op (List.map (forward_move rel') args)) (var_to_sym rel')) (if op_depends_on_memory op then (pset_add dst (mem_used rel')) - else mem_used rel'). + else mem_used rel') + (reg_used rel'). (* FIXME *) Definition oper1 (op: operation) (dst : reg) (args : list reg) (rel : RELATION.t) : RELATION.t := @@ -376,7 +415,8 @@ Definition load2 (chunk: memory_chunk) (addr : addressing) (dst : reg) (args : list reg) (rel : RELATION.t) : RELATION.t := let rel' := kill_reg dst rel in mkrel (PTree.set dst (SLoad chunk addr (List.map (forward_move rel') args)) (var_to_sym rel')) - (pset_add dst (mem_used rel')). + (pset_add dst (mem_used rel')) + (reg_used rel'). (* FIXME *) Definition load1 (chunk: memory_chunk) (addr : addressing) (dst : reg) (args : list reg) (rel : RELATION.t) := diff --git a/backend/CSE2proof.v b/backend/CSE2proof.v index 350cdb24..e0244518 100644 --- a/backend/CSE2proof.v +++ b/backend/CSE2proof.v @@ -171,6 +171,38 @@ Proof. Qed. Local Hint Resolve wellformed_mem_kill_reg : wellformed. +Lemma kill_sym_val_referenced: + forall dst, + forall sv, + (kill_sym_val dst sv)=true <-> In dst (referenced_by sv). +Proof. + intros. + destruct sv; simpl. + - destruct peq; intuition congruence. + - rewrite existsb_exists. + split. + + intros [x [IN EQ]]. + destruct peq. + * subst x; trivial. + * discriminate. + + intro. + exists dst. + split; trivial. + destruct peq; trivial. + congruence. + - rewrite existsb_exists. + split. + + intros [x [IN EQ]]. + destruct peq. + * subst x; trivial. + * discriminate. + + intro. + exists dst. + split; trivial. + destruct peq; trivial. + congruence. +Qed. + Lemma wellformed_mem_kill_mem: forall rel, (wellformed_mem rel) -> (wellformed_mem (kill_mem rel)). diff --git a/lib/Maps.v b/lib/Maps.v index 5a0e6d5a..ec1b0bee 100644 --- a/lib/Maps.v +++ b/lib/Maps.v @@ -164,6 +164,12 @@ Module Type TREE. forall (A B: Type) (f: B -> A -> B) (v: B) (m: t A), fold1 f m v = List.fold_left (fun a p => f a (snd p)) (elements m) v. + + Parameter bempty_canon : + forall (A : Type), t A -> bool. + Axiom bempty_canon_correct: + forall (A : Type) (tr : t A) (i : elt), + bempty_canon tr = true -> get i tr = None. End TREE. (** * The abstract signatures of maps *) @@ -274,6 +280,12 @@ Module PTree <: TREE. induction i; simpl; auto. Qed. + Definition bempty_canon (A : Type) (tr : t A) : bool := + match tr with + | Leaf => true + | _ => false + end. + Theorem gss: forall (A: Type) (i: positive) (x: A) (m: t A), get i (set i x m) = Some x. Proof. @@ -282,7 +294,16 @@ Module PTree <: TREE. Lemma gleaf : forall (A : Type) (i : positive), get i (Leaf : t A) = None. Proof. exact gempty. Qed. - + + Lemma bempty_canon_correct: + forall (A : Type) (tr : t A) (i : elt), + bempty_canon tr = true -> get i tr = None. + Proof. + destruct tr; intros. + - rewrite gleaf; trivial. + - discriminate. + Qed. + Theorem gso: forall (A: Type) (i j: positive) (x: A) (m: t A), i <> j -> get i (set j x m) = get i m. @@ -1045,6 +1066,104 @@ Module PTree <: TREE. intros. apply fold1_xelements with (l := @nil (positive * A)). Qed. + (* DM + Fixpoint xfind_any (A : Type) (P : elt -> A -> bool) (i : elt) (m : t A): + option elt := + match m with + | Leaf => None + | Node l None r => + match xfind_any P (xO i) l with + | None => xfind_any P (xI i) r + | r => r + end + | Node l (Some x) r => + if P i x + then Some i + else + match xfind_any P (xO i) l with + | None => xfind_any P (xI i) r + | r => r + end + end. + + Definition find_any (A : Type) (P : elt -> A -> bool) (m : t A) := + xfind_any P xH m. + + Fixpoint pos_concat (i : positive) (j : positive) := + match i with + | xI r => xI (pos_concat r j) + | xO r => xO (pos_concat r j) + | xH => j + end. + + Lemma pos_concat_assoc : + forall i j k, + (pos_concat (pos_concat i j) k) = (pos_concat i (pos_concat j k)). + Admitted. + + Lemma pos_concat_eq_l : forall i j, + (pos_concat i j) = i -> j = xH. + Proof. + induction i; simpl; intros j EQ. + - inv EQ. auto. + - inv EQ. auto. + - trivial. + Qed. + + Lemma pos_concat_eq_r : forall i j, + (pos_concat i j) = j -> i = xH. + Admitted. + + Local Hint Resolve pos_concat_eq_r : trees. + + Lemma pos_concat_xH_r: forall i, (pos_concat i xH) = i. + Proof. + induction i; simpl; try rewrite IHi; trivial. + Qed. + + Local Hint Resolve pos_concat_xH_r : trees. + + (* + Fixpoint pos_is_prefix (i j : elt) := + match i, j with + | xI i1, xI j1 => pos_is_prefix i1 j1 + | xO i1, xO j1 => pos_is_prefix i1 j1 + | xH, _ => true + | _, _ => false + end. + *) + + Lemma xfind_any_correct: + forall A P (m : t A) i k, + xfind_any P i m = Some k -> exists j v, k = pos_concat j i + /\ (get j m)=Some v /\ (P k v)=true. + Proof. + induction m; simpl. + { (* leaf *) + discriminate. + } + intros i k FOUND. + destruct o. + { destruct P eqn:Pval in FOUND. + { inv FOUND. + exists xH. exists a. + simpl. + eauto. + } + specialize IHm1 with (i := xO i) (k := k). + specialize IHm2 with (i := xI i) (k := k). + assert (Some k = Some k) as SK by trivial. + destruct (xfind_any P (xO i) m1). + { + inv FOUND. + destruct (IHm1 SK) as [j' [v [EQk [GET PP]]]]. + exists (pos_concat j' (xO xH)). exists v. + rewrite pos_concat_assoc. + simpl. + split; trivial. + split; trivial. + } + *) End PTree. (** * An implementation of maps over type [positive] *) -- cgit From 56dc34adcd147d8ac29f57edc9da1718a493dcce Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 5 Feb 2020 17:26:28 +0100 Subject: wellformed_reg_kill_reg --- backend/CSE2proof.v | 129 ++++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 129 insertions(+) diff --git a/backend/CSE2proof.v b/backend/CSE2proof.v index e0244518..e62e2ae6 100644 --- a/backend/CSE2proof.v +++ b/backend/CSE2proof.v @@ -202,7 +202,136 @@ Proof. destruct peq; trivial. congruence. Qed. +Local Hint Resolve kill_sym_val_referenced : wellformed. +Lemma in_or_not: + forall (x : reg) l, + (In x l) \/ ~(In x l). +Admitted. + +Lemma remove_from_sets_notin: + forall dst l sets j, + not (In j l) -> + (remove_from_sets dst l sets) ! j = sets ! j. +Proof. + induction l; simpl; trivial. + intros. + rewrite IHl by tauto. + destruct (@PTree.get (PTree.t unit) a sets) eqn:SETSA; trivial. + pose proof (PTree.bempty_canon_correct (PTree.remove dst t)) as CORRECT. + destruct (PTree.bempty_canon (PTree.remove dst t)). + - apply PTree.gro. + intuition. + - rewrite PTree.gso by intuition. + trivial. +Qed. + +Lemma remove_from_sets_pass: + forall dst l sets i j, + i <> dst -> + match sets ! j with + | Some uses => uses ! i = Some tt + | None => False + end -> + match (remove_from_sets dst l sets) ! j with + | Some uses => uses ! i = Some tt + | None => False + end. +Proof. + induction l; simpl; trivial. + intros. + apply IHl; trivial. + destruct (@PTree.get (PTree.t unit) a sets) eqn:SETSA; trivial. + pose proof (PTree.bempty_canon_correct (PTree.remove dst t)) as CORRECT. + specialize CORRECT with (i := i). + destruct (PTree.bempty_canon (PTree.remove dst t)). + - rewrite PTree.gro in CORRECT by congruence. + destruct (peq a j). + + subst a. + rewrite SETSA in *. + intuition congruence. + + rewrite PTree.gro by congruence. + assumption. + - destruct (peq a j). + + subst a. + rewrite SETSA in *. + rewrite PTree.gss. + rewrite PTree.gro by congruence. + assumption. + + rewrite PTree.gso by congruence. + assumption. +Qed. +Local Hint Resolve remove_from_sets_pass : wellformed. + +Lemma rem_comes_from: + forall A x y (tr: PTree.t A) v, + (PTree.remove x tr) ! y = Some v -> tr ! y = Some v. +Proof. + intros. + destruct (peq x y). + - subst y. + rewrite PTree.grs in *. + discriminate. + - rewrite PTree.gro in * by congruence. + assumption. +Qed. +Local Hint Resolve rem_comes_from : wellformed. + +Lemma rem_ineq: + forall A x y (tr: PTree.t A) v, + (PTree.remove x tr) ! y = Some v -> x<>y. +Proof. + intros. + intro. + subst y. + rewrite PTree.grs in *. + discriminate. +Qed. +Local Hint Resolve rem_ineq : wellformed. + +Lemma wellformed_reg_kill_reg: + forall dst rel, + (wellformed_reg rel) -> (wellformed_reg (kill_reg dst rel)). +Proof. + unfold wellformed_reg, kill_reg. + simpl. + intros dst rel Hrel. + intros i j sv KILLREG KILLABLE. + + rewrite PTree.gfilter1 in KILLREG. + destruct PTree.get eqn:REMi in KILLREG. + 2: discriminate. + destruct negb eqn:NEGB in KILLREG. + 2: discriminate. + inv KILLREG. + + assert ((var_to_sym rel) ! i = Some sv) as RELi by eauto with wellformed. + + destruct (peq dst j). + { (* absurd case *) + subst j. + rewrite KILLABLE in NEGB. + simpl in NEGB. + discriminate. + } + specialize Hrel with (i := i) (j := j) (sv := sv). + destruct ((var_to_sym rel) ! dst) eqn:DST; simpl. + { + destruct (in_or_not j (referenced_by s)). + { assert (dst <> i) by eauto with wellformed. + apply remove_from_sets_pass. + congruence. + rewrite PTree.gro by congruence. + apply Hrel; trivial. + } + rewrite remove_from_sets_notin by assumption. + rewrite PTree.gro by congruence. + apply Hrel; trivial. + } + rewrite PTree.gro by congruence. + apply Hrel; trivial. +Qed. + Lemma wellformed_mem_kill_mem: forall rel, (wellformed_mem rel) -> (wellformed_mem (kill_mem rel)). -- cgit From bc89c7487901c6056c84bf598ea4ab6535de68c2 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 5 Feb 2020 17:27:57 +0100 Subject: wellformed_reg_kill_reg simpliied --- backend/CSE2proof.v | 11 +++-------- 1 file changed, 3 insertions(+), 8 deletions(-) diff --git a/backend/CSE2proof.v b/backend/CSE2proof.v index e62e2ae6..3799e38b 100644 --- a/backend/CSE2proof.v +++ b/backend/CSE2proof.v @@ -317,14 +317,9 @@ Proof. specialize Hrel with (i := i) (j := j) (sv := sv). destruct ((var_to_sym rel) ! dst) eqn:DST; simpl. { - destruct (in_or_not j (referenced_by s)). - { assert (dst <> i) by eauto with wellformed. - apply remove_from_sets_pass. - congruence. - rewrite PTree.gro by congruence. - apply Hrel; trivial. - } - rewrite remove_from_sets_notin by assumption. + assert (dst <> i) by eauto with wellformed. + apply remove_from_sets_pass. + congruence. rewrite PTree.gro by congruence. apply Hrel; trivial. } -- cgit From 27b985393cd8d90a3d6f5e9f13bdf90e4300bb8e Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 5 Feb 2020 17:28:21 +0100 Subject: rm useless admitted lemma --- backend/CSE2proof.v | 5 ----- 1 file changed, 5 deletions(-) diff --git a/backend/CSE2proof.v b/backend/CSE2proof.v index 3799e38b..4d503f90 100644 --- a/backend/CSE2proof.v +++ b/backend/CSE2proof.v @@ -204,11 +204,6 @@ Proof. Qed. Local Hint Resolve kill_sym_val_referenced : wellformed. -Lemma in_or_not: - forall (x : reg) l, - (In x l) \/ ~(In x l). -Admitted. - Lemma remove_from_sets_notin: forall dst l sets j, not (In j l) -> -- cgit From e4c7ff447ad30832c42ee51f0fe7e1cf62d2eee9 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 5 Feb 2020 17:34:03 +0100 Subject: wellformed_reg_kill_mem --- backend/CSE2proof.v | 18 +++++++++++++++++- 1 file changed, 17 insertions(+), 1 deletion(-) diff --git a/backend/CSE2proof.v b/backend/CSE2proof.v index 4d503f90..0d369e7a 100644 --- a/backend/CSE2proof.v +++ b/backend/CSE2proof.v @@ -321,7 +321,8 @@ Proof. rewrite PTree.gro by congruence. apply Hrel; trivial. Qed. - +Local Hint Resolve wellformed_reg_kill_reg : wellformed. + Lemma wellformed_mem_kill_mem: forall rel, (wellformed_mem rel) -> (wellformed_mem (kill_mem rel)). @@ -343,6 +344,21 @@ Proof. Qed. Local Hint Resolve wellformed_mem_kill_mem : wellformed. +Lemma wellformed_reg_kill_mem: + forall rel, + (wellformed_reg rel) -> (wellformed_reg (kill_mem rel)). +Proof. + unfold wellformed_reg, kill_mem. + simpl. + intros rel Hrel. + intros i j sv KILLMEM KILLABLE. + apply Hrel with (sv := sv); trivial. + rewrite PTree.gremove_t in KILLMEM. + destruct ((mem_used rel) ! i) in KILLMEM. + discriminate. + assumption. +Qed. + Lemma wellformed_mem_move: forall r dst rel, (wellformed_mem rel) -> (wellformed_mem (move r dst rel)). -- cgit From e35365927d1289687aaff6d7ca5ebee1ac09d249 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 5 Feb 2020 17:40:02 +0100 Subject: add hint --- backend/CSE2proof.v | 1 + 1 file changed, 1 insertion(+) diff --git a/backend/CSE2proof.v b/backend/CSE2proof.v index 0d369e7a..0b92f5e5 100644 --- a/backend/CSE2proof.v +++ b/backend/CSE2proof.v @@ -358,6 +358,7 @@ Proof. discriminate. assumption. Qed. +Local Hint Resolve wellformed_reg_kill_mem : wellformed. Lemma wellformed_mem_move: forall r dst rel, -- cgit From 305d3d307fd1a83b17052119d75516946ce617b4 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 19 Feb 2020 17:28:43 +0100 Subject: First part of Hansen algorithm - build the sequences --- backend/Linearizeaux.ml | 44 +++++++++++++++++++++++++++++++++++++++++++- 1 file changed, 43 insertions(+), 1 deletion(-) diff --git a/backend/Linearizeaux.ml b/backend/Linearizeaux.ml index a6964233..28719207 100644 --- a/backend/Linearizeaux.ml +++ b/backend/Linearizeaux.ml @@ -122,7 +122,11 @@ let enumerate_aux_flat f reach = * rather than a branch (ifso). * * The enumeration below takes advantage of this - preferring to layout nodes - * following the fallthroughs of the Lcond branches + * following the fallthroughs of the Lcond branches. + * + * It is slightly adapted from the work of Petris and Hansen 90 on intraprocedural + * code positioning - only we do it on a broader grain, since we don't have the exact + * frequencies (we only know which branch is the preferred one) *) let get_some = function @@ -136,6 +140,7 @@ let rec last_element = function | e :: [] -> e | e' :: e :: l -> last_element (e::l) +(** old version let dfs code entrypoint = let visited = ref (PTree.map (fun n i -> false) code) in let rec dfs_list code = function @@ -159,6 +164,43 @@ let dfs code entrypoint = in dfs_list code [entrypoint] let enumerate_aux_trace f reach = dfs f.fn_code f.fn_entrypoint +*) + + +let forward_sequences code entry = + let visited = ref (PTree.map (fun n i -> false) code) in + (* returns the list of traversed nodes, and a list of nodes to start traversing next *) + let rec traverse_fallthrough code node = + if not (get_some @@ PTree.get node !visited) then begin + visited := PTree.set node true !visited; + match PTree.get node code with + | None -> failwith "No such node" + | Some bb -> + let ln, rem = match (last_element bb) with + | Lop _ | Lload _ | Lgetstack _ | Lsetstack _ | Lstore _ | Lcall _ + | Lbuiltin _ -> assert false + | Ltailcall _ | Lreturn -> ([], []) + | Lbranch n -> let ln, rem = traverse_fallthrough code n in (ln, rem) + | Lcond (_, _, ifso, ifnot) -> let ln, rem = traverse_fallthrough code ifnot in (ln, [ifso] @ rem) + | Ljumptable(_, ln) -> match ln with + | [] -> ([], []) + | n :: ln -> let lln, rem = traverse_fallthrough code n in (lln, ln @ rem) + in ([node] @ ln, rem) + end + else ([], []) + in let rec f code = function + | [] -> [] + | node :: ln -> + let fs, rem = traverse_fallthrough code node + in [fs] @ (f code rem) + in (f code [entry]) + +let order_sequences fs = fs + +let enumerate_aux_trace f reach = + let fs = forward_sequences f.fn_code f.fn_entrypoint + in let ofs = order_sequences fs + in List.flatten ofs let enumerate_aux f reach = if !Clflags.option_ftracelinearize then enumerate_aux_trace f reach -- cgit From 840f7e5af8294d121caf1c712b75ba3c13914a54 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Thu, 20 Feb 2020 11:25:58 +0100 Subject: WIP --- backend/Linearizeaux.ml | 26 ++++++++++++++++++++++++++ 1 file changed, 26 insertions(+) diff --git a/backend/Linearizeaux.ml b/backend/Linearizeaux.ml index 28719207..8dccbb4b 100644 --- a/backend/Linearizeaux.ml +++ b/backend/Linearizeaux.ml @@ -195,6 +195,32 @@ let forward_sequences code entry = in [fs] @ (f code rem) in (f code [entry]) +let iter_set f s = Seq.iter f (Set.to_seq s) + +let try_merge code fs = + let seqs = ref (Set.of_list fs) in + let oldLength = ref (Set.cardinal !seqs) in + let continue = ref true in + while !continue do + iter_set (fun s -> + iter_set (fun s' -> + if (s == s') then () + else if (can_be_merged s s') then + begin + seqs + end + else () + ) !seqs + ) !seqs + (* FIXME - FIXME - continue *) + + + if !oldLength == List.length !seqs then + continue := false + else + oldLength := List.length !seqs + done + let order_sequences fs = fs let enumerate_aux_trace f reach = -- cgit From 87d8a302892026bc831fd5c310874f2aad6be194 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Thu, 20 Feb 2020 11:50:51 +0100 Subject: WIP2 --- backend/Linearizeaux.ml | 40 ++++++++++++++++++++++++---------------- 1 file changed, 24 insertions(+), 16 deletions(-) diff --git a/backend/Linearizeaux.ml b/backend/Linearizeaux.ml index 8dccbb4b..d7150989 100644 --- a/backend/Linearizeaux.ml +++ b/backend/Linearizeaux.ml @@ -201,25 +201,33 @@ let try_merge code fs = let seqs = ref (Set.of_list fs) in let oldLength = ref (Set.cardinal !seqs) in let continue = ref true in + let found = ref false in while !continue do - iter_set (fun s -> - iter_set (fun s' -> - if (s == s') then () - else if (can_be_merged s s') then - begin - seqs - end - else () - ) !seqs - ) !seqs - (* FIXME - FIXME - continue *) + begin + found := false; + iter_set (fun s -> + if !found then () + else iter_set (fun s' -> + if (!found || s == s') then () + else if (can_be_merged s s') then + begin + seqs := Set.remove s !seqs; + seqs := Set.remove s' !seqs; + seqs := Set.add (get_some (merge s s')) !seqs; + found := true; + end + else () + ) !seqs + ) !seqs; + if !oldLength == List.length !seqs then + continue := false + else + oldLength := List.length !seqs + end + done; + (* FIXME - continue *) - if !oldLength == List.length !seqs then - continue := false - else - oldLength := List.length !seqs - done let order_sequences fs = fs -- cgit From 0271028f40c58068975170476dcaa5aadc21cb7e Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Fri, 21 Feb 2020 11:44:43 +0100 Subject: Linearizeaux: function try_merge --- backend/Linearizeaux.ml | 31 ++++++++++++++++++++----------- 1 file changed, 20 insertions(+), 11 deletions(-) diff --git a/backend/Linearizeaux.ml b/backend/Linearizeaux.ml index d7150989..44322a46 100644 --- a/backend/Linearizeaux.ml +++ b/backend/Linearizeaux.ml @@ -195,11 +195,22 @@ let forward_sequences code entry = in [fs] @ (f code rem) in (f code [entry]) -let iter_set f s = Seq.iter f (Set.to_seq s) +module PInt = struct + type t = P.t + let compare x y = compare (P.to_int x) (P.to_int y) +end + +module PSet = Set.Make(PInt) + +let iter_set f s = Seq.iter f (PSet.to_seq s) + +let can_be_merged s s' = false + +let merge s s' = Some s let try_merge code fs = - let seqs = ref (Set.of_list fs) in - let oldLength = ref (Set.cardinal !seqs) in + let seqs = ref (PSet.of_list fs) in + let oldLength = ref (PSet.cardinal !seqs) in let continue = ref true in let found = ref false in while !continue do @@ -211,23 +222,21 @@ let try_merge code fs = if (!found || s == s') then () else if (can_be_merged s s') then begin - seqs := Set.remove s !seqs; - seqs := Set.remove s' !seqs; - seqs := Set.add (get_some (merge s s')) !seqs; + seqs := PSet.remove s !seqs; + seqs := PSet.remove s' !seqs; + seqs := PSet.add (get_some (merge s s')) !seqs; found := true; end else () ) !seqs ) !seqs; - if !oldLength == List.length !seqs then + if !oldLength == PSet.cardinal !seqs then continue := false else - oldLength := List.length !seqs + oldLength := PSet.cardinal !seqs end done; - (* FIXME - continue *) - - + !seqs let order_sequences fs = fs -- cgit From 49077ae5aa2f88c843b8fae8cd60aff75a52e5e8 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Fri, 21 Feb 2020 16:02:03 +0100 Subject: Linearizeaux: can_be_merged --- backend/Linearizeaux.ml | 61 ++++++++++++++++++++++++++++++++++++++----------- 1 file changed, 48 insertions(+), 13 deletions(-) diff --git a/backend/Linearizeaux.ml b/backend/Linearizeaux.ml index 44322a46..3ef86344 100644 --- a/backend/Linearizeaux.ml +++ b/backend/Linearizeaux.ml @@ -202,38 +202,73 @@ end module PSet = Set.Make(PInt) -let iter_set f s = Seq.iter f (PSet.to_seq s) +module LPInt = struct + type t = P.t list + let rec compare x y = + match x with + | [] -> ( match y with + | [] -> 0 + | _ -> 1 ) + | e :: l -> match y with + | [] -> -1 + | e' :: l' -> + let e_cmp = PInt.compare e e' in + if e_cmp == 0 then compare l l' else e_cmp +end + +module LPSet = Set.Make(LPInt) + +let iter_lpset f s = Seq.iter f (LPSet.to_seq s) + +let first_of = function + | [] -> None + | e :: l -> Some e + +let rec last_of = function + | [] -> None + | e :: l -> (match l with [] -> Some e | e :: l -> last_of l) -let can_be_merged s s' = false +let can_be_merged code s s' = + let last_s = get_some @@ last_of s in + let first_s' = get_some @@ first_of s' in + match get_some @@ PTree.get last_s code with + | Lop _ | Lload _ | Lgetstack _ | Lsetstack _ | Lstore _ | Lcall _ + | Lbuiltin _ | Ltailcall _ | Lreturn -> false + | Lbranch n -> n == first_s' + | Lcond (_, _, ifso, ifnot) -> ifnot == first_s' + | Ljumptable (_, ln) -> + match ln with + | [] -> false + | n :: ln -> n == first_s' let merge s s' = Some s -let try_merge code fs = - let seqs = ref (PSet.of_list fs) in - let oldLength = ref (PSet.cardinal !seqs) in +let try_merge code (fs: (BinNums.positive list) list) = + let seqs = ref (LPSet.of_list fs) in + let oldLength = ref (LPSet.cardinal !seqs) in let continue = ref true in let found = ref false in while !continue do begin found := false; - iter_set (fun s -> + iter_lpset (fun s -> if !found then () - else iter_set (fun s' -> + else iter_lpset (fun s' -> if (!found || s == s') then () - else if (can_be_merged s s') then + else if (can_be_merged code s s') then begin - seqs := PSet.remove s !seqs; - seqs := PSet.remove s' !seqs; - seqs := PSet.add (get_some (merge s s')) !seqs; + seqs := LPSet.remove s !seqs; + seqs := LPSet.remove s' !seqs; + seqs := LPSet.add (get_some (merge s s')) !seqs; found := true; end else () ) !seqs ) !seqs; - if !oldLength == PSet.cardinal !seqs then + if !oldLength == LPSet.cardinal !seqs then continue := false else - oldLength := PSet.cardinal !seqs + oldLength := LPSet.cardinal !seqs end done; !seqs -- cgit From 08efc2a09b850476e39469791650faf99dd06183 Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Mon, 24 Feb 2020 13:56:07 +0100 Subject: Platform-independent implementation of Conventions.size_arguments (#222) The "size_arguments" function and its properties can be systematically derived from the "loc_arguments" function and its properties. Before, the RISC-V port used this derivation, and all other ports used hand-written "size_arguments" functions and proofs. This commit moves the definition of "size_arguments" to the platform-independent file backend/Conventions.v, using the systematic derivation, and removes the platform-specific definitions. This reduces code and proof size, and makes it easier to change the calling conventions. --- aarch64/Conventions1.v | 107 ------------------------- arm/Conventions1.v | 206 ------------------------------------------------- backend/Conventions.v | 67 ++++++++++++++++ powerpc/Conventions1.v | 126 ------------------------------ riscV/Conventions1.v | 64 --------------- x86/Asmexpand.ml | 2 +- x86/Conventions1.v | 145 ---------------------------------- 7 files changed, 68 insertions(+), 649 deletions(-) diff --git a/aarch64/Conventions1.v b/aarch64/Conventions1.v index 14cb199f..efda835d 100644 --- a/aarch64/Conventions1.v +++ b/aarch64/Conventions1.v @@ -190,27 +190,6 @@ Fixpoint loc_arguments_rec Definition loc_arguments (s: signature) : list (rpair loc) := loc_arguments_rec s.(sig_args) 0 0 0. -(** [size_arguments s] returns the number of [Outgoing] slots used - to call a function with signature [s]. *) - -Fixpoint size_arguments_rec (tyl: list typ) (ir fr ofs: Z) {struct tyl} : Z := - match tyl with - | nil => ofs - | (Tint | Tlong | Tany32 | Tany64) :: tys => - match list_nth_z int_param_regs ir with - | None => size_arguments_rec tys ir fr (ofs + 2) - | Some ireg => size_arguments_rec tys (ir + 1) fr ofs - end - | (Tfloat | Tsingle) :: tys => - match list_nth_z float_param_regs fr with - | None => size_arguments_rec tys ir fr (ofs + 2) - | Some freg => size_arguments_rec tys ir (fr + 1) ofs - end - end. - -Definition size_arguments (s: signature) : Z := - size_arguments_rec s.(sig_args) 0 0 0. - (** Argument locations are either caller-save registers or [Outgoing] stack slots at nonnegative offsets. *) @@ -285,92 +264,6 @@ Qed. Hint Resolve loc_arguments_acceptable: locs. -(** The offsets of [Outgoing] arguments are below [size_arguments s]. *) - -Remark size_arguments_rec_above: - forall tyl ir fr ofs0, - ofs0 <= size_arguments_rec tyl ir fr ofs0. -Proof. - induction tyl; simpl; intros. - omega. - assert (A: ofs0 <= - match list_nth_z int_param_regs ir with - | Some _ => size_arguments_rec tyl (ir + 1) fr ofs0 - | None => size_arguments_rec tyl ir fr (ofs0 + 2) - end). - { destruct (list_nth_z int_param_regs ir); eauto. - apply Z.le_trans with (ofs0 + 2); auto. omega. } - assert (B: ofs0 <= - match list_nth_z float_param_regs fr with - | Some _ => size_arguments_rec tyl ir (fr + 1) ofs0 - | None => size_arguments_rec tyl ir fr (ofs0 + 2) - end). - { destruct (list_nth_z float_param_regs fr); eauto. - apply Z.le_trans with (ofs0 + 2); auto. omega. } - destruct a; auto. -Qed. - -Lemma size_arguments_above: - forall s, size_arguments s >= 0. -Proof. - intros; unfold size_arguments. apply Z.le_ge. apply size_arguments_rec_above. -Qed. - -Lemma loc_arguments_rec_bounded: - forall ofs ty tyl ir fr ofs0, - In (S Outgoing ofs ty) (regs_of_rpairs (loc_arguments_rec tyl ir fr ofs0)) -> - ofs + typesize ty <= size_arguments_rec tyl ir fr ofs0. -Proof. - induction tyl; simpl; intros. -- contradiction. -- assert (T: forall ty0, typesize ty0 <= 2). - { destruct ty0; simpl; omega. } - assert (A: forall ty0, - In (S Outgoing ofs ty) (regs_of_rpairs - match list_nth_z int_param_regs ir with - | Some ireg => - One (R ireg) :: loc_arguments_rec tyl (ir + 1) fr ofs0 - | None => One (S Outgoing ofs0 ty0) :: loc_arguments_rec tyl ir fr (ofs0 + 2) - end) -> - ofs + typesize ty <= - match list_nth_z int_param_regs ir with - | Some _ => size_arguments_rec tyl (ir + 1) fr ofs0 - | None => size_arguments_rec tyl ir fr (ofs0 + 2) - end). - { intros. destruct (list_nth_z int_param_regs ir); simpl in H0; destruct H0. - - discriminate. - - eapply IHtyl; eauto. - - inv H0. apply Z.le_trans with (ofs + 2). specialize (T ty). omega. apply size_arguments_rec_above. - - eapply IHtyl; eauto. } - assert (B: forall ty0, - In (S Outgoing ofs ty) (regs_of_rpairs - match list_nth_z float_param_regs fr with - | Some ireg => - One (R ireg) :: loc_arguments_rec tyl ir (fr + 1) ofs0 - | None => One (S Outgoing ofs0 ty0) :: loc_arguments_rec tyl ir fr (ofs0 + 2) - end) -> - ofs + typesize ty <= - match list_nth_z float_param_regs fr with - | Some _ => size_arguments_rec tyl ir (fr + 1) ofs0 - | None => size_arguments_rec tyl ir fr (ofs0 + 2) - end). - { intros. destruct (list_nth_z float_param_regs fr); simpl in H0; destruct H0. - - discriminate. - - eapply IHtyl; eauto. - - inv H0. apply Z.le_trans with (ofs + 2). specialize (T ty). omega. apply size_arguments_rec_above. - - eapply IHtyl; eauto. } - destruct a; eauto. -Qed. - -Lemma loc_arguments_bounded: - forall (s: signature) (ofs: Z) (ty: typ), - In (S Outgoing ofs ty) (regs_of_rpairs (loc_arguments s)) -> - ofs + typesize ty <= size_arguments s. -Proof. - unfold loc_arguments, size_arguments; intros. - eauto using loc_arguments_rec_bounded. -Qed. - Lemma loc_arguments_main: loc_arguments signature_main = nil. Proof. diff --git a/arm/Conventions1.v b/arm/Conventions1.v index 7016c1ee..fe49a781 100644 --- a/arm/Conventions1.v +++ b/arm/Conventions1.v @@ -269,48 +269,6 @@ Definition loc_arguments (s: signature) : list (rpair loc) := else loc_arguments_hf s.(sig_args) 0 0 0 end. -(** [size_arguments s] returns the number of [Outgoing] slots used - to call a function with signature [s]. *) - -Fixpoint size_arguments_hf (tyl: list typ) (ir fr ofs: Z) {struct tyl} : Z := - match tyl with - | nil => ofs - | (Tint|Tany32) :: tys => - if zlt ir 4 - then size_arguments_hf tys (ir + 1) fr ofs - else size_arguments_hf tys ir fr (ofs + 1) - | (Tfloat|Tany64) :: tys => - if zlt fr 8 - then size_arguments_hf tys ir (fr + 1) ofs - else size_arguments_hf tys ir fr (align ofs 2 + 2) - | Tsingle :: tys => - if zlt fr 8 - then size_arguments_hf tys ir (fr + 1) ofs - else size_arguments_hf tys ir fr (ofs + 1) - | Tlong :: tys => - let ir := align ir 2 in - if zlt ir 4 - then size_arguments_hf tys (ir + 2) fr ofs - else size_arguments_hf tys ir fr (align ofs 2 + 2) - end. - -Fixpoint size_arguments_sf (tyl: list typ) (ofs: Z) {struct tyl} : Z := - match tyl with - | nil => Z.max 0 ofs - | (Tint | Tsingle | Tany32) :: tys => size_arguments_sf tys (ofs + 1) - | (Tfloat | Tlong | Tany64) :: tys => size_arguments_sf tys (align ofs 2 + 2) - end. - -Definition size_arguments (s: signature) : Z := - match Archi.abi with - | Archi.Softfloat => - size_arguments_sf s.(sig_args) (-4) - | Archi.Hardfloat => - if s.(sig_cc).(cc_vararg) - then size_arguments_sf s.(sig_args) (-4) - else size_arguments_hf s.(sig_args) 0 0 0 - end. - (** Argument locations are either non-temporary registers or [Outgoing] stack slots at nonnegative offsets. *) @@ -471,170 +429,6 @@ Qed. Hint Resolve loc_arguments_acceptable: locs. -(** The offsets of [Outgoing] arguments are below [size_arguments s]. *) - -Remark size_arguments_hf_above: - forall tyl ir fr ofs0, - ofs0 <= size_arguments_hf tyl ir fr ofs0. -Proof. - induction tyl; simpl; intros. - omega. - destruct a. - destruct (zlt ir 4); eauto. apply Z.le_trans with (ofs0 + 1); auto; omega. - destruct (zlt fr 8); eauto. - apply Z.le_trans with (align ofs0 2). apply align_le; omega. - apply Z.le_trans with (align ofs0 2 + 2); auto; omega. - set (ir' := align ir 2). - destruct (zlt ir' 4); eauto. - apply Z.le_trans with (align ofs0 2). apply align_le; omega. - apply Z.le_trans with (align ofs0 2 + 2); auto; omega. - destruct (zlt fr 8); eauto. - apply Z.le_trans with (ofs0 + 1); eauto. omega. - destruct (zlt ir 4); eauto. apply Z.le_trans with (ofs0 + 1); auto; omega. - destruct (zlt fr 8); eauto. - apply Z.le_trans with (align ofs0 2). apply align_le; omega. - apply Z.le_trans with (align ofs0 2 + 2); auto; omega. -Qed. - -Remark size_arguments_sf_above: - forall tyl ofs0, - Z.max 0 ofs0 <= size_arguments_sf tyl ofs0. -Proof. - induction tyl; simpl; intros. - omega. - destruct a; (eapply Z.le_trans; [idtac|eauto]). - xomega. - assert (ofs0 <= align ofs0 2) by (apply align_le; omega). xomega. - assert (ofs0 <= align ofs0 2) by (apply align_le; omega). xomega. - xomega. - xomega. - assert (ofs0 <= align ofs0 2) by (apply align_le; omega). xomega. -Qed. - -Lemma size_arguments_above: - forall s, size_arguments s >= 0. -Proof. - intros; unfold size_arguments. apply Z.le_ge. - assert (0 <= size_arguments_sf (sig_args s) (-4)). - { change 0 with (Z.max 0 (-4)). apply size_arguments_sf_above. } - assert (0 <= size_arguments_hf (sig_args s) 0 0 0). - { apply size_arguments_hf_above. } - destruct Archi.abi; [ | destruct (cc_vararg (sig_cc s)) ]; auto. -Qed. - -Lemma loc_arguments_hf_bounded: - forall ofs ty tyl ir fr ofs0, - In (S Outgoing ofs ty) (regs_of_rpairs (loc_arguments_hf tyl ir fr ofs0)) -> - ofs + typesize ty <= size_arguments_hf tyl ir fr ofs0. -Proof. - induction tyl; simpl; intros. - elim H. - destruct a. -- (* int *) - destruct (zlt ir 4); destruct H. - discriminate. - eauto. - inv H. apply size_arguments_hf_above. - eauto. -- (* float *) - destruct (zlt fr 8); destruct H. - discriminate. - eauto. - inv H. apply size_arguments_hf_above. - eauto. -- (* long *) - destruct (zlt (align ir 2) 4). - destruct H. discriminate. destruct H. discriminate. eauto. - destruct Archi.big_endian. - destruct H. inv H. - eapply Z.le_trans. 2: apply size_arguments_hf_above. simpl; omega. - destruct H. inv H. - rewrite <- Z.add_assoc. simpl. apply size_arguments_hf_above. - eauto. - destruct H. inv H. - rewrite <- Z.add_assoc. simpl. apply size_arguments_hf_above. - destruct H. inv H. - eapply Z.le_trans. 2: apply size_arguments_hf_above. simpl; omega. - eauto. -- (* float *) - destruct (zlt fr 8); destruct H. - discriminate. - eauto. - inv H. apply size_arguments_hf_above. - eauto. -- (* any32 *) - destruct (zlt ir 4); destruct H. - discriminate. - eauto. - inv H. apply size_arguments_hf_above. - eauto. -- (* any64 *) - destruct (zlt fr 8); destruct H. - discriminate. - eauto. - inv H. apply size_arguments_hf_above. - eauto. -Qed. - -Lemma loc_arguments_sf_bounded: - forall ofs ty tyl ofs0, - In (S Outgoing ofs ty) (regs_of_rpairs (loc_arguments_sf tyl ofs0)) -> - Z.max 0 (ofs + typesize ty) <= size_arguments_sf tyl ofs0. -Proof. - induction tyl; simpl; intros. - elim H. - destruct a. -- (* int *) - destruct H. - destruct (zlt ofs0 0); inv H. apply size_arguments_sf_above. - eauto. -- (* float *) - destruct H. - destruct (zlt (align ofs0 2) 0); inv H. apply size_arguments_sf_above. - eauto. -- (* long *) - destruct H. - destruct Archi.big_endian. - destruct (zlt (align ofs0 2) 0); inv H. - eapply Z.le_trans. 2: apply size_arguments_sf_above. simpl; xomega. - destruct (zlt (align ofs0 2) 0); inv H. - rewrite <- Z.add_assoc. simpl. apply size_arguments_sf_above. - destruct H. - destruct Archi.big_endian. - destruct (zlt (align ofs0 2) 0); inv H. - rewrite <- Z.add_assoc. simpl. apply size_arguments_sf_above. - destruct (zlt (align ofs0 2) 0); inv H. - eapply Z.le_trans. 2: apply size_arguments_sf_above. simpl; xomega. - eauto. -- (* float *) - destruct H. - destruct (zlt ofs0 0); inv H. apply size_arguments_sf_above. - eauto. -- (* any32 *) - destruct H. - destruct (zlt ofs0 0); inv H. apply size_arguments_sf_above. - eauto. -- (* any64 *) - destruct H. - destruct (zlt (align ofs0 2) 0); inv H. apply size_arguments_sf_above. - eauto. -Qed. - -Lemma loc_arguments_bounded: - forall (s: signature) (ofs: Z) (ty: typ), - In (S Outgoing ofs ty) (regs_of_rpairs (loc_arguments s)) -> - ofs + typesize ty <= size_arguments s. -Proof. - unfold loc_arguments, size_arguments; intros. - assert (In (S Outgoing ofs ty) (regs_of_rpairs (loc_arguments_sf (sig_args s) (-4))) -> - ofs + typesize ty <= size_arguments_sf (sig_args s) (-4)). - { intros. eapply Z.le_trans. 2: eapply loc_arguments_sf_bounded; eauto. xomega. } - assert (In (S Outgoing ofs ty) (regs_of_rpairs (loc_arguments_hf (sig_args s) 0 0 0)) -> - ofs + typesize ty <= size_arguments_hf (sig_args s) 0 0 0). - { intros. eapply loc_arguments_hf_bounded; eauto. } - destruct Archi.abi; [ | destruct (cc_vararg (sig_cc s)) ]; eauto. -Qed. - Lemma loc_arguments_main: loc_arguments signature_main = nil. Proof. diff --git a/backend/Conventions.v b/backend/Conventions.v index 6025c6b4..14ffb587 100644 --- a/backend/Conventions.v +++ b/backend/Conventions.v @@ -34,6 +34,73 @@ Proof. apply IHpl; auto. Qed. +(** ** Stack size of function arguments *) + +(** [size_arguments s] returns the number of [Outgoing] slots used + to call a function with signature [s]. *) + +Definition max_outgoing_1 (accu: Z) (l: loc) : Z := + match l with + | S Outgoing ofs ty => Z.max accu (ofs + typesize ty) + | _ => accu + end. + +Definition max_outgoing_2 (accu: Z) (rl: rpair loc) : Z := + match rl with + | One l => max_outgoing_1 accu l + | Twolong l1 l2 => max_outgoing_1 (max_outgoing_1 accu l1) l2 + end. + +Definition size_arguments (s: signature) : Z := + List.fold_left max_outgoing_2 (loc_arguments s) 0. + +(** The offsets of [Outgoing] arguments are below [size_arguments s]. *) + +Remark fold_max_outgoing_above: + forall l n, fold_left max_outgoing_2 l n >= n. +Proof. + assert (A: forall n l, max_outgoing_1 n l >= n). + { intros; unfold max_outgoing_1. destruct l as [_ | []]; xomega. } + induction l; simpl; intros. + - omega. + - eapply Zge_trans. eauto. + destruct a; simpl. apply A. eapply Zge_trans; eauto. +Qed. + +Lemma size_arguments_above: + forall s, size_arguments s >= 0. +Proof. + intros. apply fold_max_outgoing_above. +Qed. + +Lemma loc_arguments_bounded: + forall (s: signature) (ofs: Z) (ty: typ), + In (S Outgoing ofs ty) (regs_of_rpairs (loc_arguments s)) -> + ofs + typesize ty <= size_arguments s. +Proof. + intros until ty. + assert (A: forall n l, n <= max_outgoing_1 n l). + { intros; unfold max_outgoing_1. destruct l as [_ | []]; xomega. } + assert (B: forall p n, + In (S Outgoing ofs ty) (regs_of_rpair p) -> + ofs + typesize ty <= max_outgoing_2 n p). + { intros. destruct p; simpl in H; intuition; subst; simpl. + - xomega. + - eapply Z.le_trans. 2: apply A. xomega. + - xomega. } + assert (C: forall l n, + In (S Outgoing ofs ty) (regs_of_rpairs l) -> + ofs + typesize ty <= fold_left max_outgoing_2 l n). + { induction l; simpl; intros. + - contradiction. + - rewrite in_app_iff in H. destruct H. + + eapply Z.le_trans. eapply B; eauto. + apply Z.ge_le. apply fold_max_outgoing_above. + + apply IHl; auto. + } + apply C. +Qed. + (** ** Location of function parameters *) (** A function finds the values of its parameter in the same locations diff --git a/powerpc/Conventions1.v b/powerpc/Conventions1.v index 25d9c081..1f048694 100644 --- a/powerpc/Conventions1.v +++ b/powerpc/Conventions1.v @@ -236,33 +236,6 @@ Fixpoint loc_arguments_rec Definition loc_arguments (s: signature) : list (rpair loc) := loc_arguments_rec s.(sig_args) 0 0 0. -(** [size_arguments s] returns the number of [Outgoing] slots used - to call a function with signature [s]. *) - -Fixpoint size_arguments_rec (tyl: list typ) (ir fr ofs: Z) {struct tyl} : Z := - match tyl with - | nil => ofs - | (Tint | Tany32) :: tys => - match list_nth_z int_param_regs ir with - | None => size_arguments_rec tys ir fr (ofs + 1) - | Some ireg => size_arguments_rec tys (ir + 1) fr ofs - end - | (Tfloat | Tsingle | Tany64) :: tys => - match list_nth_z float_param_regs fr with - | None => size_arguments_rec tys ir fr (align ofs 2 + 2) - | Some freg => size_arguments_rec tys ir (fr + 1) ofs - end - | Tlong :: tys => - let ir := align ir 2 in - match list_nth_z int_param_regs ir, list_nth_z int_param_regs (ir + 1) with - | Some r1, Some r2 => size_arguments_rec tys (ir + 2) fr ofs - | _, _ => size_arguments_rec tys ir fr (align ofs 2 + 2) - end - end. - -Definition size_arguments (s: signature) : Z := - size_arguments_rec s.(sig_args) 0 0 0. - (** Argument locations are either caller-save registers or [Outgoing] stack slots at nonnegative offsets. *) @@ -359,105 +332,6 @@ Qed. Hint Resolve loc_arguments_acceptable: locs. -(** The offsets of [Outgoing] arguments are below [size_arguments s]. *) - -Remark size_arguments_rec_above: - forall tyl ir fr ofs0, - ofs0 <= size_arguments_rec tyl ir fr ofs0. -Proof. - induction tyl; simpl; intros. - omega. - destruct a. - destruct (list_nth_z int_param_regs ir); eauto. apply Z.le_trans with (ofs0 + 1); auto; omega. - destruct (list_nth_z float_param_regs fr); eauto. - apply Z.le_trans with (align ofs0 2). apply align_le; omega. - apply Z.le_trans with (align ofs0 2 + 2); auto; omega. - set (ir' := align ir 2). - destruct (list_nth_z int_param_regs ir'); eauto. - destruct (list_nth_z int_param_regs (ir' + 1)); eauto. - apply Z.le_trans with (align ofs0 2). apply align_le; omega. - apply Z.le_trans with (align ofs0 2 + 2); auto; omega. - apply Z.le_trans with (align ofs0 2). apply align_le; omega. - apply Z.le_trans with (align ofs0 2 + 2); auto; omega. - destruct (list_nth_z float_param_regs fr); eauto. - apply Z.le_trans with (align ofs0 2). apply align_le; omega. - apply Z.le_trans with (align ofs0 2 + 2); auto; omega. - destruct (list_nth_z int_param_regs ir); eauto. apply Z.le_trans with (ofs0 + 1); auto; omega. - destruct (list_nth_z float_param_regs fr); eauto. - apply Z.le_trans with (align ofs0 2). apply align_le; omega. - apply Z.le_trans with (align ofs0 2 + 2); auto; omega. -Qed. - -Lemma size_arguments_above: - forall s, size_arguments s >= 0. -Proof. - intros; unfold size_arguments. apply Z.le_ge. - apply size_arguments_rec_above. -Qed. - -Lemma loc_arguments_bounded: - forall (s: signature) (ofs: Z) (ty: typ), - In (S Outgoing ofs ty) (regs_of_rpairs (loc_arguments s)) -> - ofs + typesize ty <= size_arguments s. -Proof. - intros. - assert (forall tyl ir fr ofs0, - In (S Outgoing ofs ty) (regs_of_rpairs (loc_arguments_rec tyl ir fr ofs0)) -> - ofs + typesize ty <= size_arguments_rec tyl ir fr ofs0). -{ - induction tyl; simpl; intros. - elim H0. - destruct a. -- (* int *) - destruct (list_nth_z int_param_regs ir); destruct H0. - congruence. - eauto. - inv H0. apply size_arguments_rec_above. - eauto. -- (* float *) - destruct (list_nth_z float_param_regs fr); destruct H0. - congruence. - eauto. - inv H0. apply size_arguments_rec_above. eauto. -- (* long *) - set (ir' := align ir 2) in *. - assert (DFL: - In (S Outgoing ofs ty) (regs_of_rpairs - ((if Archi.ptr64 - then One (S Outgoing (align ofs0 2) Tlong) - else Twolong (S Outgoing (align ofs0 2) Tint) - (S Outgoing (align ofs0 2 + 1) Tint)) - :: loc_arguments_rec tyl ir' fr (align ofs0 2 + 2))) -> - ofs + typesize ty <= size_arguments_rec tyl ir' fr (align ofs0 2 + 2)). - { destruct Archi.ptr64; intros IN. - - destruct IN. inv H1. apply size_arguments_rec_above. auto. - - destruct IN. inv H1. transitivity (align ofs0 2 + 2). simpl; omega. apply size_arguments_rec_above. - destruct H1. inv H1. transitivity (align ofs0 2 + 2). simpl; omega. apply size_arguments_rec_above. - auto. } - destruct (list_nth_z int_param_regs ir'); auto. - destruct (list_nth_z int_param_regs (ir' + 1)); auto. - destruct H0. congruence. destruct H0. congruence. eauto. -- (* single *) - destruct (list_nth_z float_param_regs fr); destruct H0. - congruence. - eauto. - inv H0. transitivity (align ofs0 2 + 2). simpl; omega. apply size_arguments_rec_above. - eauto. -- (* any32 *) - destruct (list_nth_z int_param_regs ir); destruct H0. - congruence. - eauto. - inv H0. apply size_arguments_rec_above. - eauto. -- (* any64 *) - destruct (list_nth_z float_param_regs fr); destruct H0. - congruence. - eauto. - inv H0. apply size_arguments_rec_above. eauto. - } - eauto. -Qed. - Lemma loc_arguments_main: loc_arguments signature_main = nil. Proof. diff --git a/riscV/Conventions1.v b/riscV/Conventions1.v index 27d09d94..a705c954 100644 --- a/riscV/Conventions1.v +++ b/riscV/Conventions1.v @@ -265,24 +265,6 @@ Fixpoint loc_arguments_rec (va: bool) Definition loc_arguments (s: signature) : list (rpair loc) := loc_arguments_rec s.(sig_cc).(cc_vararg) s.(sig_args) 0 0. -(** [size_arguments s] returns the number of [Outgoing] slots used - to call a function with signature [s]. *) - -Definition max_outgoing_1 (accu: Z) (l: loc) : Z := - match l with - | S Outgoing ofs ty => Z.max accu (ofs + typesize ty) - | _ => accu - end. - -Definition max_outgoing_2 (accu: Z) (rl: rpair loc) : Z := - match rl with - | One l => max_outgoing_1 accu l - | Twolong l1 l2 => max_outgoing_1 (max_outgoing_1 accu l1) l2 - end. - -Definition size_arguments (s: signature) : Z := - List.fold_left max_outgoing_2 (loc_arguments s) 0. - (** Argument locations are either non-temporary registers or [Outgoing] stack slots at nonnegative offsets. *) @@ -387,52 +369,6 @@ Proof. unfold loc_arguments; intros. eapply loc_arguments_rec_charact; eauto. omega. Qed. -(** The offsets of [Outgoing] arguments are below [size_arguments s]. *) - -Remark fold_max_outgoing_above: - forall l n, fold_left max_outgoing_2 l n >= n. -Proof. - assert (A: forall n l, max_outgoing_1 n l >= n). - { intros; unfold max_outgoing_1. destruct l as [_ | []]; xomega. } - induction l; simpl; intros. - - omega. - - eapply Zge_trans. eauto. - destruct a; simpl. apply A. eapply Zge_trans; eauto. -Qed. - -Lemma size_arguments_above: - forall s, size_arguments s >= 0. -Proof. - intros. apply fold_max_outgoing_above. -Qed. - -Lemma loc_arguments_bounded: - forall (s: signature) (ofs: Z) (ty: typ), - In (S Outgoing ofs ty) (regs_of_rpairs (loc_arguments s)) -> - ofs + typesize ty <= size_arguments s. -Proof. - intros until ty. - assert (A: forall n l, n <= max_outgoing_1 n l). - { intros; unfold max_outgoing_1. destruct l as [_ | []]; xomega. } - assert (B: forall p n, - In (S Outgoing ofs ty) (regs_of_rpair p) -> - ofs + typesize ty <= max_outgoing_2 n p). - { intros. destruct p; simpl in H; intuition; subst; simpl. - - xomega. - - eapply Z.le_trans. 2: apply A. xomega. - - xomega. } - assert (C: forall l n, - In (S Outgoing ofs ty) (regs_of_rpairs l) -> - ofs + typesize ty <= fold_left max_outgoing_2 l n). - { induction l; simpl; intros. - - contradiction. - - rewrite in_app_iff in H. destruct H. - + eapply Z.le_trans. eapply B; eauto. apply Z.ge_le. apply fold_max_outgoing_above. - + apply IHl; auto. - } - apply C. -Qed. - Lemma loc_arguments_main: loc_arguments signature_main = nil. Proof. diff --git a/x86/Asmexpand.ml b/x86/Asmexpand.ml index c82d406e..b8353046 100644 --- a/x86/Asmexpand.ml +++ b/x86/Asmexpand.ml @@ -251,7 +251,7 @@ let expand_builtin_va_start_32 r = invalid_arg "Fatal error: va_start used in non-vararg function"; let ofs = Int32.(add (add !PrintAsmaux.current_function_stacksize 4l) - (mul 4l (Z.to_int32 (Conventions1.size_arguments + (mul 4l (Z.to_int32 (Conventions.size_arguments (get_current_function_sig ()))))) in emit (Pleal (RAX, linear_addr RSP (Z.of_uint32 ofs))); emit (Pmovl_mr (linear_addr r _0z, RAX)) diff --git a/x86/Conventions1.v b/x86/Conventions1.v index 01b15e98..fdd94239 100644 --- a/x86/Conventions1.v +++ b/x86/Conventions1.v @@ -220,36 +220,6 @@ Definition loc_arguments (s: signature) : list (rpair loc) := then loc_arguments_64 s.(sig_args) 0 0 0 else loc_arguments_32 s.(sig_args) 0. -(** [size_arguments s] returns the number of [Outgoing] slots used - to call a function with signature [s]. *) - -Fixpoint size_arguments_32 - (tyl: list typ) (ofs: Z) {struct tyl} : Z := - match tyl with - | nil => ofs - | ty :: tys => size_arguments_32 tys (ofs + typesize ty) - end. - -Fixpoint size_arguments_64 (tyl: list typ) (ir fr ofs: Z) {struct tyl} : Z := - match tyl with - | nil => ofs - | (Tint | Tlong | Tany32 | Tany64) :: tys => - match list_nth_z int_param_regs ir with - | None => size_arguments_64 tys ir fr (ofs + 2) - | Some ireg => size_arguments_64 tys (ir + 1) fr ofs - end - | (Tfloat | Tsingle) :: tys => - match list_nth_z float_param_regs fr with - | None => size_arguments_64 tys ir fr (ofs + 2) - | Some freg => size_arguments_64 tys ir (fr + 1) ofs - end - end. - -Definition size_arguments (s: signature) : Z := - if Archi.ptr64 - then size_arguments_64 s.(sig_args) 0 0 0 - else size_arguments_32 s.(sig_args) 0. - (** Argument locations are either caller-save registers or [Outgoing] stack slots at nonnegative offsets. *) @@ -351,121 +321,6 @@ Qed. Hint Resolve loc_arguments_acceptable: locs. -(** The offsets of [Outgoing] arguments are below [size_arguments s]. *) - -Remark size_arguments_32_above: - forall tyl ofs0, ofs0 <= size_arguments_32 tyl ofs0. -Proof. - induction tyl; simpl; intros. - omega. - apply Z.le_trans with (ofs0 + typesize a); auto. - generalize (typesize_pos a); omega. -Qed. - -Remark size_arguments_64_above: - forall tyl ir fr ofs0, - ofs0 <= size_arguments_64 tyl ir fr ofs0. -Proof. - induction tyl; simpl; intros. - omega. - assert (A: ofs0 <= - match list_nth_z int_param_regs ir with - | Some _ => size_arguments_64 tyl (ir + 1) fr ofs0 - | None => size_arguments_64 tyl ir fr (ofs0 + 2) - end). - { destruct (list_nth_z int_param_regs ir); eauto. - apply Z.le_trans with (ofs0 + 2); auto. omega. } - assert (B: ofs0 <= - match list_nth_z float_param_regs fr with - | Some _ => size_arguments_64 tyl ir (fr + 1) ofs0 - | None => size_arguments_64 tyl ir fr (ofs0 + 2) - end). - { destruct (list_nth_z float_param_regs fr); eauto. - apply Z.le_trans with (ofs0 + 2); auto. omega. } - destruct a; auto. -Qed. - -Lemma size_arguments_above: - forall s, size_arguments s >= 0. -Proof. - intros; unfold size_arguments. apply Z.le_ge. - destruct Archi.ptr64; [apply size_arguments_64_above|apply size_arguments_32_above]. -Qed. - -Lemma loc_arguments_32_bounded: - forall ofs ty tyl ofs0, - In (S Outgoing ofs ty) (regs_of_rpairs (loc_arguments_32 tyl ofs0)) -> - ofs + typesize ty <= size_arguments_32 tyl ofs0. -Proof. - induction tyl as [ | t l]; simpl; intros x IN. -- contradiction. -- rewrite in_app_iff in IN; destruct IN as [IN|IN]. -+ apply Z.le_trans with (x + typesize t); [|apply size_arguments_32_above]. - Ltac decomp := - match goal with - | [ H: _ \/ _ |- _ ] => destruct H; decomp - | [ H: S _ _ _ = S _ _ _ |- _ ] => inv H - | [ H: False |- _ ] => contradiction - end. - destruct t; simpl in IN; decomp; simpl; omega. -+ apply IHl; auto. -Qed. - -Lemma loc_arguments_64_bounded: - forall ofs ty tyl ir fr ofs0, - In (S Outgoing ofs ty) (regs_of_rpairs (loc_arguments_64 tyl ir fr ofs0)) -> - ofs + typesize ty <= size_arguments_64 tyl ir fr ofs0. -Proof. - induction tyl; simpl; intros. - contradiction. - assert (T: forall ty0, typesize ty0 <= 2). - { destruct ty0; simpl; omega. } - assert (A: forall ty0, - In (S Outgoing ofs ty) (regs_of_rpairs - match list_nth_z int_param_regs ir with - | Some ireg => - One (R ireg) :: loc_arguments_64 tyl (ir + 1) fr ofs0 - | None => One (S Outgoing ofs0 ty0) :: loc_arguments_64 tyl ir fr (ofs0 + 2) - end) -> - ofs + typesize ty <= - match list_nth_z int_param_regs ir with - | Some _ => size_arguments_64 tyl (ir + 1) fr ofs0 - | None => size_arguments_64 tyl ir fr (ofs0 + 2) - end). - { intros. destruct (list_nth_z int_param_regs ir); simpl in H0; destruct H0. - - discriminate. - - eapply IHtyl; eauto. - - inv H0. apply Z.le_trans with (ofs + 2). specialize (T ty). omega. apply size_arguments_64_above. - - eapply IHtyl; eauto. } - assert (B: forall ty0, - In (S Outgoing ofs ty) (regs_of_rpairs - match list_nth_z float_param_regs fr with - | Some ireg => - One (R ireg) :: loc_arguments_64 tyl ir (fr + 1) ofs0 - | None => One (S Outgoing ofs0 ty0) :: loc_arguments_64 tyl ir fr (ofs0 + 2) - end) -> - ofs + typesize ty <= - match list_nth_z float_param_regs fr with - | Some _ => size_arguments_64 tyl ir (fr + 1) ofs0 - | None => size_arguments_64 tyl ir fr (ofs0 + 2) - end). - { intros. destruct (list_nth_z float_param_regs fr); simpl in H0; destruct H0. - - discriminate. - - eapply IHtyl; eauto. - - inv H0. apply Z.le_trans with (ofs + 2). specialize (T ty). omega. apply size_arguments_64_above. - - eapply IHtyl; eauto. } - destruct a; eauto. -Qed. - -Lemma loc_arguments_bounded: - forall (s: signature) (ofs: Z) (ty: typ), - In (S Outgoing ofs ty) (regs_of_rpairs (loc_arguments s)) -> - ofs + typesize ty <= size_arguments s. -Proof. - unfold loc_arguments, size_arguments; intros. - destruct Archi.ptr64; eauto using loc_arguments_32_bounded, loc_arguments_64_bounded. -Qed. - Lemma loc_arguments_main: loc_arguments signature_main = nil. Proof. -- cgit From c7adc93617712acdde0ea81649eff11ada7d96b9 Mon Sep 17 00:00:00 2001 From: Bernhard Schommer Date: Mon, 24 Feb 2020 20:22:45 +0100 Subject: The type of a wide char constant is wchar_t. (#223) See ISO C2011 standard, section 6.4.4.4 para 11. --- cparser/Elab.ml | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/cparser/Elab.ml b/cparser/Elab.ml index 9aa6761c..f60e15a3 100644 --- a/cparser/Elab.ml +++ b/cparser/Elab.ml @@ -452,7 +452,8 @@ let elab_constant loc = function let (v, fk) = elab_float_constant f in CFloat(v, fk) | CONST_CHAR(wide, s) -> - CInt(elab_char_constant loc wide s, IInt, "") + let ikind = if wide then wchar_ikind () else IInt in + CInt(elab_char_constant loc wide s, ikind, "") | CONST_STRING(wide, s) -> elab_string_literal loc wide s -- cgit From 5003b8d93c2a20821b776f7f74f5096a308a03cf Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Wed, 26 Feb 2020 16:49:50 +0100 Subject: Update the RISC-V calling conventions (#221) We were implementing the ABI described in the RISC-V Instruction Set Manual, version 2.1. However, this ABI was superseded by the RISC-V ELF psABI specification. This commit changes the calling conventions to better match the ELF psABI specification. This should greatly improve interoperability with code compiled by other RISC-V compilers. One incompatibility remains: when all 8 FP argument registers have been used, further FP arguments should be passed in integer argument registers if available, while this PR passes them on stack. --- riscV/Asmexpand.ml | 36 ++++---- riscV/Conventions1.v | 250 +++++++++++++++++++++++++++------------------------ 2 files changed, 149 insertions(+), 137 deletions(-) diff --git a/riscV/Asmexpand.ml b/riscV/Asmexpand.ml index d36b6230..7e36abf8 100644 --- a/riscV/Asmexpand.ml +++ b/riscV/Asmexpand.ml @@ -63,44 +63,44 @@ let expand_storeind_ptr src base ofs = 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 pos tyl = - if pos < 8 then +let rec fixup_variadic_call ri rf tyl = + if ri < 8 then match tyl with | [] -> () | (Tint | Tany32) :: tyl -> - fixup_variadic_call (pos + 1) tyl + fixup_variadic_call (ri + 1) rf tyl | Tsingle :: tyl -> - let rs =float_param_regs.(pos) - and rd = int_param_regs.(pos) in + let rs = float_param_regs.(rf) + and rd = int_param_regs.(ri) in emit (Pfmvxs(rd, rs)); - fixup_variadic_call (pos + 1) tyl + fixup_variadic_call (ri + 1) (rf + 1) tyl | Tlong :: tyl -> - let pos' = if Archi.ptr64 then pos + 1 else align pos 2 + 2 in - fixup_variadic_call pos' tyl + let ri' = if Archi.ptr64 then ri + 1 else align ri 2 + 2 in + fixup_variadic_call ri' rf tyl | (Tfloat | Tany64) :: tyl -> if Archi.ptr64 then begin - let rs = float_param_regs.(pos) - and rd = int_param_regs.(pos) in + let rs = float_param_regs.(rf) + and rd = int_param_regs.(ri) in emit (Pfmvxd(rd, rs)); - fixup_variadic_call (pos + 1) tyl + fixup_variadic_call (ri + 1) (rf + 1) tyl end else begin - let pos = align pos 2 in - if pos < 8 then begin - let rs = float_param_regs.(pos) - and rd1 = int_param_regs.(pos) - and rd2 = int_param_regs.(pos + 1) in + 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 (pos + 2) tyl + fixup_variadic_call (ri + 2) (rf + 1) tyl end end let fixup_call sg = - if sg.sig_cc.cc_vararg then fixup_variadic_call 0 sg.sig_args + if sg.sig_cc.cc_vararg then fixup_variadic_call 0 0 sg.sig_args (* Handling of annotations *) diff --git a/riscV/Conventions1.v b/riscV/Conventions1.v index a705c954..7f8048f6 100644 --- a/riscV/Conventions1.v +++ b/riscV/Conventions1.v @@ -105,7 +105,9 @@ Definition is_float_reg (r: mreg) := of function arguments), but this leaves much liberty in choosing actual locations. To ensure binary interoperability of code generated by our compiler with libraries compiled by another compiler, we - implement the standard RISC-V conventions. *) + implement the standard RISC-V conventions as found here: + https://github.com/riscv/riscv-elf-psabi-doc/blob/master/riscv-elf.md +*) (** ** Location of function result *) @@ -169,38 +171,32 @@ Qed. (** ** Location of function arguments *) -(** The RISC-V ABI states the following convention for passing arguments +(** The RISC-V ABI states the following conventions for passing arguments to a function: -- Arguments are passed in registers when possible. - -- Up to eight integer registers (ai: int_param_regs) and up to eight - floating-point registers (fai: float_param_regs) are used for this - purpose. - -- If the arguments to a function are conceptualized as fields of a C - struct, each with pointer alignment, the argument registers are a - shadow of the first eight pointer-words of that struct. If argument - i < 8 is a floating-point type, it is passed in floating-point - register fa_i; otherwise, it is passed in integer register a_i. - -- When primitive arguments twice the size of a pointer-word are passed - on the stack, they are naturally aligned. When they are passed in the - integer registers, they reside in an aligned even-odd register pair, - with the even register holding the least-significant bits. - -- Floating-point arguments to variadic functions (except those that - are explicitly named in the parameter list) are passed in integer - registers. - -- The portion of the conceptual struct that is not passed in argument - registers is passed on the stack. The stack pointer sp points to the - first argument not passed in a register. - -The bit about variadic functions doesn't quite fit CompCert's model. -We do our best by passing the FP arguments in registers, as usual, -and reserving the corresponding integer registers, so that fixup -code can be introduced in the Asmexpand pass. +- 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. + +- 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. + +- 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. + +- RV32, variadic: same, but arguments of 64-bit types (integers as well + as floats) are passed in two consecutive aligned integer registers + (a(2i), a(2i+1)). + +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 +arguments in registers, as usual, and reserving the corresponding +integer registers, so that fixup code can be introduced in the +Asmexpand pass. *) Definition int_param_regs := @@ -208,62 +204,84 @@ Definition int_param_regs := Definition float_param_regs := F10 :: F11 :: F12 :: F13 :: F14 :: F15 :: F16 :: F17 :: nil. -Definition one_arg (regs: list mreg) (rn: Z) (ofs: Z) (ty: typ) - (rec: Z -> Z -> list (rpair loc)) := - match list_nth_z regs rn with +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 | Some r => - One(R r) :: rec (rn + 1) ofs + One(R r) :: rec (ri + 1) rf ofs | None => let ofs := align ofs (typealign ty) in - One(S Outgoing ofs ty) :: rec rn (ofs + (if Archi.ptr64 then 2 else typesize ty)) + One(S Outgoing ofs ty) + :: rec ri rf (ofs + (if Archi.ptr64 then 2 else typesize ty)) end. -Definition two_args (regs: list mreg) (rn: Z) (ofs: Z) - (rec: Z -> Z -> list (rpair loc)) := - let rn := align rn 2 in - match list_nth_z regs rn, list_nth_z regs (rn + 1) with - | Some r1, Some r2 => - Twolong (R r2) (R r1) :: rec (rn + 2) ofs - | _, _ => - let ofs := align ofs 2 in - Twolong (S Outgoing (ofs + 1) Tint) (S Outgoing ofs Tint) :: - rec rn (ofs + 2) +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 + | 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 *) + let ofs := align ofs (typealign 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 (typealign ty) in + One(S Outgoing ofs ty) + :: rec ri rf (ofs + (if Archi.ptr64 then 2 else typesize ty)) end. -Definition hybrid_arg (regs: list mreg) (rn: Z) (ofs: Z) (ty: typ) - (rec: Z -> Z -> list (rpair loc)) := - let rn := align rn 2 in - match list_nth_z regs rn with - | Some r => - One (R r) :: rec (rn + 2) ofs - | None => +Definition split_long_arg (va: bool) (ri rf ofs: Z) + (rec: Z -> Z -> Z -> list (rpair loc)) := + let ri := if va then align ri 2 else ri in + match list_nth_z int_param_regs ri, list_nth_z int_param_regs (ri + 1) with + | Some r1, Some r2 => + Twolong (R r2) (R r1) :: rec (ri + 2) rf ofs + | Some r1, None => + Twolong (S Outgoing ofs Tint) (R r1) :: rec (ri + 1) rf (ofs + 1) + | None, _ => let ofs := align ofs 2 in - One (S Outgoing ofs ty) :: rec rn (ofs + 2) + Twolong (S Outgoing (ofs + 1) Tint) (S Outgoing ofs Tint) :: + rec ri rf (ofs + 2) end. Fixpoint loc_arguments_rec (va: bool) - (tyl: list typ) (r ofs: Z) {struct tyl} : list (rpair loc) := + (tyl: list typ) (ri rf ofs: Z) {struct tyl} : list (rpair loc) := match tyl with | nil => nil | (Tint | Tany32) as ty :: tys => - one_arg int_param_regs r ofs ty (loc_arguments_rec va tys) + (* pass in one integer register or on stack *) + int_arg ri rf ofs ty (loc_arguments_rec va tys) | Tsingle as ty :: tys => - one_arg float_param_regs r ofs ty (loc_arguments_rec va 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) | Tlong as ty :: tys => - if Archi.ptr64 - then one_arg int_param_regs r ofs ty (loc_arguments_rec va tys) - else two_args int_param_regs r ofs (loc_arguments_rec va tys) + if Archi.ptr64 then + (* pass in one integer register or on stack *) + int_arg ri rf ofs ty (loc_arguments_rec va tys) + 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) | (Tfloat | Tany64) as ty :: tys => - if va && negb Archi.ptr64 - then hybrid_arg float_param_regs r ofs ty (loc_arguments_rec va tys) - else one_arg float_param_regs r ofs ty (loc_arguments_rec va 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) 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) := - loc_arguments_rec s.(sig_cc).(cc_vararg) s.(sig_args) 0 0. + loc_arguments_rec s.(sig_cc).(cc_vararg) s.(sig_args) 0 0 0. (** Argument locations are either non-temporary registers or [Outgoing] stack slots at nonnegative offsets. *) @@ -276,15 +294,18 @@ Definition loc_argument_acceptable (l: loc) : Prop := end. Lemma loc_arguments_rec_charact: - forall va tyl rn ofs p, + forall va tyl ri rf ofs p, ofs >= 0 -> - In p (loc_arguments_rec va tyl rn ofs) -> forall_rpair loc_argument_acceptable p. + In p (loc_arguments_rec va tyl ri rf ofs) -> forall_rpair loc_argument_acceptable p. Proof. set (OK := fun (l: list (rpair loc)) => forall p, In p l -> forall_rpair loc_argument_acceptable p). - set (OKF := fun (f: Z -> Z -> list (rpair loc)) => - forall rn ofs, ofs >= 0 -> OK (f rn ofs)). - set (OKREGS := fun (l: list mreg) => forall r, In r l -> is_callee_save r = false). + set (OKF := fun (f: Z -> Z -> Z -> list (rpair loc)) => + forall ri rf ofs, ofs >= 0 -> OK (f ri rf ofs)). + assert (CSI: forall r, In r int_param_regs -> is_callee_save r = false). + { decide_goal. } + assert (CSF: forall r, In r float_param_regs -> is_callee_save r = false). + { decide_goal. } assert (AL: forall ofs ty, ofs >= 0 -> align ofs (typealign ty) >= 0). { intros. assert (ofs <= align ofs (typealign ty)) by (apply align_le; apply typealign_pos). @@ -293,73 +314,64 @@ Proof. { destruct Archi.ptr64; omega. } assert (SKK: forall ty, (if Archi.ptr64 then 2 else typesize ty) > 0). { intros. destruct Archi.ptr64. omega. apply typesize_pos. } - assert (A: forall regs rn ofs ty f, - OKREGS regs -> OKF f -> ofs >= 0 -> OK (one_arg regs rn ofs ty f)). - { intros until f; intros OR OF OO; red; unfold one_arg; intros. - destruct (list_nth_z regs rn) as [r|] eqn:NTH; destruct H. - - subst p; simpl. apply OR. eapply list_nth_z_in; eauto. + 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. + destruct (list_nth_z int_param_regs ri) as [r|] eqn:NTH; destruct H. + - subst p; simpl. apply CSI. eapply list_nth_z_in; eauto. - eapply OF; eauto. - subst p; simpl. auto using align_divides, typealign_pos. - eapply OF; [idtac|eauto]. generalize (AL ofs ty OO) (SKK ty); omega. } - assert (B: forall regs rn ofs f, - OKREGS regs -> OKF f -> ofs >= 0 -> OK (two_args regs rn ofs f)). - { intros until f; intros OR OF OO; unfold two_args. - set (rn' := align rn 2). + 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 using align_divides, typealign_pos. + + eapply OF; [idtac|eauto]. generalize (AL ofs ty OO) (SKK ty); omega. + + subst p; simpl. apply CSF. eapply list_nth_z_in; eauto. + + eapply OF; eauto. + - destruct H. + + subst p; repeat split; auto using align_divides, typealign_pos. + + eapply OF; [idtac|eauto]. generalize (AL ofs ty OO) (SKK ty); omega. + } + assert (C: forall va ri rf ofs f, + OKF f -> ofs >= 0 -> OK (split_long_arg va ri rf ofs f)). + { intros until f; intros OF OO; unfold split_long_arg. + set (ri' := if va then align ri 2 else ri). set (ofs' := align ofs 2). assert (OO': ofs' >= 0) by (apply (AL ofs Tlong); auto). - assert (DFL: OK (Twolong (S Outgoing (ofs' + 1) Tint) (S Outgoing ofs' Tint) - :: f rn' (ofs' + 2))). - { red; simpl; intros. destruct H. - - subst p; simpl. - repeat split; auto using Z.divide_1_l. omega. - - eapply OF; [idtac|eauto]. omega. - } - destruct (list_nth_z regs rn') as [r1|] eqn:NTH1; - destruct (list_nth_z regs (rn' + 1)) as [r2|] eqn:NTH2; - try apply DFL. - red; simpl; intros; destruct H. - - subst p; simpl. split; apply OR; eauto using list_nth_z_in. - - eapply OF; [idtac|eauto]. auto. - } - assert (C: forall regs rn ofs ty f, - OKREGS regs -> OKF f -> ofs >= 0 -> typealign ty = 1 -> OK (hybrid_arg regs rn ofs ty f)). - { intros until f; intros OR OF OO OTY; unfold hybrid_arg; red; intros. - set (rn' := align rn 2) in *. - destruct (list_nth_z regs rn') as [r|] eqn:NTH; destruct H. - - subst p; simpl. apply OR. eapply list_nth_z_in; eauto. - - eapply OF; eauto. - - subst p; simpl. rewrite OTY. split. apply (AL ofs Tlong OO). apply Z.divide_1_l. - - eapply OF; [idtac|eauto]. generalize (AL ofs Tlong OO); simpl; omega. + destruct (list_nth_z int_param_regs ri') as [r1|] eqn:NTH1; + [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. + - 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. + - red; simpl; intros; destruct H. + + subst p; repeat split; auto using Z.divide_1_l. omega. + + eapply OF; [idtac|eauto]. omega. } - assert (D: OKREGS int_param_regs). - { red. decide_goal. } - assert (E: OKREGS float_param_regs). - { red. decide_goal. } - - cut (forall va tyl rn ofs, ofs >= 0 -> OK (loc_arguments_rec va tyl rn ofs)). + cut (forall va tyl ri rf ofs, ofs >= 0 -> OK (loc_arguments_rec va tyl 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 *) - destruct (va && negb Archi.ptr64). - apply C; auto. - apply A; auto. ++ (* float *) apply B; auto. + (* long *) destruct Archi.ptr64. apply A; auto. - apply B; auto. -+ (* single *) - apply A; auto. -+ (* any32 *) - apply A; auto. -+ (* any64 *) - destruct (va && negb Archi.ptr64). apply C; auto. - apply A; auto. ++ (* single *) apply B; auto. ++ (* any32 *) apply A; auto. ++ (* any64 *) apply B; auto. Qed. Lemma loc_arguments_acceptable: -- cgit From 222c9047d61961db9c6b19fed5ca49829223fd33 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Thu, 27 Feb 2020 05:45:13 +0100 Subject: CSE2 now uses is_trivial_op --- backend/CSE2.v | 2 +- backend/CSE2proof.v | 23 +++++++++++++++++++++-- 2 files changed, 22 insertions(+), 3 deletions(-) diff --git a/backend/CSE2.v b/backend/CSE2.v index 1e3bc3b7..19b633b0 100644 --- a/backend/CSE2.v +++ b/backend/CSE2.v @@ -525,7 +525,7 @@ Definition transf_instr (fmap : option (PMap.t RB.t)) match instr with | Iop op args dst s => let args' := subst_args fmap pc args in - match find_op_in_fmap fmap pc op args' with + match (if is_trivial_op op then None else find_op_in_fmap fmap pc op args') with | None => Iop op args' dst s | Some src => Iop Omove (src::nil) dst s end diff --git a/backend/CSE2proof.v b/backend/CSE2proof.v index 0b92f5e5..eb4268f0 100644 --- a/backend/CSE2proof.v +++ b/backend/CSE2proof.v @@ -1401,6 +1401,24 @@ Proof. Qed. Hint Resolve wellformed_mem_mpc : wellformed. +Lemma match_same_option : + forall { A B : Type}, + forall x : option A, + forall y : B, + (match x with Some _ => y | None => y end) = y. +Proof. + destruct x; trivial. +Qed. + +Lemma match_same_bool : + forall { B : Type}, + forall x : bool, + forall y : B, + (if x then y else y) = y. +Proof. + destruct x; trivial. +Qed. + Lemma step_simulation: forall S1 t S2, RTL.step ge S1 t S2 -> forall S1', match_states S1 S1' -> @@ -1428,8 +1446,9 @@ Proof. reflexivity. - (* op *) unfold transf_instr in *. - destruct find_op_in_fmap eqn:FIND_OP. + destruct (if is_trivial_op op then None else find_op_in_fmap (forward_map f) pc op (subst_args (forward_map f) pc args)) eqn:FIND_OP. { + destruct is_trivial_op. discriminate. unfold find_op_in_fmap, fmap_sem', fmap_sem in *. destruct (forward_map f) as [map |] eqn:MAP. 2: discriminate. @@ -1846,4 +1865,4 @@ Proof. exact step_simulation. Qed. -End PRESERVATION. \ No newline at end of file +End PRESERVATION. -- cgit From e66be6e05b190c51b38d628884ef3e015ebf73fd Mon Sep 17 00:00:00 2001 From: Bernhard Schommer Date: Mon, 24 Feb 2020 19:59:43 +0100 Subject: Make single arg alignment depend on toolchain. GCC does passes single arguments as singles on the stack but diab and the eabi say single arguments should be passed as double on the stack. This commit changes the alignment of single arguments to 4 for gcc based backends. --- powerpc/Archi.v | 3 +++ powerpc/Conventions1.v | 17 ++++++++++++++--- powerpc/extractionMachdep.v | 3 +++ 3 files changed, 20 insertions(+), 3 deletions(-) diff --git a/powerpc/Archi.v b/powerpc/Archi.v index ab348c14..88fff302 100644 --- a/powerpc/Archi.v +++ b/powerpc/Archi.v @@ -30,6 +30,9 @@ Definition align_float64 := 8%Z. (** Can we use the 64-bit extensions to the PowerPC architecture? *) Parameter ppc64 : bool. +(** Should singles be passed as single or double *) +Parameter single_passed_as_single : bool. + Definition splitlong := negb ppc64. Lemma splitlong_ptr32: splitlong = true -> ptr64 = false. diff --git a/powerpc/Conventions1.v b/powerpc/Conventions1.v index 1f048694..5639ff8d 100644 --- a/powerpc/Conventions1.v +++ b/powerpc/Conventions1.v @@ -208,7 +208,16 @@ Fixpoint loc_arguments_rec | Some ireg => One (R ireg) :: loc_arguments_rec tys (ir + 1) fr ofs end - | (Tfloat | Tsingle | Tany64) as ty :: tys => + | Tsingle as ty :: tys => + match list_nth_z float_param_regs fr with + | None => + let ty := if Archi.single_passed_as_single then Tsingle else Tfloat in + let ofs := align ofs (typesize ty) in + One (S Outgoing ofs Tsingle) :: loc_arguments_rec tys ir fr (ofs + (typesize ty)) + | Some freg => + One (R freg) :: loc_arguments_rec tys ir (fr + 1) ofs + end + | (Tfloat | Tany64) as ty :: tys => match list_nth_z float_param_regs fr with | None => let ofs := align ofs 2 in @@ -295,12 +304,14 @@ Opaque list_nth_z. apply align_divides; omega. apply Z.divide_1_l. apply Z.divide_1_l. eapply Y; eauto. omega. - (* single *) + assert (ofs <= align ofs 1) by (apply align_le; omega). assert (ofs <= align ofs 2) by (apply align_le; omega). destruct (list_nth_z float_param_regs fr) as [r|] eqn:E; destruct H. subst. right. eapply list_nth_z_in; eauto. eapply IHtyl; eauto. - subst. split. omega. apply Z.divide_1_l. - eapply Y; eauto. omega. + subst. split. destruct Archi.single_passed_as_single; simpl; omega. + destruct Archi.single_passed_as_single; simpl; apply Z.divide_1_l. + eapply Y; eauto. destruct Archi.single_passed_as_single; simpl; omega. - (* any32 *) destruct (list_nth_z int_param_regs ir) as [r|] eqn:E; destruct H. subst. left. eapply list_nth_z_in; eauto. diff --git a/powerpc/extractionMachdep.v b/powerpc/extractionMachdep.v index 7482435f..a3e945bf 100644 --- a/powerpc/extractionMachdep.v +++ b/powerpc/extractionMachdep.v @@ -34,3 +34,6 @@ Extract Constant Archi.ppc64 => | ""e5500"" -> true | _ -> false end". + +(* Choice of passing of single *) +Extract Constant Archi.single_passed_as_single => "Configuration.gnu_toolchain". -- cgit From 78d76b65b417b2724cc54a7e5fc5d434d8fc57b5 Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Thu, 27 Feb 2020 10:04:43 +0100 Subject: Define IRC.class_of_type for types Tany32, Tany64 Until now the types Tany32 and Tany64 were not used prior to register allocation, so IRC.class_of_type did not need to be defined for those types. However, there are possible uses of stack slots of type Tany32 and Tany64 to specify calling conventions. For this purpose, we need to define class_of_type for Tany32 and Tany64. We follow the informal convention that Tany32 goes in integer registers and Tany64 goes into integer registers if 64-bit wide, or FP registers otherwise. --- backend/IRC.ml | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/backend/IRC.ml b/backend/IRC.ml index 43955897..b359da35 100644 --- a/backend/IRC.ml +++ b/backend/IRC.ml @@ -240,7 +240,8 @@ type graph = { let class_of_type = function | Tint | Tlong -> 0 | Tfloat | Tsingle -> 1 - | Tany32 | Tany64 -> assert false + | Tany32 -> 0 + | Tany64 -> if Archi.ptr64 then 0 else 1 let class_of_reg r = if Conventions1.is_float_reg r then 1 else 0 -- cgit From 9190ca38b3ae098186421a7cc21a087666a6a677 Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Thu, 27 Feb 2020 10:10:23 +0100 Subject: In strict PPC ABI mode, pass single FP on stack in double FP format The EABI and the SVR4 ABI state that single-precision FP arguments passed on stack are passed as a 64-bit word, extended to double-precision. This commit implements this behavior by using a stack slot of type Tany64. Not only this ensures that the slot is of size and alignment 8 bytes, but it also ensures that it is accessed by stfd and lfd instructions, using single-extended-to-double format. --- powerpc/Conventions1.v | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/powerpc/Conventions1.v b/powerpc/Conventions1.v index 5639ff8d..5c9cbd4f 100644 --- a/powerpc/Conventions1.v +++ b/powerpc/Conventions1.v @@ -211,9 +211,9 @@ Fixpoint loc_arguments_rec | Tsingle as ty :: tys => match list_nth_z float_param_regs fr with | None => - let ty := if Archi.single_passed_as_single then Tsingle else Tfloat in + let ty := if Archi.single_passed_as_single then Tsingle else Tany64 in let ofs := align ofs (typesize ty) in - One (S Outgoing ofs Tsingle) :: loc_arguments_rec tys ir fr (ofs + (typesize ty)) + One (S Outgoing ofs ty) :: loc_arguments_rec tys ir fr (ofs + (typesize ty)) | Some freg => One (R freg) :: loc_arguments_rec tys ir (fr + 1) ofs end -- cgit From 35ba7f373963d8a1f0094abd457809d1e3c3cdb4 Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Thu, 27 Feb 2020 10:15:40 +0100 Subject: Documentation comment for single_passed_as_single --- powerpc/Archi.v | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/powerpc/Archi.v b/powerpc/Archi.v index 88fff302..10f38391 100644 --- a/powerpc/Archi.v +++ b/powerpc/Archi.v @@ -30,7 +30,8 @@ Definition align_float64 := 8%Z. (** Can we use the 64-bit extensions to the PowerPC architecture? *) Parameter ppc64 : bool. -(** Should singles be passed as single or double *) +(** Should single-precision FP arguments passed on stack be passed + as singles or use double FP format. *) Parameter single_passed_as_single : bool. Definition splitlong := negb ppc64. -- cgit From 8587b8310a91702e2635a689e1622a87b7bf8985 Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Mon, 2 Mar 2020 10:32:07 +0100 Subject: Weaker ec_readonly condition over external calls (#225) Currently we require the memory to be unchanged on readonly locations. This is too strong. For example, current permissions could decrease from readonly to none. This commit weakens the ec_readonly condition to the strict minimum needed to show the correctness of value analysis for const globals. --- backend/ValueAnalysis.v | 5 ++--- common/Events.v | 48 +++++++++++++++++++++++++++++++++--------------- 2 files changed, 35 insertions(+), 18 deletions(-) diff --git a/backend/ValueAnalysis.v b/backend/ValueAnalysis.v index 2b233900..b0ce019c 100644 --- a/backend/ValueAnalysis.v +++ b/backend/ValueAnalysis.v @@ -1039,9 +1039,8 @@ Proof. red; simpl; intros. destruct (plt b (Mem.nextblock m)). exploit RO; eauto. intros (R & P & Q). split; auto. - split. apply bmatch_incr with bc; auto. apply bmatch_inv with m; auto. - intros. eapply Mem.loadbytes_unchanged_on_1. eapply external_call_readonly; eauto. - auto. intros; red. apply Q. + split. apply bmatch_incr with bc; auto. apply bmatch_ext with m; auto. + intros. eapply external_call_readonly with (m2 := m'); eauto. intros; red; intros; elim (Q ofs). eapply external_call_max_perm with (m2 := m'); eauto. destruct (j' b); congruence. diff --git a/common/Events.v b/common/Events.v index 10e0c232..4431b0b7 100644 --- a/common/Events.v +++ b/common/Events.v @@ -649,9 +649,12 @@ Record extcall_properties (sem: extcall_sem) (sg: signature) : Prop := (** External call cannot modify memory unless they have [Max, Writable] permissions. *) ec_readonly: - forall ge vargs m1 t vres m2, + forall ge vargs m1 t vres m2 b ofs n bytes, sem ge vargs m1 t vres m2 -> - Mem.unchanged_on (loc_not_writable m1) m1 m2; + Mem.valid_block m1 b -> + Mem.loadbytes m2 b ofs n = Some bytes -> + (forall i, ofs <= i < ofs + n -> ~Mem.perm m1 b i Max Writable) -> + Mem.loadbytes m1 b ofs n = Some bytes; (** External calls must commute with memory extensions, in the following sense. *) @@ -784,7 +787,7 @@ Proof. (* max perms *) - inv H; auto. (* readonly *) -- inv H. apply Mem.unchanged_on_refl. +- inv H; auto. (* mem extends *) - inv H. inv H1. inv H6. inv H4. exploit volatile_load_extends; eauto. intros [v' [A B]]. @@ -833,14 +836,27 @@ Proof. rewrite C; auto. Qed. +Lemma unchanged_on_readonly: + forall m1 m2 b ofs n bytes, + Mem.unchanged_on (loc_not_writable m1) m1 m2 -> + Mem.valid_block m1 b -> + Mem.loadbytes m2 b ofs n = Some bytes -> + (forall i, ofs <= i < ofs + n -> ~Mem.perm m1 b i Max Writable) -> + Mem.loadbytes m1 b ofs n = Some bytes. +Proof. + intros. + rewrite <- H1. symmetry. + apply Mem.loadbytes_unchanged_on_1 with (P := loc_not_writable m1); auto. +Qed. + Lemma volatile_store_readonly: forall ge chunk1 m1 b1 ofs1 v t m2, volatile_store ge chunk1 m1 b1 ofs1 v t m2 -> Mem.unchanged_on (loc_not_writable m1) m1 m2. Proof. intros. inv H. - apply Mem.unchanged_on_refl. - eapply Mem.store_unchanged_on; eauto. +- apply Mem.unchanged_on_refl. +- eapply Mem.store_unchanged_on; eauto. exploit Mem.store_valid_access_3; eauto. intros [P Q]. intros. unfold loc_not_writable. red; intros. elim H2. apply Mem.perm_cur_max. apply P. auto. @@ -934,7 +950,7 @@ Proof. (* perms *) - inv H. inv H2. auto. eauto with mem. (* readonly *) -- inv H. eapply volatile_store_readonly; eauto. +- inv H. eapply unchanged_on_readonly; eauto. eapply volatile_store_readonly; eauto. (* mem extends*) - inv H. inv H1. inv H6. inv H7. inv H4. exploit volatile_store_extends; eauto. intros [m2' [A [B C]]]. @@ -994,7 +1010,7 @@ Proof. rewrite dec_eq_false. auto. apply Mem.valid_not_valid_diff with m1; eauto with mem. (* readonly *) -- inv H. eapply UNCHANGED; eauto. +- inv H. eapply unchanged_on_readonly; eauto. (* mem extends *) - inv H. inv H1. inv H7. assert (SZ: v2 = Vptrofs sz). @@ -1065,8 +1081,9 @@ Proof. (* perms *) - inv H. eapply Mem.perm_free_3; eauto. (* readonly *) -- inv H. eapply Mem.free_unchanged_on; eauto. - intros. red; intros. elim H3. +- inv H. eapply unchanged_on_readonly; eauto. + eapply Mem.free_unchanged_on; eauto. + intros. red; intros. elim H6. apply Mem.perm_cur_max. apply Mem.perm_implies with Freeable; auto with mem. eapply Mem.free_range_perm; eauto. (* mem extends *) @@ -1159,8 +1176,9 @@ Proof. - (* perms *) intros. inv H. eapply Mem.perm_storebytes_2; eauto. - (* readonly *) - intros. inv H. eapply Mem.storebytes_unchanged_on; eauto. - intros; red; intros. elim H8. + intros. inv H. eapply unchanged_on_readonly; eauto. + eapply Mem.storebytes_unchanged_on; eauto. + intros; red; intros. elim H11. apply Mem.perm_cur_max. eapply Mem.storebytes_range_perm; eauto. - (* extensions *) intros. inv H. @@ -1271,7 +1289,7 @@ Proof. (* perms *) - inv H; auto. (* readonly *) -- inv H. apply Mem.unchanged_on_refl. +- inv H; auto. (* mem extends *) - inv H. exists Vundef; exists m1'; intuition. @@ -1316,7 +1334,7 @@ Proof. (* perms *) - inv H; auto. (* readonly *) -- inv H. apply Mem.unchanged_on_refl. +- inv H; auto. (* mem extends *) - inv H. inv H1. inv H6. exists v2; exists m1'; intuition. @@ -1359,7 +1377,7 @@ Proof. (* perms *) - inv H; auto. (* readonly *) -- inv H. apply Mem.unchanged_on_refl. +- inv H; auto. (* mem extends *) - inv H. exists Vundef; exists m1'; intuition. @@ -1406,7 +1424,7 @@ Proof. (* perms *) - inv H; auto. (* readonly *) -- inv H. apply Mem.unchanged_on_refl. +- inv H; auto. (* mem extends *) - inv H. fold bsem in H2. apply val_inject_list_lessdef in H1. specialize (bs_inject _ bsem _ _ _ H1). -- cgit From f8d3d265f6ef967acf6eea017cb472809096a135 Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Mon, 2 Mar 2020 10:41:11 +0100 Subject: Define the semantics of `free(NULL)` (#226) According to ISO C, `free(NULL)` is correct and does nothing. This commit updates accordingly the formal semantics of the `free` external function and the reference interpreter. Closes: #334 --- cfrontend/Cexec.v | 73 +++++++++++++++++++++++++++++++++---------------------- common/Events.v | 44 ++++++++++++++++++++++----------- 2 files changed, 74 insertions(+), 43 deletions(-) diff --git a/cfrontend/Cexec.v b/cfrontend/Cexec.v index 2942080b..b08c3ad7 100644 --- a/cfrontend/Cexec.v +++ b/cfrontend/Cexec.v @@ -460,6 +460,14 @@ Definition do_ef_free check (zlt 0 (Ptrofs.unsigned sz)); do m' <- Mem.free m b (Ptrofs.unsigned lo - size_chunk Mptr) (Ptrofs.unsigned lo + Ptrofs.unsigned sz); Some(w, E0, Vundef, m') + | Vint n :: nil => + if Int.eq_dec n Int.zero && negb Archi.ptr64 + then Some(w, E0, Vundef, m) + else None + | Vlong n :: nil => + if Int64.eq_dec n Int64.zero && Archi.ptr64 + then Some(w, E0, Vundef, m) + else None | _ => None end. @@ -544,45 +552,51 @@ Proof with try congruence. - eapply do_external_function_sound; eauto. } destruct ef; simpl. -(* EF_external *) +- (* EF_external *) eapply do_external_function_sound; eauto. -(* EF_builtin *) +- (* EF_builtin *) eapply BF_EX; eauto. -(* EF_runtime *) +- (* EF_runtime *) eapply BF_EX; eauto. -(* EF_vload *) +- (* EF_vload *) unfold do_ef_volatile_load. destruct vargs... destruct v... destruct vargs... mydestr. destruct p as [[w'' t''] v]; mydestr. exploit do_volatile_load_sound; eauto. intuition. econstructor; eauto. - auto. -(* EF_vstore *) +- (* EF_vstore *) unfold do_ef_volatile_store. destruct vargs... destruct v... destruct vargs... destruct vargs... mydestr. destruct p as [[w'' t''] m'']. mydestr. exploit do_volatile_store_sound; eauto. intuition. econstructor; eauto. - auto. -(* EF_malloc *) +- (* EF_malloc *) unfold do_ef_malloc. destruct vargs... destruct vargs... mydestr. destruct (Mem.alloc m (- size_chunk Mptr) (Ptrofs.unsigned i)) as [m1 b] eqn:?. mydestr. split. apply SIZE in Heqo. subst v. econstructor; eauto. constructor. -(* EF_free *) - unfold do_ef_free. destruct vargs... destruct v... destruct vargs... - mydestr. split. apply SIZE in Heqo0. econstructor; eauto. congruence. omega. constructor. -(* EF_memcpy *) +- (* EF_free *) + unfold do_ef_free. destruct vargs... destruct v... ++ destruct vargs... mydestr; InvBooleans; subst i. + replace (Vint Int.zero) with Vnullptr. split; constructor. + apply negb_true_iff in H0. unfold Vnullptr; rewrite H0; auto. ++ destruct vargs... mydestr; InvBooleans; subst i. + replace (Vlong Int64.zero) with Vnullptr. split; constructor. + unfold Vnullptr; rewrite H0; auto. ++ destruct vargs... mydestr. + split. apply SIZE in Heqo0. econstructor; eauto. congruence. omega. + constructor. +- (* EF_memcpy *) unfold do_ef_memcpy. destruct vargs... destruct v... destruct vargs... destruct v... destruct vargs... mydestr. apply Decidable_sound in Heqb1. red in Heqb1. split. econstructor; eauto; tauto. constructor. -(* EF_annot *) +- (* EF_annot *) unfold do_ef_annot. mydestr. split. constructor. apply list_eventval_of_val_sound; auto. econstructor. constructor; eauto. constructor. -(* EF_annot_val *) +- (* EF_annot_val *) unfold do_ef_annot_val. destruct vargs... destruct vargs... mydestr. split. constructor. apply eventval_of_val_sound; auto. econstructor. constructor; eauto. constructor. -(* EF_inline_asm *) +- (* EF_inline_asm *) eapply do_inline_assembly_sound; eauto. -(* EF_debug *) +- (* EF_debug *) unfold do_ef_debug. mydestr. split; constructor. Qed. @@ -605,37 +619,38 @@ Proof. - eapply do_external_function_complete; eauto. } destruct ef; simpl in *. -(* EF_external *) +- (* EF_external *) eapply do_external_function_complete; eauto. -(* EF_builtin *) +- (* EF_builtin *) eapply BF_EX; eauto. -(* EF_runtime *) +- (* EF_runtime *) eapply BF_EX; eauto. -(* EF_vload *) +- (* EF_vload *) inv H; unfold do_ef_volatile_load. exploit do_volatile_load_complete; eauto. intros EQ; rewrite EQ; auto. -(* EF_vstore *) +- (* EF_vstore *) inv H; unfold do_ef_volatile_store. exploit do_volatile_store_complete; eauto. intros EQ; rewrite EQ; auto. -(* EF_malloc *) +- (* EF_malloc *) inv H; unfold do_ef_malloc. inv H0. erewrite SIZE by eauto. rewrite H1, H2. auto. -(* EF_free *) +- (* EF_free *) inv H; unfold do_ef_free. - inv H0. rewrite H1. erewrite SIZE by eauto. rewrite zlt_true. rewrite H3. auto. omega. -(* EF_memcpy *) ++ inv H0. rewrite H1. erewrite SIZE by eauto. rewrite zlt_true. rewrite H3. auto. omega. ++ inv H0. unfold Vnullptr; destruct Archi.ptr64; auto. +- (* EF_memcpy *) inv H; unfold do_ef_memcpy. inv H0. rewrite Decidable_complete. rewrite H7; rewrite H8; auto. red. tauto. -(* EF_annot *) +- (* EF_annot *) inv H; unfold do_ef_annot. inv H0. inv H6. inv H4. rewrite (list_eventval_of_val_complete _ _ _ H1). auto. -(* EF_annot_val *) +- (* EF_annot_val *) inv H; unfold do_ef_annot_val. inv H0. inv H6. inv H4. rewrite (eventval_of_val_complete _ _ _ H1). auto. -(* EF_inline_asm *) +- (* EF_inline_asm *) eapply do_inline_assembly_complete; eauto. -(* EF_debug *) +- (* EF_debug *) inv H. inv H0. reflexivity. Qed. diff --git a/common/Events.v b/common/Events.v index 4431b0b7..022adaef 100644 --- a/common/Events.v +++ b/common/Events.v @@ -1061,11 +1061,13 @@ Qed. Inductive extcall_free_sem (ge: Senv.t): list val -> mem -> trace -> val -> mem -> Prop := - | extcall_free_sem_intro: forall b lo sz m m', + | extcall_free_sem_ptr: forall b lo sz m m', Mem.load Mptr m b (Ptrofs.unsigned lo - size_chunk Mptr) = Some (Vptrofs sz) -> Ptrofs.unsigned sz > 0 -> Mem.free m b (Ptrofs.unsigned lo - size_chunk Mptr) (Ptrofs.unsigned lo + Ptrofs.unsigned sz) = Some m' -> - extcall_free_sem ge (Vptr b lo :: nil) m E0 Vundef m'. + extcall_free_sem ge (Vptr b lo :: nil) m E0 Vundef m' + | extcall_free_sem_null: forall m, + extcall_free_sem ge (Vnullptr :: nil) m E0 Vundef m. Lemma extcall_free_ok: extcall_properties extcall_free_sem @@ -1073,27 +1075,29 @@ Lemma extcall_free_ok: Proof. constructor; intros. (* well typed *) -- inv H. simpl. auto. +- inv H; simpl; auto. (* symbols preserved *) - inv H0; econstructor; eauto. (* valid block *) -- inv H. eauto with mem. +- inv H; eauto with mem. (* perms *) -- inv H. eapply Mem.perm_free_3; eauto. +- inv H; eauto using Mem.perm_free_3. (* readonly *) -- inv H. eapply unchanged_on_readonly; eauto. - eapply Mem.free_unchanged_on; eauto. +- eapply unchanged_on_readonly; eauto. inv H. ++ eapply Mem.free_unchanged_on; eauto. intros. red; intros. elim H6. apply Mem.perm_cur_max. apply Mem.perm_implies with Freeable; auto with mem. eapply Mem.free_range_perm; eauto. ++ apply Mem.unchanged_on_refl. (* mem extends *) -- inv H. inv H1. inv H8. inv H6. +- inv H. ++ inv H1. inv H8. inv H6. exploit Mem.load_extends; eauto. intros [v' [A B]]. assert (v' = Vptrofs sz). { unfold Vptrofs in *; destruct Archi.ptr64; inv B; auto. } subst v'. exploit Mem.free_parallel_extends; eauto. intros [m2' [C D]]. - exists Vundef; exists m2'; intuition. + exists Vundef; exists m2'; intuition auto. econstructor; eauto. eapply Mem.free_unchanged_on; eauto. unfold loc_out_of_bounds; intros. @@ -1101,8 +1105,14 @@ Proof. { apply Mem.perm_cur_max. apply Mem.perm_implies with Freeable; auto with mem. eapply Mem.free_range_perm. eexact H4. eauto. } tauto. ++ inv H1. inv H5. replace v2 with Vnullptr. + exists Vundef; exists m1'; intuition auto. + constructor. + apply Mem.unchanged_on_refl. + unfold Vnullptr in *; destruct Archi.ptr64; inv H3; auto. (* mem inject *) -- inv H0. inv H2. inv H7. inv H9. +- inv H0. ++ inv H2. inv H7. inv H9. exploit Mem.load_inject; eauto. intros [v' [A B]]. assert (v' = Vptrofs sz). { unfold Vptrofs in *; destruct Archi.ptr64; inv B; auto. } @@ -1116,7 +1126,7 @@ Proof. intro EQ. exploit Mem.free_parallel_inject; eauto. intros (m2' & C & D). exists f, Vundef, m2'; split. - apply extcall_free_sem_intro with (sz := sz) (m' := m2'). + apply extcall_free_sem_ptr with (sz := sz) (m' := m2'). rewrite EQ. rewrite <- A. f_equal. omega. auto. auto. rewrite ! EQ. rewrite <- C. f_equal; omega. @@ -1129,14 +1139,19 @@ Proof. apply P. omega. split. auto. red; intros. congruence. ++ inv H2. inv H6. replace v' with Vnullptr. + exists f, Vundef, m1'; intuition auto using Mem.unchanged_on_refl. + constructor. + red; intros; congruence. + unfold Vnullptr in *; destruct Archi.ptr64; inv H4; auto. (* trace length *) - inv H; simpl; omega. (* receptive *) -- assert (t1 = t2). inv H; inv H0; auto. subst t2. +- assert (t1 = t2) by (inv H; inv H0; auto). subst t2. exists vres1; exists m1; auto. (* determ *) -- inv H; inv H0. - assert (EQ1: Vptrofs sz0 = Vptrofs sz) by congruence. +- inv H; inv H0; try (unfold Vnullptr in *; discriminate). ++ assert (EQ1: Vptrofs sz0 = Vptrofs sz) by congruence. assert (EQ2: sz0 = sz). { unfold Vptrofs in EQ1; destruct Archi.ptr64 eqn:SF. rewrite <- (Ptrofs.of_int64_to_int64 SF sz0), <- (Ptrofs.of_int64_to_int64 SF sz). congruence. @@ -1144,6 +1159,7 @@ Proof. } subst sz0. split. constructor. intuition congruence. ++ split. constructor. intuition auto. Qed. (** ** Semantics of [memcpy] operations. *) -- cgit From 94558ecb3e48261f12c644045d40c7d0784415e0 Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Mon, 2 Mar 2020 11:39:06 +0100 Subject: Define the semantics of `free(NULL)`, continued The proof script for Events.excall_free_ok was incomplete if Archi.ptr64 is unknown (as in the RISC-V case). --- common/Events.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/common/Events.v b/common/Events.v index 022adaef..28bb992a 100644 --- a/common/Events.v +++ b/common/Events.v @@ -1150,7 +1150,7 @@ Proof. - assert (t1 = t2) by (inv H; inv H0; auto). subst t2. exists vres1; exists m1; auto. (* determ *) -- inv H; inv H0; try (unfold Vnullptr in *; discriminate). +- inv H; inv H0; try (unfold Vnullptr in *; destruct Archi.ptr64; discriminate). + assert (EQ1: Vptrofs sz0 = Vptrofs sz) by congruence. assert (EQ2: sz0 = sz). { unfold Vptrofs in EQ1; destruct Archi.ptr64 eqn:SF. -- cgit From 6af8f4275f7f9572d4d0783818cbfb85357807c6 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 2 Mar 2020 13:17:54 +0100 Subject: loadv_storev_same --- common/Memory.v | 18 ++++++++++++++++++ 1 file changed, 18 insertions(+) diff --git a/common/Memory.v b/common/Memory.v index b68a5049..89b920b3 100644 --- a/common/Memory.v +++ b/common/Memory.v @@ -1169,6 +1169,24 @@ Local Hint Resolve store_valid_block_1 store_valid_block_2: mem. Local Hint Resolve store_valid_access_1 store_valid_access_2 store_valid_access_3: mem. + +Section STOREV. +Variable chunk: memory_chunk. +Variable m1: mem. +Variables addr v: val. +Variable m2: mem. +Hypothesis STORE: storev chunk m1 addr v = Some m2. + + +Theorem loadv_storev_same: + loadv chunk m2 addr = Some (Val.load_result chunk v). +Proof. + destruct addr; simpl in *; try discriminate. + eapply load_store_same. + eassumption. +Qed. +End STOREV. + Lemma load_store_overlap: forall chunk m1 b ofs v m2 chunk' ofs' v', store chunk m1 b ofs v = Some m2 -> -- cgit From fbfbc3c4cbe250a40513e5dabcd6930b39043ea3 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 2 Mar 2020 18:14:23 +0100 Subject: playing with offsets --- backend/CSE2proof.v | 68 ++++++++++++++++++++++++++++++++++++++++++++++++++++- 1 file changed, 67 insertions(+), 1 deletion(-) diff --git a/backend/CSE2proof.v b/backend/CSE2proof.v index 82fa8a28..4dd8b054 100644 --- a/backend/CSE2proof.v +++ b/backend/CSE2proof.v @@ -14,6 +14,7 @@ Require Import Globalenvs Values. Require Import Linking Values Memory Globalenvs Events Smallstep. Require Import Registers Op RTL. Require Import CSE2. +Require Import Lia. Lemma args_unaffected: forall rs : regset, @@ -697,6 +698,71 @@ Proof. Qed. End SAME_MEMORY. +Section MEMORY_WRITE. + Variable m m2 : mem. + Variable chunkw chunkr : memory_chunk. + Variable base : val. + Variable ofsw ofsr : Z. + + Hypothesis RANGEW : 0 <= ofsw <= 18446744073709551608. + Hypothesis RANGER : 0 <= ofsr <= 18446744073709551608. + + Lemma size_chunk_bounded : + forall chunk, 0 <= size_chunk chunk <= 8. + Proof. + destruct chunk; simpl; lia. + Qed. + + Hypothesis no_overlap: + ofsw + size_chunk chunkw <= ofsr + \/ ofsr + size_chunk chunkr <= ofsw. + + Variable addrw addrr valw valr : val. + + Hypothesis ADDRW : eval_addressing genv sp + (Aindexed ofsw) (base :: nil) = Some addrw. + Hypothesis ADDRR : eval_addressing genv sp + (Aindexed ofsr) (base :: nil) = Some addrr. + Hypothesis STORE : Mem.storev chunkw m addrw valw = Some m2. + Hypothesis READ : Mem.loadv chunkr m addrr = Some valr. + + Theorem load_store_away : + Mem.loadv chunkr m2 addrr = Some valr. + Proof. + pose proof (size_chunk_bounded chunkr) as size_chunkr_bounded. + pose proof (size_chunk_bounded chunkw) as size_chunkw_bounded. + + destruct addrr ; simpl in * ; try discriminate. + unfold eval_addressing in *. + destruct Archi.ptr64 eqn:PTR64. + { + unfold eval_addressing64 in *. + inv ADDRW. + destruct base; simpl in *; try discriminate. + rewrite PTR64 in *. + inv ADDRR. + rewrite <- READ. + eapply Mem.load_store_other. + exact STORE. + right. + destruct (Ptrofs.unsigned_add_either i0 + (Ptrofs.of_int64 (Int64.repr ofsr))) as [OFSR | OFSR]; + rewrite OFSR. + all: destruct (Ptrofs.unsigned_add_either i0 + (Ptrofs.of_int64 (Int64.repr ofsw))) as [OFSW | OFSW]; + rewrite OFSW. + all: unfold Ptrofs.of_int64. + all: repeat rewrite Int64.unsigned_repr by (change Int64.max_unsigned with 18446744073709551615; lia). + all: repeat rewrite Ptrofs.unsigned_repr by (change Ptrofs.max_unsigned with 18446744073709551615; lia). + all: change Ptrofs.modulus with 18446744073709551616. + all: intuition lia. + } + { + destruct base; simpl in *; try discriminate. + } + Qed. +End MEMORY_WRITE. + Lemma kill_mem_sound : forall m m' : mem, forall rel : RELATION.t, @@ -1305,4 +1371,4 @@ Proof. exact step_simulation. Qed. -End PRESERVATION. \ No newline at end of file +End PRESERVATION. -- cgit From c11b19619e82377be9c43e926d66086124637044 Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Mon, 2 Mar 2020 18:59:02 +0100 Subject: Update the RISC-V calling conventions, continued (#227) Double FP arguments passed on stack were incorrectly aligned: they must be 8-aligned but were 4-aligned only. This was due to the use of `Location.typealign`, which is the minimal hardware-supported alignment for a given type, namely 1 word for type Tfloat. To get the correct alignments, `Location.typesize` must be used instead. --- riscV/Conventions1.v | 17 ++++++++++------- 1 file changed, 10 insertions(+), 7 deletions(-) diff --git a/riscV/Conventions1.v b/riscV/Conventions1.v index 7f8048f6..17326139 100644 --- a/riscV/Conventions1.v +++ b/riscV/Conventions1.v @@ -210,7 +210,7 @@ Definition int_arg (ri rf ofs: Z) (ty: typ) | Some r => One(R r) :: rec (ri + 1) rf ofs | None => - let ofs := align ofs (typealign ty) in + let ofs := align ofs (typesize ty) in One(S Outgoing ofs ty) :: rec ri rf (ofs + (if Archi.ptr64 then 2 else typesize ty)) end. @@ -228,13 +228,13 @@ Definition float_arg (va: bool) (ri rf ofs: Z) (ty: typ) One (R r) :: rec ri' (rf + 1) ofs else (* we are out of integer registers, pass argument on stack *) - let ofs := align ofs (typealign ty) in + 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 (typealign ty) in + let ofs := align ofs (typesize ty) in One(S Outgoing ofs ty) :: rec ri rf (ofs + (if Archi.ptr64 then 2 else typesize ty)) end. @@ -306,10 +306,13 @@ Proof. { decide_goal. } assert (CSF: forall r, In r float_param_regs -> is_callee_save r = false). { decide_goal. } - assert (AL: forall ofs ty, ofs >= 0 -> align ofs (typealign ty) >= 0). + assert (AL: forall ofs ty, ofs >= 0 -> align ofs (typesize ty) >= 0). { intros. - assert (ofs <= align ofs (typealign ty)) by (apply align_le; apply typealign_pos). + assert (ofs <= align ofs (typesize ty)) by (apply align_le; apply typesize_pos). omega. } + 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. } assert (SKK: forall ty, (if Archi.ptr64 then 2 else typesize ty) > 0). @@ -332,12 +335,12 @@ Proof. 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 using align_divides, typealign_pos. + + subst p; repeat split; auto. + eapply OF; [idtac|eauto]. generalize (AL ofs ty OO) (SKK ty); omega. + subst p; simpl. apply CSF. eapply list_nth_z_in; eauto. + eapply OF; eauto. - destruct H. - + subst p; repeat split; auto using align_divides, typealign_pos. + + subst p; repeat split; auto. + eapply OF; [idtac|eauto]. generalize (AL ofs ty OO) (SKK ty); omega. } assert (C: forall va ri rf ofs f, -- cgit From 036fc22224c8d171a90b608f6146e742a51e0a25 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 2 Mar 2020 19:10:35 +0100 Subject: works on x86 x86_64 --- backend/CSE2proof.v | 76 +++++++++++++++++++++++++++++++++-------------------- 1 file changed, 47 insertions(+), 29 deletions(-) diff --git a/backend/CSE2proof.v b/backend/CSE2proof.v index 4dd8b054..55ec653c 100644 --- a/backend/CSE2proof.v +++ b/backend/CSE2proof.v @@ -704,12 +704,18 @@ Section MEMORY_WRITE. Variable base : val. Variable ofsw ofsr : Z. + Definition max_chunk_size := 8. + (* Hypothesis RANGEW : 0 <= ofsw <= 18446744073709551608. Hypothesis RANGER : 0 <= ofsr <= 18446744073709551608. - + *) + Hypothesis RANGEW : 0 <= ofsw <= Ptrofs.modulus - max_chunk_size. + Hypothesis RANGER : 0 <= ofsr <= Ptrofs.modulus - max_chunk_size. + Lemma size_chunk_bounded : - forall chunk, 0 <= size_chunk chunk <= 8. + forall chunk, 0 <= size_chunk chunk <= max_chunk_size. Proof. + unfold max_chunk_size. destruct chunk; simpl; lia. Qed. @@ -731,35 +737,47 @@ Section MEMORY_WRITE. Proof. pose proof (size_chunk_bounded chunkr) as size_chunkr_bounded. pose proof (size_chunk_bounded chunkw) as size_chunkw_bounded. - + unfold max_chunk_size in size_chunkr_bounded, size_chunkw_bounded. + try change (Ptrofs.modulus - max_chunk_size) with 4294967288 in *. + try change (Ptrofs.modulus - max_chunk_size) with 18446744073709551608 in *. destruct addrr ; simpl in * ; try discriminate. unfold eval_addressing in *. - destruct Archi.ptr64 eqn:PTR64. - { - unfold eval_addressing64 in *. - inv ADDRW. - destruct base; simpl in *; try discriminate. - rewrite PTR64 in *. - inv ADDRR. - rewrite <- READ. - eapply Mem.load_store_other. - exact STORE. - right. - destruct (Ptrofs.unsigned_add_either i0 - (Ptrofs.of_int64 (Int64.repr ofsr))) as [OFSR | OFSR]; - rewrite OFSR. - all: destruct (Ptrofs.unsigned_add_either i0 - (Ptrofs.of_int64 (Int64.repr ofsw))) as [OFSW | OFSW]; - rewrite OFSW. - all: unfold Ptrofs.of_int64. - all: repeat rewrite Int64.unsigned_repr by (change Int64.max_unsigned with 18446744073709551615; lia). - all: repeat rewrite Ptrofs.unsigned_repr by (change Ptrofs.max_unsigned with 18446744073709551615; lia). - all: change Ptrofs.modulus with 18446744073709551616. - all: intuition lia. - } - { - destruct base; simpl in *; try discriminate. - } + destruct Archi.ptr64 eqn:PTR64; destruct base; simpl in *; try discriminate. + rewrite PTR64 in *. + + inv ADDRR. + inv ADDRW. + rewrite <- READ. + eapply Mem.load_store_other with (chunk := chunkw) (v := valw) (b := b). + exact STORE. + right. + + all: try (destruct (Ptrofs.unsigned_add_either i0 + (Ptrofs.of_int (Int.repr ofsr))) as [OFSR | OFSR]; + rewrite OFSR). + all: try (destruct (Ptrofs.unsigned_add_either i0 + (Ptrofs.of_int64 (Int64.repr ofsr))) as [OFSR | OFSR]; + rewrite OFSR). + all: try (destruct (Ptrofs.unsigned_add_either i0 + (Ptrofs.of_int (Int.repr ofsw))) as [OFSW | OFSW]; + rewrite OFSW). + all: try (destruct (Ptrofs.unsigned_add_either i0 + (Ptrofs.of_int64 (Int64.repr ofsw))) as [OFSW | OFSW]; + rewrite OFSW). + + all: unfold Ptrofs.of_int64. + all: unfold Ptrofs.of_int. + + + all: repeat rewrite Int.unsigned_repr by (change Int.max_unsigned with 4294967295; lia). + all: repeat rewrite Ptrofs.unsigned_repr by (change Ptrofs.max_unsigned with 4294967295; lia). + all: repeat rewrite Int64.unsigned_repr by (change Int64.max_unsigned with 18446744073709551615; lia). + all: repeat rewrite Ptrofs.unsigned_repr by (change Ptrofs.max_unsigned with 18446744073709551615; lia). + + all: try change Ptrofs.modulus with 4294967296. + all: try change Ptrofs.modulus with 18446744073709551616. + + all: intuition lia. Qed. End MEMORY_WRITE. -- cgit From f7ea436ac282b6a4a8ddb2281b6e1d86eee46286 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 2 Mar 2020 19:50:21 +0100 Subject: swap predicate --- backend/CSE2.v | 8 ++++++++ backend/CSE2proof.v | 32 +++++++++++++++++++------------- 2 files changed, 27 insertions(+), 13 deletions(-) diff --git a/backend/CSE2.v b/backend/CSE2.v index b7665097..a03e02a4 100644 --- a/backend/CSE2.v +++ b/backend/CSE2.v @@ -266,6 +266,14 @@ Definition kill_reg (dst : reg) (rel : RELATION.t) := PTree.filter1 (fun x => negb (kill_sym_val dst x)) (PTree.remove dst rel). +Definition max_chunk_size := 8. + +Definition can_swap_accesses_ofs ofsr chunkr ofsw chunkw := + (0 <=? ofsw) && (ofsw <=? (Ptrofs.modulus - max_chunk_size)) + && (0 <=? ofsr) && (ofsr <=? (Ptrofs.modulus - max_chunk_size)) + && ((ofsw + size_chunk chunkw <=? ofsr) || + (ofsr + size_chunk chunkr <=? ofsw)). + Definition kill_sym_val_mem (sv: sym_val) := match sv with | SMove _ => false diff --git a/backend/CSE2proof.v b/backend/CSE2proof.v index 55ec653c..6d5496fd 100644 --- a/backend/CSE2proof.v +++ b/backend/CSE2proof.v @@ -704,14 +704,6 @@ Section MEMORY_WRITE. Variable base : val. Variable ofsw ofsr : Z. - Definition max_chunk_size := 8. - (* - Hypothesis RANGEW : 0 <= ofsw <= 18446744073709551608. - Hypothesis RANGER : 0 <= ofsr <= 18446744073709551608. - *) - Hypothesis RANGEW : 0 <= ofsw <= Ptrofs.modulus - max_chunk_size. - Hypothesis RANGER : 0 <= ofsr <= Ptrofs.modulus - max_chunk_size. - Lemma size_chunk_bounded : forall chunk, 0 <= size_chunk chunk <= max_chunk_size. Proof. @@ -719,10 +711,6 @@ Section MEMORY_WRITE. destruct chunk; simpl; lia. Qed. - Hypothesis no_overlap: - ofsw + size_chunk chunkw <= ofsr - \/ ofsr + size_chunk chunkr <= ofsw. - Variable addrw addrr valw valr : val. Hypothesis ADDRW : eval_addressing genv sp @@ -732,9 +720,15 @@ Section MEMORY_WRITE. Hypothesis STORE : Mem.storev chunkw m addrw valw = Some m2. Hypothesis READ : Mem.loadv chunkr m addrr = Some valr. - Theorem load_store_away : + Lemma load_store_away1 : + forall RANGEW : 0 <= ofsw <= Ptrofs.modulus - max_chunk_size, + forall RANGER : 0 <= ofsr <= Ptrofs.modulus - max_chunk_size, + forall SWAPPABLE : ofsw + size_chunk chunkw <= ofsr + \/ ofsr + size_chunk chunkr <= ofsw, Mem.loadv chunkr m2 addrr = Some valr. Proof. + intros. + pose proof (size_chunk_bounded chunkr) as size_chunkr_bounded. pose proof (size_chunk_bounded chunkw) as size_chunkw_bounded. unfold max_chunk_size in size_chunkr_bounded, size_chunkw_bounded. @@ -779,6 +773,18 @@ Section MEMORY_WRITE. all: intuition lia. Qed. + + Theorem load_store_away : + can_swap_accesses_ofs ofsr chunkr ofsw chunkw = true -> + Mem.loadv chunkr m2 addrr = Some valr. + Proof. + intro SWAP. + unfold can_swap_accesses_ofs in SWAP. + repeat rewrite andb_true_iff in SWAP. + repeat rewrite orb_true_iff in SWAP. + repeat rewrite Z.leb_le in SWAP. + apply load_store_away1; intuition. + Qed. End MEMORY_WRITE. Lemma kill_mem_sound : -- cgit From cf56eea4e093f25a5c01ccac5ede2a69b568af7a Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 2 Mar 2020 20:16:16 +0100 Subject: load_store_away --- backend/CSE2proof.v | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/backend/CSE2proof.v b/backend/CSE2proof.v index 6d5496fd..c6bb00dd 100644 --- a/backend/CSE2proof.v +++ b/backend/CSE2proof.v @@ -783,7 +783,8 @@ Section MEMORY_WRITE. repeat rewrite andb_true_iff in SWAP. repeat rewrite orb_true_iff in SWAP. repeat rewrite Z.leb_le in SWAP. - apply load_store_away1; intuition. + apply load_store_away1. + all: tauto. Qed. End MEMORY_WRITE. -- cgit From 3601929b68ced3777c05cd2861847a111603abee Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 2 Mar 2020 21:34:35 +0100 Subject: kill_store1_sound --- backend/CSE2proof.v | 44 +++++++++++++++++++++++++++++++++++++++++--- 1 file changed, 41 insertions(+), 3 deletions(-) diff --git a/backend/CSE2proof.v b/backend/CSE2proof.v index c6bb00dd..3c42f6e1 100644 --- a/backend/CSE2proof.v +++ b/backend/CSE2proof.v @@ -811,6 +811,44 @@ Proof. } Qed. +Lemma kill_store1_sound : + forall m m' : mem, + forall rel : RELATION.t, + forall chunk addr args a v rs, + (eval_addressing genv sp addr (rs ## args)) = Some a -> + (Mem.storev chunk m a v) = Some m' -> + sem_rel m rel rs -> sem_rel m' (kill_store1 chunk addr args rel) rs. +Proof. + unfold sem_rel, sem_reg. + intros until rs. + intros ADDR STORE SEM x. + pose proof (SEM x) as SEMx. + unfold kill_store1. + rewrite PTree.gfilter1. + destruct (rel ! x) as [ sv | ]. + 2: reflexivity. + destruct sv; simpl in *; trivial. + { + destruct op_depends_on_memory eqn:DEPENDS; simpl; trivial. + rewrite SEMx. + apply op_depends_on_memory_correct; auto. + } + destruct addr; destruct addr0; simpl; trivial. + destruct args as [ | base [ | ]]; simpl; trivial. + destruct args0 as [ | base0 [ | ]]; simpl; trivial. + destruct (peq base base0); simpl; trivial. + subst base0. + destruct (can_swap_accesses_ofs z0 chunk0 z chunk) eqn:SWAP; simpl; trivial. + simpl in *. + destruct (eval_addressing genv sp (Aindexed z0) (rs # base :: nil)) eqn:ADDR0; simpl; trivial. + symmetry. + eapply load_store_away. + exact ADDR. + exact ADDR0. + exact STORE. + congruence. + assumption. +Qed. End SOUNDNESS. Definition match_prog (p tp: RTL.program) := @@ -1193,9 +1231,9 @@ Proof. unfold fmap_sem', fmap_sem in *. destruct (forward_map _) as [map |] eqn:MAP in *; trivial. destruct (map # pc) as [mpc |] eqn:MPC in *; try contradiction. - apply sem_rel_b_ge with (rb2 := Some (kill_mem mpc)); trivial. + apply sem_rel_b_ge with (rb2 := Some (kill_store chunk addr args mpc)); trivial. { - replace (Some (kill_mem mpc)) with (apply_instr' (fn_code f) pc (map # pc)). + replace (Some (kill_store chunk addr args mpc)) with (apply_instr' (fn_code f) pc (map # pc)). { eapply DS.fixpoint_solution with (code := fn_code f) (successors := successors_instr); try eassumption. 2: apply apply_instr'_bot. @@ -1207,7 +1245,7 @@ Proof. rewrite H. reflexivity. } - apply (kill_mem_sound' sp m). + apply (kill_store_sound' sp m). assumption. (* call *) -- cgit From 93bf7e0925b1c11e1874ae5f651970db2bd9823d Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 2 Mar 2020 21:55:54 +0100 Subject: kill_store_sound --- backend/CSE2proof.v | 35 +++++++++++++++++++++++++---------- 1 file changed, 25 insertions(+), 10 deletions(-) diff --git a/backend/CSE2proof.v b/backend/CSE2proof.v index 3c42f6e1..cd9f5f46 100644 --- a/backend/CSE2proof.v +++ b/backend/CSE2proof.v @@ -502,6 +502,20 @@ Proof. destruct s; congruence. Qed. + +Lemma forward_move_rs: + forall rel arg rs, + sem_rel rel rs -> + rs # (forward_move rel arg) = rs # arg. +Proof. + unfold forward_move, sem_rel, sem_reg, sem_sym_val in *. + intros until rs. + intro REL. + pose proof (REL arg) as RELarg. + destruct (rel ! arg); trivial. + destruct s; congruence. +Qed. + Lemma oper_sound : forall rel : RELATION.t, forall op : operation, @@ -811,19 +825,19 @@ Proof. } Qed. -Lemma kill_store1_sound : +Lemma kill_store_sound : forall m m' : mem, forall rel : RELATION.t, forall chunk addr args a v rs, (eval_addressing genv sp addr (rs ## args)) = Some a -> (Mem.storev chunk m a v) = Some m' -> - sem_rel m rel rs -> sem_rel m' (kill_store1 chunk addr args rel) rs. + sem_rel m rel rs -> sem_rel m' (kill_store chunk addr args rel) rs. Proof. unfold sem_rel, sem_reg. intros until rs. intros ADDR STORE SEM x. pose proof (SEM x) as SEMx. - unfold kill_store1. + unfold kill_store, kill_store1. rewrite PTree.gfilter1. destruct (rel ! x) as [ sv | ]. 2: reflexivity. @@ -836,18 +850,19 @@ Proof. destruct addr; destruct addr0; simpl; trivial. destruct args as [ | base [ | ]]; simpl; trivial. destruct args0 as [ | base0 [ | ]]; simpl; trivial. - destruct (peq base base0); simpl; trivial. + destruct (peq (forward_move rel base) base0); simpl; trivial. subst base0. destruct (can_swap_accesses_ofs z0 chunk0 z chunk) eqn:SWAP; simpl; trivial. simpl in *. - destruct (eval_addressing genv sp (Aindexed z0) (rs # base :: nil)) eqn:ADDR0; simpl; trivial. + erewrite forward_move_rs in * by exact SEM. + destruct (eval_addressing genv sp (Aindexed z0) ((rs # base) :: nil)) eqn:ADDR0; simpl; trivial. symmetry. eapply load_store_away. - exact ADDR. - exact ADDR0. - exact STORE. - congruence. - assumption. + - exact ADDR. + - exact ADDR0. + - exact STORE. + - congruence. + - assumption. Qed. End SOUNDNESS. -- cgit From a398b5750ceeeab90a44b2e1d34fe6d5ff8b1f08 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 3 Mar 2020 06:13:05 +0100 Subject: with indexed/indexed alias analysis for x86 --- backend/CSE2.v | 26 ++++++++++++++++++++++++-- backend/CSE2proof.v | 4 ++-- test/cse2/indexed_addr.c | 6 ++++++ 3 files changed, 32 insertions(+), 4 deletions(-) create mode 100644 test/cse2/indexed_addr.c diff --git a/backend/CSE2.v b/backend/CSE2.v index a03e02a4..f5ff8748 100644 --- a/backend/CSE2.v +++ b/backend/CSE2.v @@ -281,16 +281,38 @@ Definition kill_sym_val_mem (sv: sym_val) := | SLoad _ _ _ => true end. +Definition may_overlap chunk addr args chunk' addr' args' := + match addr, addr', args, args' with + | (Aindexed ofs), (Aindexed ofs'), + (base :: nil), (base' :: nil) => + if peq base base' + then negb (can_swap_accesses_ofs ofs' chunk' ofs chunk) + else true + | _, _, _, _ => true + end. + +Definition kill_sym_val_store chunk addr args (sv: sym_val) := + match sv with + | SMove _ => false + | SOp op _ => op_depends_on_memory op + | SLoad chunk' addr' args' => may_overlap chunk addr args chunk' addr' args' + end. + Definition kill_mem (rel : RELATION.t) := PTree.filter1 (fun x => negb (kill_sym_val_mem x)) rel. - Definition forward_move (rel : RELATION.t) (x : reg) : reg := match rel ! x with | Some (SMove org) => org | _ => x end. +Definition kill_store1 chunk addr args rel := + PTree.filter1 (fun x => negb (kill_sym_val_store chunk addr args x)) rel. + +Definition kill_store chunk addr args rel := + kill_store1 chunk addr (List.map (forward_move rel) args) rel. + Definition move (src dst : reg) (rel : RELATION.t) := PTree.set dst (SMove (forward_move rel src)) (kill_reg dst rel). @@ -403,7 +425,7 @@ Definition apply_instr instr (rel : RELATION.t) : RB.t := | Inop _ | Icond _ _ _ _ | Ijumptable _ _ => Some rel - | Istore _ _ _ _ _ => Some (kill_mem rel) + | Istore chunk addr args _ _ => Some (kill_store chunk addr args rel) | Iop op args dst _ => Some (gen_oper op dst args rel) | Iload chunk addr args dst _ => Some (load chunk addr dst args rel) | Icall _ _ _ dst _ => Some (kill_reg dst (kill_mem rel)) diff --git a/backend/CSE2proof.v b/backend/CSE2proof.v index cd9f5f46..e65d9194 100644 --- a/backend/CSE2proof.v +++ b/backend/CSE2proof.v @@ -951,6 +951,7 @@ Definition fmap_sem' := fmap_sem fundef unit ge. Definition subst_arg_ok' := subst_arg_ok fundef unit ge. Definition subst_args_ok' := subst_args_ok fundef unit ge. Definition kill_mem_sound' := kill_mem_sound fundef unit ge. +Definition kill_store_sound' := kill_store_sound fundef unit ge. Lemma sem_rel_b_ge: forall rb1 rb2 : RB.t, @@ -1260,8 +1261,7 @@ Proof. rewrite H. reflexivity. } - apply (kill_store_sound' sp m). - assumption. + eapply (kill_store_sound' sp m); eassumption. (* call *) - econstructor; split. diff --git a/test/cse2/indexed_addr.c b/test/cse2/indexed_addr.c new file mode 100644 index 00000000..30a7c571 --- /dev/null +++ b/test/cse2/indexed_addr.c @@ -0,0 +1,6 @@ +void foo(int *t) { + if (t[0] > 4) { + t[1] ++; + t[0] --; + } +} -- cgit From c4f88ed5581ffb71e7ed5824c7503e8ce08165df Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 3 Mar 2020 08:58:01 +0100 Subject: globals alias analysis for x86 --- backend/CSE2.v | 2 ++ backend/CSE2proof.v | 67 +++++++++++++++++++++++++++++++++++++++++++++-------- test/cse2/globals.c | 8 +++++++ 3 files changed, 67 insertions(+), 10 deletions(-) create mode 100644 test/cse2/globals.c diff --git a/backend/CSE2.v b/backend/CSE2.v index f5ff8748..5b0556aa 100644 --- a/backend/CSE2.v +++ b/backend/CSE2.v @@ -288,6 +288,8 @@ Definition may_overlap chunk addr args chunk' addr' args' := if peq base base' then negb (can_swap_accesses_ofs ofs' chunk' ofs chunk) else true + | (Aglobal symb ofs), (Aglobal symb' ofs'), + nil, nil => peq symb symb' | _, _, _, _ => true end. diff --git a/backend/CSE2proof.v b/backend/CSE2proof.v index e65d9194..ada0eb00 100644 --- a/backend/CSE2proof.v +++ b/backend/CSE2proof.v @@ -716,7 +716,6 @@ Section MEMORY_WRITE. Variable m m2 : mem. Variable chunkw chunkr : memory_chunk. Variable base : val. - Variable ofsw ofsr : Z. Lemma size_chunk_bounded : forall chunk, 0 <= size_chunk chunk <= max_chunk_size. @@ -726,13 +725,15 @@ Section MEMORY_WRITE. Qed. Variable addrw addrr valw valr : val. - + Hypothesis STORE : Mem.storev chunkw m addrw valw = Some m2. + Hypothesis READ : Mem.loadv chunkr m addrr = Some valr. + + Section INDEXED_AWAY. + Variable ofsw ofsr : Z. Hypothesis ADDRW : eval_addressing genv sp (Aindexed ofsw) (base :: nil) = Some addrw. Hypothesis ADDRR : eval_addressing genv sp (Aindexed ofsr) (base :: nil) = Some addrr. - Hypothesis STORE : Mem.storev chunkw m addrw valw = Some m2. - Hypothesis READ : Mem.loadv chunkr m addrr = Some valr. Lemma load_store_away1 : forall RANGEW : 0 <= ofsw <= Ptrofs.modulus - max_chunk_size, @@ -800,6 +801,47 @@ Section MEMORY_WRITE. apply load_store_away1. all: tauto. Qed. + End INDEXED_AWAY. + + Section DIFFERENT_GLOBALS. + Variable ofsw ofsr : ptrofs. + Hypothesis symw symr : ident. + Hypothesis ADDRW : eval_addressing genv sp + (Aglobal symw ofsw) nil = Some addrw. + Hypothesis ADDRR : eval_addressing genv sp + (Aglobal symr ofsr) nil = Some addrr. + + Lemma load_store_diff_globals : + symw <> symr -> + Mem.loadv chunkr m2 addrr = Some valr. + Proof. + intros. + unfold eval_addressing in *. + destruct Archi.ptr64 eqn:PTR64; simpl in *. + rewrite PTR64 in *. + 2: simpl in *; discriminate. + unfold Genv.symbol_address in *. + unfold Genv.find_symbol in *. + destruct ((Genv.genv_symb genv) ! symw) as [bw |] eqn:SYMW; inv ADDRW. + 2: simpl in STORE; discriminate. + destruct ((Genv.genv_symb genv) ! symr) as [br |] eqn:SYMR; inv ADDRR. + 2: simpl in READ; discriminate. + assert (br <> bw). + { + intro EQ. + subst br. + assert (symr = symw). + { + eapply Genv.genv_vars_inj; eauto. + } + congruence. + } + rewrite <- READ. + eapply Mem.load_store_other with (chunk := chunkw) (v := valw) (b := bw). + - exact STORE. + - left. assumption. + Qed. + End DIFFERENT_GLOBALS. End MEMORY_WRITE. Lemma kill_mem_sound : @@ -848,6 +890,7 @@ Proof. apply op_depends_on_memory_correct; auto. } destruct addr; destruct addr0; simpl; trivial. + { (* Aindexed / Aindexed *) destruct args as [ | base [ | ]]; simpl; trivial. destruct args0 as [ | base0 [ | ]]; simpl; trivial. destruct (peq (forward_move rel base) base0); simpl; trivial. @@ -857,12 +900,16 @@ Proof. erewrite forward_move_rs in * by exact SEM. destruct (eval_addressing genv sp (Aindexed z0) ((rs # base) :: nil)) eqn:ADDR0; simpl; trivial. symmetry. - eapply load_store_away. - - exact ADDR. - - exact ADDR0. - - exact STORE. - - congruence. - - assumption. + eapply load_store_away; eauto. + } + { (* Aglobal / Aglobal *) + destruct args; destruct args0; simpl; trivial. + destruct (peq i i1); simpl; trivial. + simpl in *. + destruct (eval_addressing genv sp (Aglobal i1 i2) nil) eqn:ADDR1; simpl; trivial. + symmetry. + eapply load_store_diff_globals; eauto. + } Qed. End SOUNDNESS. diff --git a/test/cse2/globals.c b/test/cse2/globals.c new file mode 100644 index 00000000..c6dd59cd --- /dev/null +++ b/test/cse2/globals.c @@ -0,0 +1,8 @@ +int glob1, glob2; + +void toto() { + if (glob1 > 4) { + glob2 ++; + glob1 --; + } +} -- cgit From af7afc0a388986a94f21d2657cc13b823d456781 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 3 Mar 2020 09:39:03 +0100 Subject: offsets in globals for x86 --- backend/CSE2.v | 5 +++- backend/CSE2proof.v | 79 ++++++++++++++++++++++++++++++++++++++++++++++++++++- 2 files changed, 82 insertions(+), 2 deletions(-) diff --git a/backend/CSE2.v b/backend/CSE2.v index 5b0556aa..9c45b476 100644 --- a/backend/CSE2.v +++ b/backend/CSE2.v @@ -289,7 +289,10 @@ Definition may_overlap chunk addr args chunk' addr' args' := then negb (can_swap_accesses_ofs ofs' chunk' ofs chunk) else true | (Aglobal symb ofs), (Aglobal symb' ofs'), - nil, nil => peq symb symb' + nil, nil => + if peq symb symb' + then negb (can_swap_accesses_ofs (Ptrofs.unsigned ofs') chunk' (Ptrofs.unsigned ofs) chunk) + else false | _, _, _, _ => true end. diff --git a/backend/CSE2proof.v b/backend/CSE2proof.v index ada0eb00..5acffc71 100644 --- a/backend/CSE2proof.v +++ b/backend/CSE2proof.v @@ -842,6 +842,74 @@ Section MEMORY_WRITE. - left. assumption. Qed. End DIFFERENT_GLOBALS. + + Section SAME_GLOBALS. + Variable ofsw ofsr : ptrofs. + Hypothesis sym : ident. + Hypothesis ADDRW : eval_addressing genv sp + (Aglobal sym ofsw) nil = Some addrw. + Hypothesis ADDRR : eval_addressing genv sp + (Aglobal sym ofsr) nil = Some addrr. + + Lemma load_store_glob_away1 : + forall RANGEW : 0 <= (Ptrofs.unsigned ofsw) <= Ptrofs.modulus - max_chunk_size, + forall RANGER : 0 <= (Ptrofs.unsigned ofsr) <= Ptrofs.modulus - max_chunk_size, + forall SWAPPABLE : (Ptrofs.unsigned ofsw) + size_chunk chunkw <= (Ptrofs.unsigned ofsr) + \/ (Ptrofs.unsigned ofsr) + size_chunk chunkr <= (Ptrofs.unsigned ofsw), + Mem.loadv chunkr m2 addrr = Some valr. + Proof. + intros. + + pose proof (size_chunk_bounded chunkr) as size_chunkr_bounded. + pose proof (size_chunk_bounded chunkw) as size_chunkw_bounded. + unfold max_chunk_size in size_chunkr_bounded, size_chunkw_bounded. + try change (Ptrofs.modulus - max_chunk_size) with 4294967288 in *. + try change (Ptrofs.modulus - max_chunk_size) with 18446744073709551608 in *. + unfold eval_addressing, eval_addressing32, eval_addressing64 in *. + destruct Archi.ptr64 eqn:PTR64. + + { + unfold Genv.symbol_address in *. + inv ADDRR. + inv ADDRW. + destruct (Genv.find_symbol genv sym). + 2: discriminate. + + rewrite <- READ. + eapply Mem.load_store_other with (chunk := chunkw) (v := valw) (b := b). + exact STORE. + right. + tauto. + } + + { + unfold Genv.symbol_address in *. + inv ADDRR. + inv ADDRW. + destruct (Genv.find_symbol genv sym). + 2: discriminate. + + rewrite <- READ. + eapply Mem.load_store_other with (chunk := chunkw) (v := valw) (b := b). + exact STORE. + right. + tauto. + } + Qed. + + Lemma load_store_glob_away : + (can_swap_accesses_ofs (Ptrofs.unsigned ofsr) chunkr (Ptrofs.unsigned ofsw) chunkw) = true -> + Mem.loadv chunkr m2 addrr = Some valr. + Proof. + intro SWAP. + unfold can_swap_accesses_ofs in SWAP. + repeat rewrite andb_true_iff in SWAP. + repeat rewrite orb_true_iff in SWAP. + repeat rewrite Z.leb_le in SWAP. + apply load_store_glob_away1. + all: tauto. + Qed. + End SAME_GLOBALS. End MEMORY_WRITE. Lemma kill_mem_sound : @@ -904,7 +972,16 @@ Proof. } { (* Aglobal / Aglobal *) destruct args; destruct args0; simpl; trivial. - destruct (peq i i1); simpl; trivial. + destruct (peq i i1). + { + subst i1. + destruct (can_swap_accesses_ofs (Ptrofs.unsigned i2) chunk0 + (Ptrofs.unsigned i0) chunk) eqn:SWAP; simpl; trivial. + simpl in *. + destruct (eval_addressing genv sp (Aglobal i i2) nil) eqn:ADDR1; simpl; trivial. + symmetry. + eapply load_store_glob_away; eauto. + } simpl in *. destruct (eval_addressing genv sp (Aglobal i1 i2) nil) eqn:ADDR1; simpl; trivial. symmetry. -- cgit From b11cbccf7eb4a6696c5285cb0bcde57dd0c0447e Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 3 Mar 2020 10:41:33 +0100 Subject: starting to move x86 stuff to x86 --- backend/CSE2proof.v | 202 +----------------------------------------------- x86/CSE2depsproof.v | 215 ++++++++++++++++++++++++++++++++++++++++++++++++++++ 2 files changed, 216 insertions(+), 201 deletions(-) create mode 100644 x86/CSE2depsproof.v diff --git a/backend/CSE2proof.v b/backend/CSE2proof.v index 5acffc71..eeb9442f 100644 --- a/backend/CSE2proof.v +++ b/backend/CSE2proof.v @@ -13,7 +13,7 @@ Require Import Memory Registers Op RTL Maps. Require Import Globalenvs Values. Require Import Linking Values Memory Globalenvs Events Smallstep. Require Import Registers Op RTL. -Require Import CSE2. +Require Import CSE2 CSE2depsproof. Require Import Lia. Lemma args_unaffected: @@ -712,206 +712,6 @@ Proof. Qed. End SAME_MEMORY. -Section MEMORY_WRITE. - Variable m m2 : mem. - Variable chunkw chunkr : memory_chunk. - Variable base : val. - - Lemma size_chunk_bounded : - forall chunk, 0 <= size_chunk chunk <= max_chunk_size. - Proof. - unfold max_chunk_size. - destruct chunk; simpl; lia. - Qed. - - Variable addrw addrr valw valr : val. - Hypothesis STORE : Mem.storev chunkw m addrw valw = Some m2. - Hypothesis READ : Mem.loadv chunkr m addrr = Some valr. - - Section INDEXED_AWAY. - Variable ofsw ofsr : Z. - Hypothesis ADDRW : eval_addressing genv sp - (Aindexed ofsw) (base :: nil) = Some addrw. - Hypothesis ADDRR : eval_addressing genv sp - (Aindexed ofsr) (base :: nil) = Some addrr. - - Lemma load_store_away1 : - forall RANGEW : 0 <= ofsw <= Ptrofs.modulus - max_chunk_size, - forall RANGER : 0 <= ofsr <= Ptrofs.modulus - max_chunk_size, - forall SWAPPABLE : ofsw + size_chunk chunkw <= ofsr - \/ ofsr + size_chunk chunkr <= ofsw, - Mem.loadv chunkr m2 addrr = Some valr. - Proof. - intros. - - pose proof (size_chunk_bounded chunkr) as size_chunkr_bounded. - pose proof (size_chunk_bounded chunkw) as size_chunkw_bounded. - unfold max_chunk_size in size_chunkr_bounded, size_chunkw_bounded. - try change (Ptrofs.modulus - max_chunk_size) with 4294967288 in *. - try change (Ptrofs.modulus - max_chunk_size) with 18446744073709551608 in *. - destruct addrr ; simpl in * ; try discriminate. - unfold eval_addressing in *. - destruct Archi.ptr64 eqn:PTR64; destruct base; simpl in *; try discriminate. - rewrite PTR64 in *. - - inv ADDRR. - inv ADDRW. - rewrite <- READ. - eapply Mem.load_store_other with (chunk := chunkw) (v := valw) (b := b). - exact STORE. - right. - - all: try (destruct (Ptrofs.unsigned_add_either i0 - (Ptrofs.of_int (Int.repr ofsr))) as [OFSR | OFSR]; - rewrite OFSR). - all: try (destruct (Ptrofs.unsigned_add_either i0 - (Ptrofs.of_int64 (Int64.repr ofsr))) as [OFSR | OFSR]; - rewrite OFSR). - all: try (destruct (Ptrofs.unsigned_add_either i0 - (Ptrofs.of_int (Int.repr ofsw))) as [OFSW | OFSW]; - rewrite OFSW). - all: try (destruct (Ptrofs.unsigned_add_either i0 - (Ptrofs.of_int64 (Int64.repr ofsw))) as [OFSW | OFSW]; - rewrite OFSW). - - all: unfold Ptrofs.of_int64. - all: unfold Ptrofs.of_int. - - - all: repeat rewrite Int.unsigned_repr by (change Int.max_unsigned with 4294967295; lia). - all: repeat rewrite Ptrofs.unsigned_repr by (change Ptrofs.max_unsigned with 4294967295; lia). - all: repeat rewrite Int64.unsigned_repr by (change Int64.max_unsigned with 18446744073709551615; lia). - all: repeat rewrite Ptrofs.unsigned_repr by (change Ptrofs.max_unsigned with 18446744073709551615; lia). - - all: try change Ptrofs.modulus with 4294967296. - all: try change Ptrofs.modulus with 18446744073709551616. - - all: intuition lia. - Qed. - - Theorem load_store_away : - can_swap_accesses_ofs ofsr chunkr ofsw chunkw = true -> - Mem.loadv chunkr m2 addrr = Some valr. - Proof. - intro SWAP. - unfold can_swap_accesses_ofs in SWAP. - repeat rewrite andb_true_iff in SWAP. - repeat rewrite orb_true_iff in SWAP. - repeat rewrite Z.leb_le in SWAP. - apply load_store_away1. - all: tauto. - Qed. - End INDEXED_AWAY. - - Section DIFFERENT_GLOBALS. - Variable ofsw ofsr : ptrofs. - Hypothesis symw symr : ident. - Hypothesis ADDRW : eval_addressing genv sp - (Aglobal symw ofsw) nil = Some addrw. - Hypothesis ADDRR : eval_addressing genv sp - (Aglobal symr ofsr) nil = Some addrr. - - Lemma load_store_diff_globals : - symw <> symr -> - Mem.loadv chunkr m2 addrr = Some valr. - Proof. - intros. - unfold eval_addressing in *. - destruct Archi.ptr64 eqn:PTR64; simpl in *. - rewrite PTR64 in *. - 2: simpl in *; discriminate. - unfold Genv.symbol_address in *. - unfold Genv.find_symbol in *. - destruct ((Genv.genv_symb genv) ! symw) as [bw |] eqn:SYMW; inv ADDRW. - 2: simpl in STORE; discriminate. - destruct ((Genv.genv_symb genv) ! symr) as [br |] eqn:SYMR; inv ADDRR. - 2: simpl in READ; discriminate. - assert (br <> bw). - { - intro EQ. - subst br. - assert (symr = symw). - { - eapply Genv.genv_vars_inj; eauto. - } - congruence. - } - rewrite <- READ. - eapply Mem.load_store_other with (chunk := chunkw) (v := valw) (b := bw). - - exact STORE. - - left. assumption. - Qed. - End DIFFERENT_GLOBALS. - - Section SAME_GLOBALS. - Variable ofsw ofsr : ptrofs. - Hypothesis sym : ident. - Hypothesis ADDRW : eval_addressing genv sp - (Aglobal sym ofsw) nil = Some addrw. - Hypothesis ADDRR : eval_addressing genv sp - (Aglobal sym ofsr) nil = Some addrr. - - Lemma load_store_glob_away1 : - forall RANGEW : 0 <= (Ptrofs.unsigned ofsw) <= Ptrofs.modulus - max_chunk_size, - forall RANGER : 0 <= (Ptrofs.unsigned ofsr) <= Ptrofs.modulus - max_chunk_size, - forall SWAPPABLE : (Ptrofs.unsigned ofsw) + size_chunk chunkw <= (Ptrofs.unsigned ofsr) - \/ (Ptrofs.unsigned ofsr) + size_chunk chunkr <= (Ptrofs.unsigned ofsw), - Mem.loadv chunkr m2 addrr = Some valr. - Proof. - intros. - - pose proof (size_chunk_bounded chunkr) as size_chunkr_bounded. - pose proof (size_chunk_bounded chunkw) as size_chunkw_bounded. - unfold max_chunk_size in size_chunkr_bounded, size_chunkw_bounded. - try change (Ptrofs.modulus - max_chunk_size) with 4294967288 in *. - try change (Ptrofs.modulus - max_chunk_size) with 18446744073709551608 in *. - unfold eval_addressing, eval_addressing32, eval_addressing64 in *. - destruct Archi.ptr64 eqn:PTR64. - - { - unfold Genv.symbol_address in *. - inv ADDRR. - inv ADDRW. - destruct (Genv.find_symbol genv sym). - 2: discriminate. - - rewrite <- READ. - eapply Mem.load_store_other with (chunk := chunkw) (v := valw) (b := b). - exact STORE. - right. - tauto. - } - - { - unfold Genv.symbol_address in *. - inv ADDRR. - inv ADDRW. - destruct (Genv.find_symbol genv sym). - 2: discriminate. - - rewrite <- READ. - eapply Mem.load_store_other with (chunk := chunkw) (v := valw) (b := b). - exact STORE. - right. - tauto. - } - Qed. - - Lemma load_store_glob_away : - (can_swap_accesses_ofs (Ptrofs.unsigned ofsr) chunkr (Ptrofs.unsigned ofsw) chunkw) = true -> - Mem.loadv chunkr m2 addrr = Some valr. - Proof. - intro SWAP. - unfold can_swap_accesses_ofs in SWAP. - repeat rewrite andb_true_iff in SWAP. - repeat rewrite orb_true_iff in SWAP. - repeat rewrite Z.leb_le in SWAP. - apply load_store_glob_away1. - all: tauto. - Qed. - End SAME_GLOBALS. -End MEMORY_WRITE. - Lemma kill_mem_sound : forall m m' : mem, forall rel : RELATION.t, diff --git a/x86/CSE2depsproof.v b/x86/CSE2depsproof.v new file mode 100644 index 00000000..f4eace6f --- /dev/null +++ b/x86/CSE2depsproof.v @@ -0,0 +1,215 @@ +Require Import Coqlib Maps Errors Integers Floats Lattice Kildall. +Require Import AST Linking. +Require Import Memory Registers Op RTL Maps. + +Require Import Globalenvs Values. +Require Import Linking Values Memory Globalenvs Events Smallstep. +Require Import Registers Op RTL. +Require Import CSE2. +Require Import Lia. + +Section SOUNDNESS. + Variable F V : Type. + Variable genv: Genv.t F V. + Variable sp : val. + +Section MEMORY_WRITE. + Variable m m2 : mem. + Variable chunkw chunkr : memory_chunk. + Variable base : val. + + Lemma size_chunk_bounded : + forall chunk, 0 <= size_chunk chunk <= max_chunk_size. + Proof. + unfold max_chunk_size. + destruct chunk; simpl; lia. + Qed. + + Variable addrw addrr valw valr : val. + Hypothesis STORE : Mem.storev chunkw m addrw valw = Some m2. + Hypothesis READ : Mem.loadv chunkr m addrr = Some valr. + + Section INDEXED_AWAY. + Variable ofsw ofsr : Z. + Hypothesis ADDRW : eval_addressing genv sp + (Aindexed ofsw) (base :: nil) = Some addrw. + Hypothesis ADDRR : eval_addressing genv sp + (Aindexed ofsr) (base :: nil) = Some addrr. + + Lemma load_store_away1 : + forall RANGEW : 0 <= ofsw <= Ptrofs.modulus - max_chunk_size, + forall RANGER : 0 <= ofsr <= Ptrofs.modulus - max_chunk_size, + forall SWAPPABLE : ofsw + size_chunk chunkw <= ofsr + \/ ofsr + size_chunk chunkr <= ofsw, + Mem.loadv chunkr m2 addrr = Some valr. + Proof. + intros. + + pose proof (size_chunk_bounded chunkr) as size_chunkr_bounded. + pose proof (size_chunk_bounded chunkw) as size_chunkw_bounded. + unfold max_chunk_size in size_chunkr_bounded, size_chunkw_bounded. + try change (Ptrofs.modulus - max_chunk_size) with 4294967288 in *. + try change (Ptrofs.modulus - max_chunk_size) with 18446744073709551608 in *. + destruct addrr ; simpl in * ; try discriminate. + unfold eval_addressing in *. + destruct Archi.ptr64 eqn:PTR64; destruct base; simpl in *; try discriminate. + rewrite PTR64 in *. + + inv ADDRR. + inv ADDRW. + rewrite <- READ. + eapply Mem.load_store_other with (chunk := chunkw) (v := valw) (b := b). + exact STORE. + right. + + all: try (destruct (Ptrofs.unsigned_add_either i0 + (Ptrofs.of_int (Int.repr ofsr))) as [OFSR | OFSR]; + rewrite OFSR). + all: try (destruct (Ptrofs.unsigned_add_either i0 + (Ptrofs.of_int64 (Int64.repr ofsr))) as [OFSR | OFSR]; + rewrite OFSR). + all: try (destruct (Ptrofs.unsigned_add_either i0 + (Ptrofs.of_int (Int.repr ofsw))) as [OFSW | OFSW]; + rewrite OFSW). + all: try (destruct (Ptrofs.unsigned_add_either i0 + (Ptrofs.of_int64 (Int64.repr ofsw))) as [OFSW | OFSW]; + rewrite OFSW). + + all: unfold Ptrofs.of_int64. + all: unfold Ptrofs.of_int. + + + all: repeat rewrite Int.unsigned_repr by (change Int.max_unsigned with 4294967295; lia). + all: repeat rewrite Ptrofs.unsigned_repr by (change Ptrofs.max_unsigned with 4294967295; lia). + all: repeat rewrite Int64.unsigned_repr by (change Int64.max_unsigned with 18446744073709551615; lia). + all: repeat rewrite Ptrofs.unsigned_repr by (change Ptrofs.max_unsigned with 18446744073709551615; lia). + + all: try change Ptrofs.modulus with 4294967296. + all: try change Ptrofs.modulus with 18446744073709551616. + + all: intuition lia. + Qed. + + Theorem load_store_away : + can_swap_accesses_ofs ofsr chunkr ofsw chunkw = true -> + Mem.loadv chunkr m2 addrr = Some valr. + Proof. + intro SWAP. + unfold can_swap_accesses_ofs in SWAP. + repeat rewrite andb_true_iff in SWAP. + repeat rewrite orb_true_iff in SWAP. + repeat rewrite Z.leb_le in SWAP. + apply load_store_away1. + all: tauto. + Qed. + End INDEXED_AWAY. + + Section DIFFERENT_GLOBALS. + Variable ofsw ofsr : ptrofs. + Hypothesis symw symr : ident. + Hypothesis ADDRW : eval_addressing genv sp + (Aglobal symw ofsw) nil = Some addrw. + Hypothesis ADDRR : eval_addressing genv sp + (Aglobal symr ofsr) nil = Some addrr. + + Lemma load_store_diff_globals : + symw <> symr -> + Mem.loadv chunkr m2 addrr = Some valr. + Proof. + intros. + unfold eval_addressing in *. + destruct Archi.ptr64 eqn:PTR64; simpl in *. + rewrite PTR64 in *. + 2: simpl in *; discriminate. + unfold Genv.symbol_address in *. + unfold Genv.find_symbol in *. + destruct ((Genv.genv_symb genv) ! symw) as [bw |] eqn:SYMW; inv ADDRW. + 2: simpl in STORE; discriminate. + destruct ((Genv.genv_symb genv) ! symr) as [br |] eqn:SYMR; inv ADDRR. + 2: simpl in READ; discriminate. + assert (br <> bw). + { + intro EQ. + subst br. + assert (symr = symw). + { + eapply Genv.genv_vars_inj; eauto. + } + congruence. + } + rewrite <- READ. + eapply Mem.load_store_other with (chunk := chunkw) (v := valw) (b := bw). + - exact STORE. + - left. assumption. + Qed. + End DIFFERENT_GLOBALS. + + Section SAME_GLOBALS. + Variable ofsw ofsr : ptrofs. + Hypothesis sym : ident. + Hypothesis ADDRW : eval_addressing genv sp + (Aglobal sym ofsw) nil = Some addrw. + Hypothesis ADDRR : eval_addressing genv sp + (Aglobal sym ofsr) nil = Some addrr. + + Lemma load_store_glob_away1 : + forall RANGEW : 0 <= (Ptrofs.unsigned ofsw) <= Ptrofs.modulus - max_chunk_size, + forall RANGER : 0 <= (Ptrofs.unsigned ofsr) <= Ptrofs.modulus - max_chunk_size, + forall SWAPPABLE : (Ptrofs.unsigned ofsw) + size_chunk chunkw <= (Ptrofs.unsigned ofsr) + \/ (Ptrofs.unsigned ofsr) + size_chunk chunkr <= (Ptrofs.unsigned ofsw), + Mem.loadv chunkr m2 addrr = Some valr. + Proof. + intros. + + pose proof (size_chunk_bounded chunkr) as size_chunkr_bounded. + pose proof (size_chunk_bounded chunkw) as size_chunkw_bounded. + unfold max_chunk_size in size_chunkr_bounded, size_chunkw_bounded. + try change (Ptrofs.modulus - max_chunk_size) with 4294967288 in *. + try change (Ptrofs.modulus - max_chunk_size) with 18446744073709551608 in *. + unfold eval_addressing, eval_addressing32, eval_addressing64 in *. + destruct Archi.ptr64 eqn:PTR64. + + { + unfold Genv.symbol_address in *. + inv ADDRR. + inv ADDRW. + destruct (Genv.find_symbol genv sym). + 2: discriminate. + + rewrite <- READ. + eapply Mem.load_store_other with (chunk := chunkw) (v := valw) (b := b). + exact STORE. + right. + tauto. + } + + { + unfold Genv.symbol_address in *. + inv ADDRR. + inv ADDRW. + destruct (Genv.find_symbol genv sym). + 2: discriminate. + + rewrite <- READ. + eapply Mem.load_store_other with (chunk := chunkw) (v := valw) (b := b). + exact STORE. + right. + tauto. + } + Qed. + + Lemma load_store_glob_away : + (can_swap_accesses_ofs (Ptrofs.unsigned ofsr) chunkr (Ptrofs.unsigned ofsw) chunkw) = true -> + Mem.loadv chunkr m2 addrr = Some valr. + Proof. + intro SWAP. + unfold can_swap_accesses_ofs in SWAP. + repeat rewrite andb_true_iff in SWAP. + repeat rewrite orb_true_iff in SWAP. + repeat rewrite Z.leb_le in SWAP. + apply load_store_glob_away1. + all: tauto. + Qed. + End SAME_GLOBALS. +End MEMORY_WRITE. +End SOUNDNESS. -- cgit From 44811b4917b69e9a33efe5ab75ceb3b01f6594fc Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 3 Mar 2020 11:21:27 +0100 Subject: may_overlap_sound --- backend/CSE2proof.v | 38 ++++++++++++++++++++++++++++++++++++++ 1 file changed, 38 insertions(+) diff --git a/backend/CSE2proof.v b/backend/CSE2proof.v index eeb9442f..206dbf30 100644 --- a/backend/CSE2proof.v +++ b/backend/CSE2proof.v @@ -735,6 +735,44 @@ Proof. } Qed. +Lemma may_overlap_sound: + forall m m' : mem, + forall chunk addr args chunk' addr' args' v a a' vl rs, + (eval_addressing genv sp addr (rs ## args)) = Some a -> + (eval_addressing genv sp addr' (rs ## args')) = Some a' -> + (may_overlap chunk addr args chunk' addr' args') = false -> + (Mem.storev chunk m a v) = Some m' -> + (Mem.loadv chunk' m a') = Some vl -> + (Mem.loadv chunk' m' a') = Some vl. +Proof. + intros until rs. + intros ADDR ADDR' OVERLAP STORE LOAD. + destruct addr; destruct addr'; try discriminate. + { (* Aindexed / Aindexed *) + destruct args as [ | base [ | ]]. 1,3: discriminate. + destruct args' as [ | base' [ | ]]. 1,3: discriminate. + simpl in OVERLAP. + destruct (peq base base'). 2: discriminate. + subst base'. + destruct (can_swap_accesses_ofs z0 chunk' z chunk) eqn:SWAP. + 2: discriminate. + simpl in *. + eapply load_store_away; eassumption. + } + { (* Aglobal / Aglobal *) + destruct args. 2: discriminate. + destruct args'. 2: discriminate. + simpl in *. + destruct (peq i i1). + { + subst i1. + rewrite negb_false_iff in OVERLAP. + eapply load_store_glob_away; eassumption. + } + eapply load_store_diff_globals; eassumption. + } +Qed. + Lemma kill_store_sound : forall m m' : mem, forall rel : RELATION.t, -- cgit From 3fc937ddc8f82525081bca67818ca87f448f618e Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 3 Mar 2020 11:27:59 +0100 Subject: modularize proof --- backend/CSE2proof.v | 39 ++++++++++----------------------------- 1 file changed, 10 insertions(+), 29 deletions(-) diff --git a/backend/CSE2proof.v b/backend/CSE2proof.v index 206dbf30..90179f82 100644 --- a/backend/CSE2proof.v +++ b/backend/CSE2proof.v @@ -795,36 +795,17 @@ Proof. rewrite SEMx. apply op_depends_on_memory_correct; auto. } - destruct addr; destruct addr0; simpl; trivial. - { (* Aindexed / Aindexed *) - destruct args as [ | base [ | ]]; simpl; trivial. - destruct args0 as [ | base0 [ | ]]; simpl; trivial. - destruct (peq (forward_move rel base) base0); simpl; trivial. - subst base0. - destruct (can_swap_accesses_ofs z0 chunk0 z chunk) eqn:SWAP; simpl; trivial. - simpl in *. - erewrite forward_move_rs in * by exact SEM. - destruct (eval_addressing genv sp (Aindexed z0) ((rs # base) :: nil)) eqn:ADDR0; simpl; trivial. + destruct may_overlap eqn:OVERLAP; simpl; trivial. + destruct (eval_addressing genv sp addr0 rs ## args0) eqn:ADDR0. symmetry. - eapply load_store_away; eauto. - } - { (* Aglobal / Aglobal *) - destruct args; destruct args0; simpl; trivial. - destruct (peq i i1). - { - subst i1. - destruct (can_swap_accesses_ofs (Ptrofs.unsigned i2) chunk0 - (Ptrofs.unsigned i0) chunk) eqn:SWAP; simpl; trivial. - simpl in *. - destruct (eval_addressing genv sp (Aglobal i i2) nil) eqn:ADDR1; simpl; trivial. - symmetry. - eapply load_store_glob_away; eauto. - } - simpl in *. - destruct (eval_addressing genv sp (Aglobal i1 i2) nil) eqn:ADDR1; simpl; trivial. - symmetry. - eapply load_store_diff_globals; eauto. - } + eapply may_overlap_sound with (args := (map (forward_move rel) args)). + erewrite forward_move_map by eassumption. + exact ADDR. + exact ADDR0. + exact OVERLAP. + exact STORE. + symmetry; assumption. + assumption. Qed. End SOUNDNESS. -- cgit From 577d3dbeb54aaf23db19dddf149c48764e20c58d Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 3 Mar 2020 11:30:23 +0100 Subject: moved away x86-dependent parts --- backend/CSE2proof.v | 38 -------------------------------------- x86/CSE2depsproof.v | 46 ++++++++++++++++++++++++++++++++++++++++++++++ 2 files changed, 46 insertions(+), 38 deletions(-) diff --git a/backend/CSE2proof.v b/backend/CSE2proof.v index 90179f82..8cc87275 100644 --- a/backend/CSE2proof.v +++ b/backend/CSE2proof.v @@ -734,44 +734,6 @@ Proof. apply op_depends_on_memory_correct; auto. } Qed. - -Lemma may_overlap_sound: - forall m m' : mem, - forall chunk addr args chunk' addr' args' v a a' vl rs, - (eval_addressing genv sp addr (rs ## args)) = Some a -> - (eval_addressing genv sp addr' (rs ## args')) = Some a' -> - (may_overlap chunk addr args chunk' addr' args') = false -> - (Mem.storev chunk m a v) = Some m' -> - (Mem.loadv chunk' m a') = Some vl -> - (Mem.loadv chunk' m' a') = Some vl. -Proof. - intros until rs. - intros ADDR ADDR' OVERLAP STORE LOAD. - destruct addr; destruct addr'; try discriminate. - { (* Aindexed / Aindexed *) - destruct args as [ | base [ | ]]. 1,3: discriminate. - destruct args' as [ | base' [ | ]]. 1,3: discriminate. - simpl in OVERLAP. - destruct (peq base base'). 2: discriminate. - subst base'. - destruct (can_swap_accesses_ofs z0 chunk' z chunk) eqn:SWAP. - 2: discriminate. - simpl in *. - eapply load_store_away; eassumption. - } - { (* Aglobal / Aglobal *) - destruct args. 2: discriminate. - destruct args'. 2: discriminate. - simpl in *. - destruct (peq i i1). - { - subst i1. - rewrite negb_false_iff in OVERLAP. - eapply load_store_glob_away; eassumption. - } - eapply load_store_diff_globals; eassumption. - } -Qed. Lemma kill_store_sound : forall m m' : mem, diff --git a/x86/CSE2depsproof.v b/x86/CSE2depsproof.v index f4eace6f..84b22c69 100644 --- a/x86/CSE2depsproof.v +++ b/x86/CSE2depsproof.v @@ -213,3 +213,49 @@ Section MEMORY_WRITE. End SAME_GLOBALS. End MEMORY_WRITE. End SOUNDNESS. + + +Section SOUNDNESS. + Variable F V : Type. + Variable genv: Genv.t F V. + Variable sp : val. + +Lemma may_overlap_sound: + forall m m' : mem, + forall chunk addr args chunk' addr' args' v a a' vl rs, + (eval_addressing genv sp addr (rs ## args)) = Some a -> + (eval_addressing genv sp addr' (rs ## args')) = Some a' -> + (may_overlap chunk addr args chunk' addr' args') = false -> + (Mem.storev chunk m a v) = Some m' -> + (Mem.loadv chunk' m a') = Some vl -> + (Mem.loadv chunk' m' a') = Some vl. +Proof. + intros until rs. + intros ADDR ADDR' OVERLAP STORE LOAD. + destruct addr; destruct addr'; try discriminate. + { (* Aindexed / Aindexed *) + destruct args as [ | base [ | ]]. 1,3: discriminate. + destruct args' as [ | base' [ | ]]. 1,3: discriminate. + simpl in OVERLAP. + destruct (peq base base'). 2: discriminate. + subst base'. + destruct (can_swap_accesses_ofs z0 chunk' z chunk) eqn:SWAP. + 2: discriminate. + simpl in *. + eapply load_store_away; eassumption. + } + { (* Aglobal / Aglobal *) + destruct args. 2: discriminate. + destruct args'. 2: discriminate. + simpl in *. + destruct (peq i i1). + { + subst i1. + rewrite negb_false_iff in OVERLAP. + eapply load_store_glob_away; eassumption. + } + eapply load_store_diff_globals; eassumption. + } +Qed. + +End SOUNDNESS. -- cgit From 74efac13484e4767f793cf9ccc94835825ca30ba Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 3 Mar 2020 12:27:58 +0100 Subject: CSE2 alias analysis for Risc-V --- backend/CSE2.v | 25 +--------- backend/CSE2proof.v | 2 +- backend/ValueDomain.v | 5 -- common/Memdata.v | 7 +++ riscV/CSE2deps.v | 20 ++++++++ riscV/CSE2depsproof.v | 129 ++++++++++++++++++++++++++++++++++++++++++++++++++ 6 files changed, 158 insertions(+), 30 deletions(-) create mode 100644 riscV/CSE2deps.v create mode 100644 riscV/CSE2depsproof.v diff --git a/backend/CSE2.v b/backend/CSE2.v index 9c45b476..8142ee5d 100644 --- a/backend/CSE2.v +++ b/backend/CSE2.v @@ -6,7 +6,7 @@ David Monniaux, CNRS, VERIMAG Require Import Coqlib Maps Errors Integers Floats Lattice Kildall. Require Import AST Linking. -Require Import Memory Registers Op RTL Maps. +Require Import Memory Registers Op RTL Maps CSE2deps. (* Static analysis *) @@ -265,14 +265,6 @@ Definition kill_sym_val (dst : reg) (sv : sym_val) := Definition kill_reg (dst : reg) (rel : RELATION.t) := PTree.filter1 (fun x => negb (kill_sym_val dst x)) (PTree.remove dst rel). - -Definition max_chunk_size := 8. - -Definition can_swap_accesses_ofs ofsr chunkr ofsw chunkw := - (0 <=? ofsw) && (ofsw <=? (Ptrofs.modulus - max_chunk_size)) - && (0 <=? ofsr) && (ofsr <=? (Ptrofs.modulus - max_chunk_size)) - && ((ofsw + size_chunk chunkw <=? ofsr) || - (ofsr + size_chunk chunkr <=? ofsw)). Definition kill_sym_val_mem (sv: sym_val) := match sv with @@ -281,21 +273,6 @@ Definition kill_sym_val_mem (sv: sym_val) := | SLoad _ _ _ => true end. -Definition may_overlap chunk addr args chunk' addr' args' := - match addr, addr', args, args' with - | (Aindexed ofs), (Aindexed ofs'), - (base :: nil), (base' :: nil) => - if peq base base' - then negb (can_swap_accesses_ofs ofs' chunk' ofs chunk) - else true - | (Aglobal symb ofs), (Aglobal symb' ofs'), - nil, nil => - if peq symb symb' - then negb (can_swap_accesses_ofs (Ptrofs.unsigned ofs') chunk' (Ptrofs.unsigned ofs) chunk) - else false - | _, _, _, _ => true - end. - Definition kill_sym_val_store chunk addr args (sv: sym_val) := match sv with | SMove _ => false diff --git a/backend/CSE2proof.v b/backend/CSE2proof.v index 8cc87275..09490c4d 100644 --- a/backend/CSE2proof.v +++ b/backend/CSE2proof.v @@ -13,7 +13,7 @@ Require Import Memory Registers Op RTL Maps. Require Import Globalenvs Values. Require Import Linking Values Memory Globalenvs Events Smallstep. Require Import Registers Op RTL. -Require Import CSE2 CSE2depsproof. +Require Import CSE2 CSE2deps CSE2depsproof. Require Import Lia. Lemma args_unaffected: diff --git a/backend/ValueDomain.v b/backend/ValueDomain.v index c132ce7c..779e7bb9 100644 --- a/backend/ValueDomain.v +++ b/backend/ValueDomain.v @@ -3502,11 +3502,6 @@ Proof. - omegaContradiction. Qed. -Lemma max_size_chunk: forall chunk, size_chunk chunk <= 8. -Proof. - destruct chunk; simpl; omega. -Qed. - Remark inval_before_contents: forall i c chunk' av' j, (inval_before i (i - 7) c)##j = Some (ACval chunk' av') -> diff --git a/common/Memdata.v b/common/Memdata.v index f3016efe..a09b90f5 100644 --- a/common/Memdata.v +++ b/common/Memdata.v @@ -44,6 +44,13 @@ Definition size_chunk (chunk: memory_chunk) : Z := | Many64 => 8 end. +Definition largest_size_chunk := 8. + +Lemma max_size_chunk: forall chunk, size_chunk chunk <= 8. +Proof. + destruct chunk; simpl; omega. +Qed. + Lemma size_chunk_pos: forall chunk, size_chunk chunk > 0. Proof. diff --git a/riscV/CSE2deps.v b/riscV/CSE2deps.v new file mode 100644 index 00000000..8ab9242a --- /dev/null +++ b/riscV/CSE2deps.v @@ -0,0 +1,20 @@ +Require Import BoolEqual Coqlib. +Require Import AST Integers Floats. +Require Import Values Memory Globalenvs Events. +Require Import Op. + + +Definition can_swap_accesses_ofs ofsr chunkr ofsw chunkw := + (0 <=? ofsw) && (ofsw <=? (Ptrofs.modulus - largest_size_chunk)) + && (0 <=? ofsr) && (ofsr <=? (Ptrofs.modulus - largest_size_chunk)) + && ((ofsw + size_chunk chunkw <=? ofsr) || + (ofsr + size_chunk chunkr <=? ofsw)). + +Definition may_overlap chunk addr args chunk' addr' args' := + match addr, addr', args, args' with + | (Aindexed ofs), (Aindexed ofs'), + (base :: nil), (base' :: nil) => + if peq base base' + then negb (can_swap_accesses_ofs (Ptrofs.unsigned ofs') chunk' (Ptrofs.unsigned ofs) chunk) + else true | _, _, _, _ => true + end. diff --git a/riscV/CSE2depsproof.v b/riscV/CSE2depsproof.v new file mode 100644 index 00000000..ee500965 --- /dev/null +++ b/riscV/CSE2depsproof.v @@ -0,0 +1,129 @@ +Require Import Coqlib Maps Errors Integers Floats Lattice Kildall. +Require Import AST Linking. +Require Import Memory Registers Op RTL Maps. + +Require Import Globalenvs Values. +Require Import Linking Values Memory Globalenvs Events Smallstep. +Require Import Registers Op RTL. +Require Import CSE2 CSE2deps. +Require Import Lia. + +Lemma ptrofs_size : + Ptrofs.wordsize = (if Archi.ptr64 then 64 else 32)%nat. +Proof. + unfold Ptrofs.wordsize. + unfold Wordsize_Ptrofs.wordsize. + trivial. +Qed. + +Lemma ptrofs_modulus : + Ptrofs.modulus = if Archi.ptr64 then 18446744073709551616 else 4294967296. +Proof. + unfold Ptrofs.modulus. + rewrite ptrofs_size. + destruct Archi.ptr64; reflexivity. +Qed. + +Section SOUNDNESS. + Variable F V : Type. + Variable genv: Genv.t F V. + Variable sp : val. + +Section MEMORY_WRITE. + Variable m m2 : mem. + Variable chunkw chunkr : memory_chunk. + Variable base : val. + + Variable addrw addrr valw valr : val. + Hypothesis STORE : Mem.storev chunkw m addrw valw = Some m2. + Hypothesis READ : Mem.loadv chunkr m addrr = Some valr. + + Section INDEXED_AWAY. + Variable ofsw ofsr : ptrofs. + Hypothesis ADDRW : eval_addressing genv sp + (Aindexed ofsw) (base :: nil) = Some addrw. + Hypothesis ADDRR : eval_addressing genv sp + (Aindexed ofsr) (base :: nil) = Some addrr. + + Lemma load_store_away1 : + forall RANGEW : 0 <= Ptrofs.unsigned ofsw <= Ptrofs.modulus - largest_size_chunk, + forall RANGER : 0 <= Ptrofs.unsigned ofsr <= Ptrofs.modulus - largest_size_chunk, + forall SWAPPABLE : Ptrofs.unsigned ofsw + size_chunk chunkw <= Ptrofs.unsigned ofsr + \/ Ptrofs.unsigned ofsr + size_chunk chunkr <= Ptrofs.unsigned ofsw, + Mem.loadv chunkr m2 addrr = Some valr. + Proof. + intros. + + pose proof (max_size_chunk chunkr) as size_chunkr_bounded. + pose proof (max_size_chunk chunkw) as size_chunkw_bounded. + unfold largest_size_chunk in *. + + rewrite ptrofs_modulus in *. + simpl in *. + inv ADDRR. + inv ADDRW. + rewrite <- READ. + destruct base; try discriminate. + eapply Mem.load_store_other with (chunk := chunkw) (v := valw) (b := b). + exact STORE. + right. + + all: try (destruct (Ptrofs.unsigned_add_either i ofsr) as [OFSR | OFSR]; + rewrite OFSR). + all: try (destruct (Ptrofs.unsigned_add_either i ofsw) as [OFSW | OFSW]; + rewrite OFSW). + all: try rewrite ptrofs_modulus in *. + all: destruct Archi.ptr64. + + all: intuition lia. + Qed. + + Theorem load_store_away : + can_swap_accesses_ofs (Ptrofs.unsigned ofsr) chunkr (Ptrofs.unsigned ofsw) chunkw = true -> + Mem.loadv chunkr m2 addrr = Some valr. + Proof. + intro SWAP. + unfold can_swap_accesses_ofs in SWAP. + repeat rewrite andb_true_iff in SWAP. + repeat rewrite orb_true_iff in SWAP. + repeat rewrite Z.leb_le in SWAP. + apply load_store_away1. + all: tauto. + Qed. + End INDEXED_AWAY. +End MEMORY_WRITE. +End SOUNDNESS. + + +Section SOUNDNESS. + Variable F V : Type. + Variable genv: Genv.t F V. + Variable sp : val. + +Lemma may_overlap_sound: + forall m m' : mem, + forall chunk addr args chunk' addr' args' v a a' vl rs, + (eval_addressing genv sp addr (rs ## args)) = Some a -> + (eval_addressing genv sp addr' (rs ## args')) = Some a' -> + (may_overlap chunk addr args chunk' addr' args') = false -> + (Mem.storev chunk m a v) = Some m' -> + (Mem.loadv chunk' m a') = Some vl -> + (Mem.loadv chunk' m' a') = Some vl. +Proof. + intros until rs. + intros ADDR ADDR' OVERLAP STORE LOAD. + destruct addr; destruct addr'; try discriminate. + { (* Aindexed / Aindexed *) + destruct args as [ | base [ | ]]. 1,3: discriminate. + destruct args' as [ | base' [ | ]]. 1,3: discriminate. + simpl in OVERLAP. + destruct (peq base base'). 2: discriminate. + subst base'. + destruct (can_swap_accesses_ofs (Ptrofs.unsigned i0) chunk' (Ptrofs.unsigned i) chunk) eqn:SWAP. + 2: discriminate. + simpl in *. + eapply load_store_away with (F:=F) (V:=V) (genv:=genv) (sp:=sp); eassumption. + } +Qed. + +End SOUNDNESS. -- cgit From 4096e8c1b1e3d4fcdb44e81844d65a74f881aa47 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 3 Mar 2020 13:33:54 +0100 Subject: better 32/64-bit handling --- x86/CSE2deps.v | 24 ++++++++++++++++++++++++ x86/CSE2depsproof.v | 53 +++++++++++++++++++++++++++-------------------------- 2 files changed, 51 insertions(+), 26 deletions(-) create mode 100644 x86/CSE2deps.v diff --git a/x86/CSE2deps.v b/x86/CSE2deps.v new file mode 100644 index 00000000..f4d9e254 --- /dev/null +++ b/x86/CSE2deps.v @@ -0,0 +1,24 @@ +Require Import BoolEqual Coqlib. +Require Import AST Integers Floats. +Require Import Values Memory Globalenvs Events. +Require Import Op. + +Definition can_swap_accesses_ofs ofsr chunkr ofsw chunkw := + (0 <=? ofsw) && (ofsw <=? (Ptrofs.modulus - largest_size_chunk)) + && (0 <=? ofsr) && (ofsr <=? (Ptrofs.modulus - largest_size_chunk)) + && ((ofsw + size_chunk chunkw <=? ofsr) || + (ofsr + size_chunk chunkr <=? ofsw)). + +Definition may_overlap chunk addr args chunk' addr' args' := + match addr, addr', args, args' with + | (Aindexed ofs), (Aindexed ofs'), + (base :: nil), (base' :: nil) => + if peq base base' + then negb (can_swap_accesses_ofs ofs' chunk' ofs chunk) + else true + | (Aglobal symb ofs), (Aglobal symb' ofs'), nil, nil => + if peq symb symb' + then negb (can_swap_accesses_ofs (Ptrofs.unsigned ofs') chunk' (Ptrofs.unsigned ofs) chunk) + else false + | _, _, _, _ => true + end. diff --git a/x86/CSE2depsproof.v b/x86/CSE2depsproof.v index 84b22c69..37e16dfe 100644 --- a/x86/CSE2depsproof.v +++ b/x86/CSE2depsproof.v @@ -5,7 +5,7 @@ Require Import Memory Registers Op RTL Maps. Require Import Globalenvs Values. Require Import Linking Values Memory Globalenvs Events Smallstep. Require Import Registers Op RTL. -Require Import CSE2. +Require Import CSE2 CSE2deps. Require Import Lia. Section SOUNDNESS. @@ -17,13 +17,6 @@ Section MEMORY_WRITE. Variable m m2 : mem. Variable chunkw chunkr : memory_chunk. Variable base : val. - - Lemma size_chunk_bounded : - forall chunk, 0 <= size_chunk chunk <= max_chunk_size. - Proof. - unfold max_chunk_size. - destruct chunk; simpl; lia. - Qed. Variable addrw addrr valw valr : val. Hypothesis STORE : Mem.storev chunkw m addrw valw = Some m2. @@ -37,19 +30,18 @@ Section MEMORY_WRITE. (Aindexed ofsr) (base :: nil) = Some addrr. Lemma load_store_away1 : - forall RANGEW : 0 <= ofsw <= Ptrofs.modulus - max_chunk_size, - forall RANGER : 0 <= ofsr <= Ptrofs.modulus - max_chunk_size, + forall RANGEW : 0 <= ofsw <= Ptrofs.modulus - largest_size_chunk, + forall RANGER : 0 <= ofsr <= Ptrofs.modulus - largest_size_chunk, forall SWAPPABLE : ofsw + size_chunk chunkw <= ofsr \/ ofsr + size_chunk chunkr <= ofsw, Mem.loadv chunkr m2 addrr = Some valr. Proof. intros. - pose proof (size_chunk_bounded chunkr) as size_chunkr_bounded. - pose proof (size_chunk_bounded chunkw) as size_chunkw_bounded. - unfold max_chunk_size in size_chunkr_bounded, size_chunkw_bounded. - try change (Ptrofs.modulus - max_chunk_size) with 4294967288 in *. - try change (Ptrofs.modulus - max_chunk_size) with 18446744073709551608 in *. + pose proof (max_size_chunk chunkr) as size_chunkr_bounded. + pose proof (max_size_chunk chunkw) as size_chunkw_bounded. + try change (Ptrofs.modulus - largest_size_chunk) with 4294967288 in *. + try change (Ptrofs.modulus - largest_size_chunk) with 18446744073709551608 in *. destruct addrr ; simpl in * ; try discriminate. unfold eval_addressing in *. destruct Archi.ptr64 eqn:PTR64; destruct base; simpl in *; try discriminate. @@ -111,16 +103,25 @@ Section MEMORY_WRITE. (Aglobal symw ofsw) nil = Some addrw. Hypothesis ADDRR : eval_addressing genv sp (Aglobal symr ofsr) nil = Some addrr. - + + Lemma ptr64_cases: + forall {T : Type}, + forall b : bool, + forall x y : T, + (if b then (if b then x else y) else (if b then y else x)) = x. + Proof. + destruct b; reflexivity. + Qed. + Lemma load_store_diff_globals : symw <> symr -> Mem.loadv chunkr m2 addrr = Some valr. Proof. intros. unfold eval_addressing in *. - destruct Archi.ptr64 eqn:PTR64; simpl in *. - rewrite PTR64 in *. - 2: simpl in *; discriminate. + simpl in *. + rewrite ptr64_cases in ADDRR. + rewrite ptr64_cases in ADDRW. unfold Genv.symbol_address in *. unfold Genv.find_symbol in *. destruct ((Genv.genv_symb genv) ! symw) as [bw |] eqn:SYMW; inv ADDRW. @@ -153,19 +154,19 @@ Section MEMORY_WRITE. (Aglobal sym ofsr) nil = Some addrr. Lemma load_store_glob_away1 : - forall RANGEW : 0 <= (Ptrofs.unsigned ofsw) <= Ptrofs.modulus - max_chunk_size, - forall RANGER : 0 <= (Ptrofs.unsigned ofsr) <= Ptrofs.modulus - max_chunk_size, + forall RANGEW : 0 <= (Ptrofs.unsigned ofsw) <= Ptrofs.modulus - largest_size_chunk, + forall RANGER : 0 <= (Ptrofs.unsigned ofsr) <= Ptrofs.modulus - largest_size_chunk, forall SWAPPABLE : (Ptrofs.unsigned ofsw) + size_chunk chunkw <= (Ptrofs.unsigned ofsr) \/ (Ptrofs.unsigned ofsr) + size_chunk chunkr <= (Ptrofs.unsigned ofsw), Mem.loadv chunkr m2 addrr = Some valr. Proof. intros. - pose proof (size_chunk_bounded chunkr) as size_chunkr_bounded. - pose proof (size_chunk_bounded chunkw) as size_chunkw_bounded. - unfold max_chunk_size in size_chunkr_bounded, size_chunkw_bounded. - try change (Ptrofs.modulus - max_chunk_size) with 4294967288 in *. - try change (Ptrofs.modulus - max_chunk_size) with 18446744073709551608 in *. + pose proof (max_size_chunk chunkr) as size_chunkr_bounded. + pose proof (max_size_chunk chunkw) as size_chunkw_bounded. + unfold largest_size_chunk in size_chunkr_bounded, size_chunkw_bounded. + try change (Ptrofs.modulus - largest_size_chunk) with 4294967288 in *. + try change (Ptrofs.modulus - largest_size_chunk) with 18446744073709551608 in *. unfold eval_addressing, eval_addressing32, eval_addressing64 in *. destruct Archi.ptr64 eqn:PTR64. -- cgit From 091e00ed16d4189c27a05ad7056eab47bd29f5b7 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 3 Mar 2020 14:39:31 +0100 Subject: CSE2 for ARM --- arm/CSE2depsproof.v | 132 ++++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 132 insertions(+) create mode 100644 arm/CSE2depsproof.v diff --git a/arm/CSE2depsproof.v b/arm/CSE2depsproof.v new file mode 100644 index 00000000..2112a230 --- /dev/null +++ b/arm/CSE2depsproof.v @@ -0,0 +1,132 @@ +Require Import Coqlib Maps Errors Integers Floats Lattice Kildall. +Require Import AST Linking. +Require Import Memory Registers Op RTL Maps. + +Require Import Globalenvs Values. +Require Import Linking Values Memory Globalenvs Events Smallstep. +Require Import Registers Op RTL. +Require Import CSE2 CSE2deps. +Require Import Lia. + +Lemma ptrofs_size : + Ptrofs.wordsize = 32%nat. +Proof. + unfold Ptrofs.wordsize. + unfold Wordsize_Ptrofs.wordsize. + trivial. +Qed. + +Lemma ptrofs_modulus : + Ptrofs.modulus = 4294967296. +Proof. + unfold Ptrofs.modulus. + rewrite ptrofs_size. + destruct Archi.ptr64; reflexivity. +Qed. + +Section SOUNDNESS. + Variable F V : Type. + Variable genv: Genv.t F V. + Variable sp : val. + +Section MEMORY_WRITE. + Variable m m2 : mem. + Variable chunkw chunkr : memory_chunk. + Variable base : val. + + Variable addrw addrr valw valr : val. + Hypothesis STORE : Mem.storev chunkw m addrw valw = Some m2. + Hypothesis READ : Mem.loadv chunkr m addrr = Some valr. + + Section INDEXED_AWAY. + Variable ofsw ofsr : int. + Hypothesis ADDRW : eval_addressing genv sp + (Aindexed ofsw) (base :: nil) = Some addrw. + Hypothesis ADDRR : eval_addressing genv sp + (Aindexed ofsr) (base :: nil) = Some addrr. + + Lemma load_store_away1 : + forall RANGEW : 0 <= Int.unsigned ofsw <= Ptrofs.modulus - largest_size_chunk, + forall RANGER : 0 <= Int.unsigned ofsr <= Ptrofs.modulus - largest_size_chunk, + forall SWAPPABLE : Int.unsigned ofsw + size_chunk chunkw <= Int.unsigned ofsr + \/ Int.unsigned ofsr + size_chunk chunkr <= Int.unsigned ofsw, + Mem.loadv chunkr m2 addrr = Some valr. + Proof. + intros. + + pose proof (max_size_chunk chunkr) as size_chunkr_bounded. + pose proof (max_size_chunk chunkw) as size_chunkw_bounded. + unfold largest_size_chunk in *. + + rewrite ptrofs_modulus in *. + simpl in *. + inv ADDRR. + inv ADDRW. + rewrite <- READ. + destruct base; try discriminate. + eapply Mem.load_store_other with (chunk := chunkw) (v := valw) (b := b). + exact STORE. + right. + + all: try (destruct (Ptrofs.unsigned_add_either i (Ptrofs.of_int ofsr)) as [OFSR | OFSR]; + rewrite OFSR). + all: try (destruct (Ptrofs.unsigned_add_either i (Ptrofs.of_int ofsw)) as [OFSW | OFSW]; + rewrite OFSW). + + all: try rewrite ptrofs_modulus in *. + + all: unfold Ptrofs.of_int. + + all: repeat rewrite Ptrofs.unsigned_repr by (change Ptrofs.max_unsigned with 4294967295; lia). + all: intuition lia. + Qed. + + Theorem load_store_away : + can_swap_accesses_ofs (Int.unsigned ofsr) chunkr (Int.unsigned ofsw) chunkw = true -> + Mem.loadv chunkr m2 addrr = Some valr. + Proof. + intro SWAP. + unfold can_swap_accesses_ofs in SWAP. + repeat rewrite andb_true_iff in SWAP. + repeat rewrite orb_true_iff in SWAP. + repeat rewrite Z.leb_le in SWAP. + apply load_store_away1. + all: tauto. + Qed. + End INDEXED_AWAY. +End MEMORY_WRITE. +End SOUNDNESS. + + +Section SOUNDNESS. + Variable F V : Type. + Variable genv: Genv.t F V. + Variable sp : val. + +Lemma may_overlap_sound: + forall m m' : mem, + forall chunk addr args chunk' addr' args' v a a' vl rs, + (eval_addressing genv sp addr (rs ## args)) = Some a -> + (eval_addressing genv sp addr' (rs ## args')) = Some a' -> + (may_overlap chunk addr args chunk' addr' args') = false -> + (Mem.storev chunk m a v) = Some m' -> + (Mem.loadv chunk' m a') = Some vl -> + (Mem.loadv chunk' m' a') = Some vl. +Proof. + intros until rs. + intros ADDR ADDR' OVERLAP STORE LOAD. + destruct addr; destruct addr'; try discriminate. + { (* Aindexed / Aindexed *) + destruct args as [ | base [ | ]]. 1,3: discriminate. + destruct args' as [ | base' [ | ]]. 1,3: discriminate. + simpl in OVERLAP. + destruct (peq base base'). 2: discriminate. + subst base'. + destruct (can_swap_accesses_ofs (Int.unsigned i0) chunk' (Int.unsigned i) chunk) eqn:SWAP. + 2: discriminate. + simpl in *. + eapply load_store_away with (F:=F) (V:=V) (genv:=genv) (sp:=sp); eassumption. + } +Qed. + +End SOUNDNESS. -- cgit From a2b5e7c85dbbc6a27d941dcd931b36c4aa747fb5 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 3 Mar 2020 14:57:20 +0100 Subject: aarch64 --- aarch64/CSE2deps.v | 20 ++++++++ aarch64/CSE2depsproof.v | 130 ++++++++++++++++++++++++++++++++++++++++++++++++ 2 files changed, 150 insertions(+) create mode 100644 aarch64/CSE2deps.v create mode 100644 aarch64/CSE2depsproof.v diff --git a/aarch64/CSE2deps.v b/aarch64/CSE2deps.v new file mode 100644 index 00000000..90b514a2 --- /dev/null +++ b/aarch64/CSE2deps.v @@ -0,0 +1,20 @@ +Require Import BoolEqual Coqlib. +Require Import AST Integers Floats. +Require Import Values Memory Globalenvs Events. +Require Import Op. + + +Definition can_swap_accesses_ofs ofsr chunkr ofsw chunkw := + (0 <=? ofsw) && (ofsw <=? (Ptrofs.modulus - largest_size_chunk)) + && (0 <=? ofsr) && (ofsr <=? (Ptrofs.modulus - largest_size_chunk)) + && ((ofsw + size_chunk chunkw <=? ofsr) || + (ofsr + size_chunk chunkr <=? ofsw)). + +Definition may_overlap chunk addr args chunk' addr' args' := + match addr, addr', args, args' with + | (Aindexed ofs), (Aindexed ofs'), + (base :: nil), (base' :: nil) => + if peq base base' + then negb (can_swap_accesses_ofs (Int64.unsigned ofs') chunk' (Int64.unsigned ofs) chunk) + else true | _, _, _, _ => true + end. diff --git a/aarch64/CSE2depsproof.v b/aarch64/CSE2depsproof.v new file mode 100644 index 00000000..e20824e3 --- /dev/null +++ b/aarch64/CSE2depsproof.v @@ -0,0 +1,130 @@ +Require Import Coqlib Maps Errors Integers Floats Lattice Kildall. +Require Import AST Linking. +Require Import Memory Registers Op RTL Maps. + +Require Import Globalenvs Values. +Require Import Linking Values Memory Globalenvs Events Smallstep. +Require Import Registers Op RTL. +Require Import CSE2 CSE2deps. +Require Import Lia. + +Lemma ptrofs_size : + Ptrofs.wordsize = 64%nat. +Proof. + unfold Ptrofs.wordsize. + unfold Wordsize_Ptrofs.wordsize. + trivial. +Qed. + +Lemma ptrofs_modulus : + Ptrofs.modulus = 18446744073709551616. +Proof. + reflexivity. +Qed. + +Section SOUNDNESS. + Variable F V : Type. + Variable genv: Genv.t F V. + Variable sp : val. + +Section MEMORY_WRITE. + Variable m m2 : mem. + Variable chunkw chunkr : memory_chunk. + Variable base : val. + + Variable addrw addrr valw valr : val. + Hypothesis STORE : Mem.storev chunkw m addrw valw = Some m2. + Hypothesis READ : Mem.loadv chunkr m addrr = Some valr. + + Section INDEXED_AWAY. + Variable ofsw ofsr : int64. + Hypothesis ADDRW : eval_addressing genv sp + (Aindexed ofsw) (base :: nil) = Some addrw. + Hypothesis ADDRR : eval_addressing genv sp + (Aindexed ofsr) (base :: nil) = Some addrr. + + Lemma load_store_away1 : + forall RANGEW : 0 <= Int64.unsigned ofsw <= Ptrofs.modulus - largest_size_chunk, + forall RANGER : 0 <= Int64.unsigned ofsr <= Ptrofs.modulus - largest_size_chunk, + forall SWAPPABLE : Int64.unsigned ofsw + size_chunk chunkw <= Int64.unsigned ofsr + \/ Int64.unsigned ofsr + size_chunk chunkr <= Int64.unsigned ofsw, + Mem.loadv chunkr m2 addrr = Some valr. + Proof. + intros. + + pose proof (max_size_chunk chunkr) as size_chunkr_bounded. + pose proof (max_size_chunk chunkw) as size_chunkw_bounded. + unfold largest_size_chunk in *. + + rewrite ptrofs_modulus in *. + simpl in *. + inv ADDRR. + inv ADDRW. + rewrite <- READ. + destruct base; try discriminate. + eapply Mem.load_store_other with (chunk := chunkw) (v := valw) (b := b). + exact STORE. + right. + + all: unfold Ptrofs.of_int64. + + all: try (destruct (Ptrofs.unsigned_add_either i (Ptrofs.repr (Int64.unsigned ofsr))) as [OFSR | OFSR]; + rewrite OFSR). + all: try (destruct (Ptrofs.unsigned_add_either i (Ptrofs.repr (Int64.unsigned ofsw))) as [OFSW | OFSW]; + rewrite OFSW). + all: repeat rewrite Ptrofs.unsigned_repr by (change Ptrofs.max_unsigned with 18446744073709551615; lia). + + all: try rewrite ptrofs_modulus in *. + + all: intuition lia. + Qed. + + Theorem load_store_away : + can_swap_accesses_ofs (Ptrofs.unsigned ofsr) chunkr (Ptrofs.unsigned ofsw) chunkw = true -> + Mem.loadv chunkr m2 addrr = Some valr. + Proof. + intro SWAP. + unfold can_swap_accesses_ofs in SWAP. + repeat rewrite andb_true_iff in SWAP. + repeat rewrite orb_true_iff in SWAP. + repeat rewrite Z.leb_le in SWAP. + apply load_store_away1. + all: tauto. + Qed. + End INDEXED_AWAY. +End MEMORY_WRITE. +End SOUNDNESS. + + +Section SOUNDNESS. + Variable F V : Type. + Variable genv: Genv.t F V. + Variable sp : val. + +Lemma may_overlap_sound: + forall m m' : mem, + forall chunk addr args chunk' addr' args' v a a' vl rs, + (eval_addressing genv sp addr (rs ## args)) = Some a -> + (eval_addressing genv sp addr' (rs ## args')) = Some a' -> + (may_overlap chunk addr args chunk' addr' args') = false -> + (Mem.storev chunk m a v) = Some m' -> + (Mem.loadv chunk' m a') = Some vl -> + (Mem.loadv chunk' m' a') = Some vl. +Proof. + intros until rs. + intros ADDR ADDR' OVERLAP STORE LOAD. + destruct addr; destruct addr'; try discriminate. + { (* Aindexed / Aindexed *) + destruct args as [ | base [ | ]]. 1,3: discriminate. + destruct args' as [ | base' [ | ]]. 1,3: discriminate. + simpl in OVERLAP. + destruct (peq base base'). 2: discriminate. + subst base'. + destruct (can_swap_accesses_ofs (Ptrofs.unsigned i0) chunk' (Ptrofs.unsigned i) chunk) eqn:SWAP. + 2: discriminate. + simpl in *. + eapply load_store_away with (F:=F) (V:=V) (genv:=genv) (sp:=sp); eassumption. + } +Qed. + +End SOUNDNESS. -- cgit From 63c878610c5ef531731f5d9f83570f19c8c1acbc Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 3 Mar 2020 15:08:43 +0100 Subject: CSE2 for powerpc --- powerpc/CSE2deps.v | 20 ++++++++ powerpc/CSE2depsproof.v | 132 ++++++++++++++++++++++++++++++++++++++++++++++++ 2 files changed, 152 insertions(+) create mode 100644 powerpc/CSE2deps.v create mode 100644 powerpc/CSE2depsproof.v diff --git a/powerpc/CSE2deps.v b/powerpc/CSE2deps.v new file mode 100644 index 00000000..9db51bbb --- /dev/null +++ b/powerpc/CSE2deps.v @@ -0,0 +1,20 @@ +Require Import BoolEqual Coqlib. +Require Import AST Integers Floats. +Require Import Values Memory Globalenvs Events. +Require Import Op. + + +Definition can_swap_accesses_ofs ofsr chunkr ofsw chunkw := + (0 <=? ofsw) && (ofsw <=? (Ptrofs.modulus - largest_size_chunk)) + && (0 <=? ofsr) && (ofsr <=? (Ptrofs.modulus - largest_size_chunk)) + && ((ofsw + size_chunk chunkw <=? ofsr) || + (ofsr + size_chunk chunkr <=? ofsw)). + +Definition may_overlap chunk addr args chunk' addr' args' := + match addr, addr', args, args' with + | (Aindexed ofs), (Aindexed ofs'), + (base :: nil), (base' :: nil) => + if peq base base' + then negb (can_swap_accesses_ofs (Int.unsigned ofs') chunk' (Int.unsigned ofs) chunk) + else true | _, _, _, _ => true + end. diff --git a/powerpc/CSE2depsproof.v b/powerpc/CSE2depsproof.v new file mode 100644 index 00000000..2112a230 --- /dev/null +++ b/powerpc/CSE2depsproof.v @@ -0,0 +1,132 @@ +Require Import Coqlib Maps Errors Integers Floats Lattice Kildall. +Require Import AST Linking. +Require Import Memory Registers Op RTL Maps. + +Require Import Globalenvs Values. +Require Import Linking Values Memory Globalenvs Events Smallstep. +Require Import Registers Op RTL. +Require Import CSE2 CSE2deps. +Require Import Lia. + +Lemma ptrofs_size : + Ptrofs.wordsize = 32%nat. +Proof. + unfold Ptrofs.wordsize. + unfold Wordsize_Ptrofs.wordsize. + trivial. +Qed. + +Lemma ptrofs_modulus : + Ptrofs.modulus = 4294967296. +Proof. + unfold Ptrofs.modulus. + rewrite ptrofs_size. + destruct Archi.ptr64; reflexivity. +Qed. + +Section SOUNDNESS. + Variable F V : Type. + Variable genv: Genv.t F V. + Variable sp : val. + +Section MEMORY_WRITE. + Variable m m2 : mem. + Variable chunkw chunkr : memory_chunk. + Variable base : val. + + Variable addrw addrr valw valr : val. + Hypothesis STORE : Mem.storev chunkw m addrw valw = Some m2. + Hypothesis READ : Mem.loadv chunkr m addrr = Some valr. + + Section INDEXED_AWAY. + Variable ofsw ofsr : int. + Hypothesis ADDRW : eval_addressing genv sp + (Aindexed ofsw) (base :: nil) = Some addrw. + Hypothesis ADDRR : eval_addressing genv sp + (Aindexed ofsr) (base :: nil) = Some addrr. + + Lemma load_store_away1 : + forall RANGEW : 0 <= Int.unsigned ofsw <= Ptrofs.modulus - largest_size_chunk, + forall RANGER : 0 <= Int.unsigned ofsr <= Ptrofs.modulus - largest_size_chunk, + forall SWAPPABLE : Int.unsigned ofsw + size_chunk chunkw <= Int.unsigned ofsr + \/ Int.unsigned ofsr + size_chunk chunkr <= Int.unsigned ofsw, + Mem.loadv chunkr m2 addrr = Some valr. + Proof. + intros. + + pose proof (max_size_chunk chunkr) as size_chunkr_bounded. + pose proof (max_size_chunk chunkw) as size_chunkw_bounded. + unfold largest_size_chunk in *. + + rewrite ptrofs_modulus in *. + simpl in *. + inv ADDRR. + inv ADDRW. + rewrite <- READ. + destruct base; try discriminate. + eapply Mem.load_store_other with (chunk := chunkw) (v := valw) (b := b). + exact STORE. + right. + + all: try (destruct (Ptrofs.unsigned_add_either i (Ptrofs.of_int ofsr)) as [OFSR | OFSR]; + rewrite OFSR). + all: try (destruct (Ptrofs.unsigned_add_either i (Ptrofs.of_int ofsw)) as [OFSW | OFSW]; + rewrite OFSW). + + all: try rewrite ptrofs_modulus in *. + + all: unfold Ptrofs.of_int. + + all: repeat rewrite Ptrofs.unsigned_repr by (change Ptrofs.max_unsigned with 4294967295; lia). + all: intuition lia. + Qed. + + Theorem load_store_away : + can_swap_accesses_ofs (Int.unsigned ofsr) chunkr (Int.unsigned ofsw) chunkw = true -> + Mem.loadv chunkr m2 addrr = Some valr. + Proof. + intro SWAP. + unfold can_swap_accesses_ofs in SWAP. + repeat rewrite andb_true_iff in SWAP. + repeat rewrite orb_true_iff in SWAP. + repeat rewrite Z.leb_le in SWAP. + apply load_store_away1. + all: tauto. + Qed. + End INDEXED_AWAY. +End MEMORY_WRITE. +End SOUNDNESS. + + +Section SOUNDNESS. + Variable F V : Type. + Variable genv: Genv.t F V. + Variable sp : val. + +Lemma may_overlap_sound: + forall m m' : mem, + forall chunk addr args chunk' addr' args' v a a' vl rs, + (eval_addressing genv sp addr (rs ## args)) = Some a -> + (eval_addressing genv sp addr' (rs ## args')) = Some a' -> + (may_overlap chunk addr args chunk' addr' args') = false -> + (Mem.storev chunk m a v) = Some m' -> + (Mem.loadv chunk' m a') = Some vl -> + (Mem.loadv chunk' m' a') = Some vl. +Proof. + intros until rs. + intros ADDR ADDR' OVERLAP STORE LOAD. + destruct addr; destruct addr'; try discriminate. + { (* Aindexed / Aindexed *) + destruct args as [ | base [ | ]]. 1,3: discriminate. + destruct args' as [ | base' [ | ]]. 1,3: discriminate. + simpl in OVERLAP. + destruct (peq base base'). 2: discriminate. + subst base'. + destruct (can_swap_accesses_ofs (Int.unsigned i0) chunk' (Int.unsigned i) chunk) eqn:SWAP. + 2: discriminate. + simpl in *. + eapply load_store_away with (F:=F) (V:=V) (genv:=genv) (sp:=sp); eassumption. + } +Qed. + +End SOUNDNESS. -- cgit From 3b640f041be480b82f1b3a1f695ed8a57193bf28 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 3 Mar 2020 15:28:27 +0100 Subject: CSE2 with alias analysis --- arm/CSE2deps.v | 20 ++++++++++++++++++++ 1 file changed, 20 insertions(+) create mode 100644 arm/CSE2deps.v diff --git a/arm/CSE2deps.v b/arm/CSE2deps.v new file mode 100644 index 00000000..9db51bbb --- /dev/null +++ b/arm/CSE2deps.v @@ -0,0 +1,20 @@ +Require Import BoolEqual Coqlib. +Require Import AST Integers Floats. +Require Import Values Memory Globalenvs Events. +Require Import Op. + + +Definition can_swap_accesses_ofs ofsr chunkr ofsw chunkw := + (0 <=? ofsw) && (ofsw <=? (Ptrofs.modulus - largest_size_chunk)) + && (0 <=? ofsr) && (ofsr <=? (Ptrofs.modulus - largest_size_chunk)) + && ((ofsw + size_chunk chunkw <=? ofsr) || + (ofsr + size_chunk chunkr <=? ofsw)). + +Definition may_overlap chunk addr args chunk' addr' args' := + match addr, addr', args, args' with + | (Aindexed ofs), (Aindexed ofs'), + (base :: nil), (base' :: nil) => + if peq base base' + then negb (can_swap_accesses_ofs (Int.unsigned ofs') chunk' (Int.unsigned ofs) chunk) + else true | _, _, _, _ => true + end. -- cgit From d14c78013f654ca586681136ba291f1487f1b586 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 3 Mar 2020 17:23:01 +0100 Subject: adjust for x86 --- x86/CSE2depsproof.v | 79 ++++++++++++++++++++++++----------------------------- 1 file changed, 35 insertions(+), 44 deletions(-) diff --git a/x86/CSE2depsproof.v b/x86/CSE2depsproof.v index 37e16dfe..1e913254 100644 --- a/x86/CSE2depsproof.v +++ b/x86/CSE2depsproof.v @@ -18,9 +18,8 @@ Section MEMORY_WRITE. Variable chunkw chunkr : memory_chunk. Variable base : val. - Variable addrw addrr valw valr : val. + Variable addrw addrr valw : val. Hypothesis STORE : Mem.storev chunkw m addrw valw = Some m2. - Hypothesis READ : Mem.loadv chunkr m addrr = Some valr. Section INDEXED_AWAY. Variable ofsw ofsr : Z. @@ -34,7 +33,7 @@ Section MEMORY_WRITE. forall RANGER : 0 <= ofsr <= Ptrofs.modulus - largest_size_chunk, forall SWAPPABLE : ofsw + size_chunk chunkw <= ofsr \/ ofsr + size_chunk chunkr <= ofsw, - Mem.loadv chunkr m2 addrr = Some valr. + Mem.loadv chunkr m2 addrr = Mem.loadv chunkr m addrr. Proof. intros. @@ -42,14 +41,13 @@ Section MEMORY_WRITE. pose proof (max_size_chunk chunkw) as size_chunkw_bounded. try change (Ptrofs.modulus - largest_size_chunk) with 4294967288 in *. try change (Ptrofs.modulus - largest_size_chunk) with 18446744073709551608 in *. - destruct addrr ; simpl in * ; try discriminate. - unfold eval_addressing in *. + destruct addrr ; simpl in * ; trivial. + unfold eval_addressing, eval_addressing32, eval_addressing64 in *. destruct Archi.ptr64 eqn:PTR64; destruct base; simpl in *; try discriminate. rewrite PTR64 in *. inv ADDRR. inv ADDRW. - rewrite <- READ. eapply Mem.load_store_other with (chunk := chunkw) (v := valw) (b := b). exact STORE. right. @@ -84,7 +82,7 @@ Section MEMORY_WRITE. Theorem load_store_away : can_swap_accesses_ofs ofsr chunkr ofsw chunkw = true -> - Mem.loadv chunkr m2 addrr = Some valr. + Mem.loadv chunkr m2 addrr = Mem.loadv chunkr m addrr. Proof. intro SWAP. unfold can_swap_accesses_ofs in SWAP. @@ -113,9 +111,20 @@ Section MEMORY_WRITE. destruct b; reflexivity. Qed. + (* not needed + Lemma bool_cases_same: + forall {T : Type}, + forall b : bool, + forall x : T, + (if b then x else x) = x. + Proof. + destruct b; reflexivity. + Qed. + *) + Lemma load_store_diff_globals : symw <> symr -> - Mem.loadv chunkr m2 addrr = Some valr. + Mem.loadv chunkr m2 addrr = Mem.loadv chunkr m addrr. Proof. intros. unfold eval_addressing in *. @@ -127,7 +136,7 @@ Section MEMORY_WRITE. destruct ((Genv.genv_symb genv) ! symw) as [bw |] eqn:SYMW; inv ADDRW. 2: simpl in STORE; discriminate. destruct ((Genv.genv_symb genv) ! symr) as [br |] eqn:SYMR; inv ADDRR. - 2: simpl in READ; discriminate. + 2: reflexivity. assert (br <> bw). { intro EQ. @@ -138,7 +147,6 @@ Section MEMORY_WRITE. } congruence. } - rewrite <- READ. eapply Mem.load_store_other with (chunk := chunkw) (v := valw) (b := bw). - exact STORE. - left. assumption. @@ -158,7 +166,7 @@ Section MEMORY_WRITE. forall RANGER : 0 <= (Ptrofs.unsigned ofsr) <= Ptrofs.modulus - largest_size_chunk, forall SWAPPABLE : (Ptrofs.unsigned ofsw) + size_chunk chunkw <= (Ptrofs.unsigned ofsr) \/ (Ptrofs.unsigned ofsr) + size_chunk chunkr <= (Ptrofs.unsigned ofsw), - Mem.loadv chunkr m2 addrr = Some valr. + Mem.loadv chunkr m2 addrr = Mem.loadv chunkr m addrr. Proof. intros. @@ -168,40 +176,24 @@ Section MEMORY_WRITE. try change (Ptrofs.modulus - largest_size_chunk) with 4294967288 in *. try change (Ptrofs.modulus - largest_size_chunk) with 18446744073709551608 in *. unfold eval_addressing, eval_addressing32, eval_addressing64 in *. - destruct Archi.ptr64 eqn:PTR64. - - { - unfold Genv.symbol_address in *. - inv ADDRR. - inv ADDRW. - destruct (Genv.find_symbol genv sym). - 2: discriminate. - - rewrite <- READ. - eapply Mem.load_store_other with (chunk := chunkw) (v := valw) (b := b). - exact STORE. - right. - tauto. - } - { - unfold Genv.symbol_address in *. - inv ADDRR. - inv ADDRW. - destruct (Genv.find_symbol genv sym). - 2: discriminate. - - rewrite <- READ. - eapply Mem.load_store_other with (chunk := chunkw) (v := valw) (b := b). - exact STORE. - right. - tauto. - } + rewrite ptr64_cases in ADDRR. + rewrite ptr64_cases in ADDRW. + unfold Genv.symbol_address in *. + inv ADDRR. + inv ADDRW. + destruct (Genv.find_symbol genv sym). + 2: discriminate. + + eapply Mem.load_store_other with (chunk := chunkw) (v := valw) (b := b). + exact STORE. + right. + tauto. Qed. Lemma load_store_glob_away : (can_swap_accesses_ofs (Ptrofs.unsigned ofsr) chunkr (Ptrofs.unsigned ofsw) chunkw) = true -> - Mem.loadv chunkr m2 addrr = Some valr. + Mem.loadv chunkr m2 addrr = Mem.loadv chunkr m addrr. Proof. intro SWAP. unfold can_swap_accesses_ofs in SWAP. @@ -223,16 +215,15 @@ Section SOUNDNESS. Lemma may_overlap_sound: forall m m' : mem, - forall chunk addr args chunk' addr' args' v a a' vl rs, + forall chunk addr args chunk' addr' args' v a a' rs, (eval_addressing genv sp addr (rs ## args)) = Some a -> (eval_addressing genv sp addr' (rs ## args')) = Some a' -> (may_overlap chunk addr args chunk' addr' args') = false -> (Mem.storev chunk m a v) = Some m' -> - (Mem.loadv chunk' m a') = Some vl -> - (Mem.loadv chunk' m' a') = Some vl. + (Mem.loadv chunk' m' a') = (Mem.loadv chunk' m a'). Proof. intros until rs. - intros ADDR ADDR' OVERLAP STORE LOAD. + intros ADDR ADDR' OVERLAP STORE. destruct addr; destruct addr'; try discriminate. { (* Aindexed / Aindexed *) destruct args as [ | base [ | ]]. 1,3: discriminate. -- cgit From ab15e9d17f999637ae16b2913b3c6f4f71f79e37 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 3 Mar 2020 17:36:30 +0100 Subject: fix for risc-V --- riscV/CSE2depsproof.v | 16 +++++++--------- 1 file changed, 7 insertions(+), 9 deletions(-) diff --git a/riscV/CSE2depsproof.v b/riscV/CSE2depsproof.v index ee500965..2ed12658 100644 --- a/riscV/CSE2depsproof.v +++ b/riscV/CSE2depsproof.v @@ -36,7 +36,6 @@ Section MEMORY_WRITE. Variable addrw addrr valw valr : val. Hypothesis STORE : Mem.storev chunkw m addrw valw = Some m2. - Hypothesis READ : Mem.loadv chunkr m addrr = Some valr. Section INDEXED_AWAY. Variable ofsw ofsr : ptrofs. @@ -49,8 +48,9 @@ Section MEMORY_WRITE. forall RANGEW : 0 <= Ptrofs.unsigned ofsw <= Ptrofs.modulus - largest_size_chunk, forall RANGER : 0 <= Ptrofs.unsigned ofsr <= Ptrofs.modulus - largest_size_chunk, forall SWAPPABLE : Ptrofs.unsigned ofsw + size_chunk chunkw <= Ptrofs.unsigned ofsr - \/ Ptrofs.unsigned ofsr + size_chunk chunkr <= Ptrofs.unsigned ofsw, - Mem.loadv chunkr m2 addrr = Some valr. + \/ Ptrofs.unsigned ofsr + size_chunk chunkr <= Ptrofs.unsigned ofsw, + Mem.loadv chunkr m2 addrr = Mem.loadv chunkr m addrr. + Proof. intros. @@ -62,7 +62,6 @@ Section MEMORY_WRITE. simpl in *. inv ADDRR. inv ADDRW. - rewrite <- READ. destruct base; try discriminate. eapply Mem.load_store_other with (chunk := chunkw) (v := valw) (b := b). exact STORE. @@ -80,7 +79,7 @@ Section MEMORY_WRITE. Theorem load_store_away : can_swap_accesses_ofs (Ptrofs.unsigned ofsr) chunkr (Ptrofs.unsigned ofsw) chunkw = true -> - Mem.loadv chunkr m2 addrr = Some valr. + Mem.loadv chunkr m2 addrr = Mem.loadv chunkr m addrr. Proof. intro SWAP. unfold can_swap_accesses_ofs in SWAP. @@ -102,16 +101,15 @@ Section SOUNDNESS. Lemma may_overlap_sound: forall m m' : mem, - forall chunk addr args chunk' addr' args' v a a' vl rs, + forall chunk addr args chunk' addr' args' v a a' rs, (eval_addressing genv sp addr (rs ## args)) = Some a -> (eval_addressing genv sp addr' (rs ## args')) = Some a' -> (may_overlap chunk addr args chunk' addr' args') = false -> (Mem.storev chunk m a v) = Some m' -> - (Mem.loadv chunk' m a') = Some vl -> - (Mem.loadv chunk' m' a') = Some vl. + (Mem.loadv chunk' m' a') = (Mem.loadv chunk' m a'). Proof. intros until rs. - intros ADDR ADDR' OVERLAP STORE LOAD. + intros ADDR ADDR' OVERLAP STORE. destruct addr; destruct addr'; try discriminate. { (* Aindexed / Aindexed *) destruct args as [ | base [ | ]]. 1,3: discriminate. -- cgit From 10cb118e1201268b993973852499d38ce6b8d890 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 3 Mar 2020 17:44:04 +0100 Subject: ported for ppc --- powerpc/CSE2depsproof.v | 39 +++++++++++++++++---------------------- 1 file changed, 17 insertions(+), 22 deletions(-) diff --git a/powerpc/CSE2depsproof.v b/powerpc/CSE2depsproof.v index 2112a230..a047b02a 100644 --- a/powerpc/CSE2depsproof.v +++ b/powerpc/CSE2depsproof.v @@ -9,7 +9,7 @@ Require Import CSE2 CSE2deps. Require Import Lia. Lemma ptrofs_size : - Ptrofs.wordsize = 32%nat. + Ptrofs.wordsize = if Archi.ptr64 then 64%nat else 32%nat. Proof. unfold Ptrofs.wordsize. unfold Wordsize_Ptrofs.wordsize. @@ -17,7 +17,7 @@ Proof. Qed. Lemma ptrofs_modulus : - Ptrofs.modulus = 4294967296. + Ptrofs.modulus = if Archi.ptr64 then 18446744073709551616 else 4294967296. Proof. unfold Ptrofs.modulus. rewrite ptrofs_size. @@ -34,23 +34,22 @@ Section MEMORY_WRITE. Variable chunkw chunkr : memory_chunk. Variable base : val. - Variable addrw addrr valw valr : val. + Variable addrw addrr valw : val. Hypothesis STORE : Mem.storev chunkw m addrw valw = Some m2. - Hypothesis READ : Mem.loadv chunkr m addrr = Some valr. Section INDEXED_AWAY. - Variable ofsw ofsr : int. + Variable ofsw ofsr : ptrofs. Hypothesis ADDRW : eval_addressing genv sp (Aindexed ofsw) (base :: nil) = Some addrw. Hypothesis ADDRR : eval_addressing genv sp (Aindexed ofsr) (base :: nil) = Some addrr. Lemma load_store_away1 : - forall RANGEW : 0 <= Int.unsigned ofsw <= Ptrofs.modulus - largest_size_chunk, - forall RANGER : 0 <= Int.unsigned ofsr <= Ptrofs.modulus - largest_size_chunk, - forall SWAPPABLE : Int.unsigned ofsw + size_chunk chunkw <= Int.unsigned ofsr - \/ Int.unsigned ofsr + size_chunk chunkr <= Int.unsigned ofsw, - Mem.loadv chunkr m2 addrr = Some valr. + forall RANGEW : 0 <= Ptrofs.unsigned ofsw <= Ptrofs.modulus - largest_size_chunk, + forall RANGER : 0 <= Ptrofs.unsigned ofsr <= Ptrofs.modulus - largest_size_chunk, + forall SWAPPABLE : Ptrofs.unsigned ofsw + size_chunk chunkw <= Ptrofs.unsigned ofsr + \/ Ptrofs.unsigned ofsr + size_chunk chunkr <= Ptrofs.unsigned ofsw, + Mem.loadv chunkr m2 addrr = Mem.loadv chunkr m addrr. Proof. intros. @@ -62,28 +61,25 @@ Section MEMORY_WRITE. simpl in *. inv ADDRR. inv ADDRW. - rewrite <- READ. destruct base; try discriminate. eapply Mem.load_store_other with (chunk := chunkw) (v := valw) (b := b). exact STORE. right. - all: try (destruct (Ptrofs.unsigned_add_either i (Ptrofs.of_int ofsr)) as [OFSR | OFSR]; + all: try (destruct (Ptrofs.unsigned_add_either i ofsr) as [OFSR | OFSR]; rewrite OFSR). - all: try (destruct (Ptrofs.unsigned_add_either i (Ptrofs.of_int ofsw)) as [OFSW | OFSW]; + all: try (destruct (Ptrofs.unsigned_add_either i ofsw) as [OFSW | OFSW]; rewrite OFSW). all: try rewrite ptrofs_modulus in *. - - all: unfold Ptrofs.of_int. all: repeat rewrite Ptrofs.unsigned_repr by (change Ptrofs.max_unsigned with 4294967295; lia). all: intuition lia. Qed. Theorem load_store_away : - can_swap_accesses_ofs (Int.unsigned ofsr) chunkr (Int.unsigned ofsw) chunkw = true -> - Mem.loadv chunkr m2 addrr = Some valr. + can_swap_accesses_ofs (Ptrofs.unsigned ofsr) chunkr (Ptrofs.unsigned ofsw) chunkw = true -> + Mem.loadv chunkr m2 addrr = Mem.loadv chunkr m addrr. Proof. intro SWAP. unfold can_swap_accesses_ofs in SWAP. @@ -105,16 +101,15 @@ Section SOUNDNESS. Lemma may_overlap_sound: forall m m' : mem, - forall chunk addr args chunk' addr' args' v a a' vl rs, + forall chunk addr args chunk' addr' args' v a a' rs, (eval_addressing genv sp addr (rs ## args)) = Some a -> (eval_addressing genv sp addr' (rs ## args')) = Some a' -> (may_overlap chunk addr args chunk' addr' args') = false -> (Mem.storev chunk m a v) = Some m' -> - (Mem.loadv chunk' m a') = Some vl -> - (Mem.loadv chunk' m' a') = Some vl. + (Mem.loadv chunk' m' a') = (Mem.loadv chunk' m a'). Proof. intros until rs. - intros ADDR ADDR' OVERLAP STORE LOAD. + intros ADDR ADDR' OVERLAP STORE. destruct addr; destruct addr'; try discriminate. { (* Aindexed / Aindexed *) destruct args as [ | base [ | ]]. 1,3: discriminate. @@ -122,7 +117,7 @@ Proof. simpl in OVERLAP. destruct (peq base base'). 2: discriminate. subst base'. - destruct (can_swap_accesses_ofs (Int.unsigned i0) chunk' (Int.unsigned i) chunk) eqn:SWAP. + destruct (can_swap_accesses_ofs (Ptrofs.unsigned i0) chunk' (Ptrofs.unsigned i) chunk) eqn:SWAP. 2: discriminate. simpl in *. eapply load_store_away with (F:=F) (V:=V) (genv:=genv) (sp:=sp); eassumption. -- cgit From f503e4287bc76150fd3ec5be8c076bf734361493 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 3 Mar 2020 17:56:53 +0100 Subject: ported to arm --- arm/CSE2depsproof.v | 15 ++++++--------- 1 file changed, 6 insertions(+), 9 deletions(-) diff --git a/arm/CSE2depsproof.v b/arm/CSE2depsproof.v index 2112a230..61fe5980 100644 --- a/arm/CSE2depsproof.v +++ b/arm/CSE2depsproof.v @@ -34,9 +34,8 @@ Section MEMORY_WRITE. Variable chunkw chunkr : memory_chunk. Variable base : val. - Variable addrw addrr valw valr : val. + Variable addrw addrr valw : val. Hypothesis STORE : Mem.storev chunkw m addrw valw = Some m2. - Hypothesis READ : Mem.loadv chunkr m addrr = Some valr. Section INDEXED_AWAY. Variable ofsw ofsr : int. @@ -50,7 +49,7 @@ Section MEMORY_WRITE. forall RANGER : 0 <= Int.unsigned ofsr <= Ptrofs.modulus - largest_size_chunk, forall SWAPPABLE : Int.unsigned ofsw + size_chunk chunkw <= Int.unsigned ofsr \/ Int.unsigned ofsr + size_chunk chunkr <= Int.unsigned ofsw, - Mem.loadv chunkr m2 addrr = Some valr. + Mem.loadv chunkr m2 addrr = Mem.loadv chunkr m addrr. Proof. intros. @@ -62,7 +61,6 @@ Section MEMORY_WRITE. simpl in *. inv ADDRR. inv ADDRW. - rewrite <- READ. destruct base; try discriminate. eapply Mem.load_store_other with (chunk := chunkw) (v := valw) (b := b). exact STORE. @@ -83,7 +81,7 @@ Section MEMORY_WRITE. Theorem load_store_away : can_swap_accesses_ofs (Int.unsigned ofsr) chunkr (Int.unsigned ofsw) chunkw = true -> - Mem.loadv chunkr m2 addrr = Some valr. + Mem.loadv chunkr m2 addrr = Mem.loadv chunkr m addrr. Proof. intro SWAP. unfold can_swap_accesses_ofs in SWAP. @@ -105,16 +103,15 @@ Section SOUNDNESS. Lemma may_overlap_sound: forall m m' : mem, - forall chunk addr args chunk' addr' args' v a a' vl rs, + forall chunk addr args chunk' addr' args' v a a' rs, (eval_addressing genv sp addr (rs ## args)) = Some a -> (eval_addressing genv sp addr' (rs ## args')) = Some a' -> (may_overlap chunk addr args chunk' addr' args') = false -> (Mem.storev chunk m a v) = Some m' -> - (Mem.loadv chunk' m a') = Some vl -> - (Mem.loadv chunk' m' a') = Some vl. + (Mem.loadv chunk' m' a') = (Mem.loadv chunk' m a'). Proof. intros until rs. - intros ADDR ADDR' OVERLAP STORE LOAD. + intros ADDR ADDR' OVERLAP STORE. destruct addr; destruct addr'; try discriminate. { (* Aindexed / Aindexed *) destruct args as [ | base [ | ]]. 1,3: discriminate. -- cgit From 50fbe4a02ab8deab82c4f137dc9575bac6b9b573 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 3 Mar 2020 18:08:59 +0100 Subject: fix for aarch64 --- aarch64/CSE2depsproof.v | 20 +++++++++----------- 1 file changed, 9 insertions(+), 11 deletions(-) diff --git a/aarch64/CSE2depsproof.v b/aarch64/CSE2depsproof.v index e20824e3..4aac23af 100644 --- a/aarch64/CSE2depsproof.v +++ b/aarch64/CSE2depsproof.v @@ -32,9 +32,8 @@ Section MEMORY_WRITE. Variable chunkw chunkr : memory_chunk. Variable base : val. - Variable addrw addrr valw valr : val. + Variable addrw addrr valw : val. Hypothesis STORE : Mem.storev chunkw m addrw valw = Some m2. - Hypothesis READ : Mem.loadv chunkr m addrr = Some valr. Section INDEXED_AWAY. Variable ofsw ofsr : int64. @@ -48,7 +47,7 @@ Section MEMORY_WRITE. forall RANGER : 0 <= Int64.unsigned ofsr <= Ptrofs.modulus - largest_size_chunk, forall SWAPPABLE : Int64.unsigned ofsw + size_chunk chunkw <= Int64.unsigned ofsr \/ Int64.unsigned ofsr + size_chunk chunkr <= Int64.unsigned ofsw, - Mem.loadv chunkr m2 addrr = Some valr. + Mem.loadv chunkr m2 addrr = Mem.loadv chunkr m addrr. Proof. intros. @@ -60,7 +59,7 @@ Section MEMORY_WRITE. simpl in *. inv ADDRR. inv ADDRW. - rewrite <- READ. + destruct base; try discriminate. eapply Mem.load_store_other with (chunk := chunkw) (v := valw) (b := b). exact STORE. @@ -80,8 +79,8 @@ Section MEMORY_WRITE. Qed. Theorem load_store_away : - can_swap_accesses_ofs (Ptrofs.unsigned ofsr) chunkr (Ptrofs.unsigned ofsw) chunkw = true -> - Mem.loadv chunkr m2 addrr = Some valr. + can_swap_accesses_ofs (Int64.unsigned ofsr) chunkr (Int64.unsigned ofsw) chunkw = true -> + Mem.loadv chunkr m2 addrr = Mem.loadv chunkr m addrr. Proof. intro SWAP. unfold can_swap_accesses_ofs in SWAP. @@ -103,16 +102,15 @@ Section SOUNDNESS. Lemma may_overlap_sound: forall m m' : mem, - forall chunk addr args chunk' addr' args' v a a' vl rs, + forall chunk addr args chunk' addr' args' v a a' rs, (eval_addressing genv sp addr (rs ## args)) = Some a -> (eval_addressing genv sp addr' (rs ## args')) = Some a' -> (may_overlap chunk addr args chunk' addr' args') = false -> (Mem.storev chunk m a v) = Some m' -> - (Mem.loadv chunk' m a') = Some vl -> - (Mem.loadv chunk' m' a') = Some vl. + (Mem.loadv chunk' m' a') = (Mem.loadv chunk' m a'). Proof. intros until rs. - intros ADDR ADDR' OVERLAP STORE LOAD. + intros ADDR ADDR' OVERLAP STORE. destruct addr; destruct addr'; try discriminate. { (* Aindexed / Aindexed *) destruct args as [ | base [ | ]]. 1,3: discriminate. @@ -120,7 +118,7 @@ Proof. simpl in OVERLAP. destruct (peq base base'). 2: discriminate. subst base'. - destruct (can_swap_accesses_ofs (Ptrofs.unsigned i0) chunk' (Ptrofs.unsigned i) chunk) eqn:SWAP. + destruct (can_swap_accesses_ofs (Int64.unsigned ofs0) chunk' (Int64.unsigned ofs) chunk) eqn:SWAP. 2: discriminate. simpl in *. eapply load_store_away with (F:=F) (V:=V) (genv:=genv) (sp:=sp); eassumption. -- cgit From 5996f8d84a61f76292f1a40c39faeb838011de6e Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 3 Mar 2020 18:10:00 +0100 Subject: fixes for risc-V --- riscV/CSE2depsproof.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/riscV/CSE2depsproof.v b/riscV/CSE2depsproof.v index 2ed12658..a3811e78 100644 --- a/riscV/CSE2depsproof.v +++ b/riscV/CSE2depsproof.v @@ -34,7 +34,7 @@ Section MEMORY_WRITE. Variable chunkw chunkr : memory_chunk. Variable base : val. - Variable addrw addrr valw valr : val. + Variable addrw addrr valw : val. Hypothesis STORE : Mem.storev chunkw m addrw valw = Some m2. Section INDEXED_AWAY. -- cgit From 4f659bb46bb3e2d2c1f297d65e71bb8e66782f79 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 3 Mar 2020 18:13:04 +0100 Subject: forgot k1C --- mppa_k1c/CSE2deps.v | 20 ++++++++ mppa_k1c/CSE2depsproof.v | 127 +++++++++++++++++++++++++++++++++++++++++++++++ 2 files changed, 147 insertions(+) create mode 100644 mppa_k1c/CSE2deps.v create mode 100644 mppa_k1c/CSE2depsproof.v diff --git a/mppa_k1c/CSE2deps.v b/mppa_k1c/CSE2deps.v new file mode 100644 index 00000000..8ab9242a --- /dev/null +++ b/mppa_k1c/CSE2deps.v @@ -0,0 +1,20 @@ +Require Import BoolEqual Coqlib. +Require Import AST Integers Floats. +Require Import Values Memory Globalenvs Events. +Require Import Op. + + +Definition can_swap_accesses_ofs ofsr chunkr ofsw chunkw := + (0 <=? ofsw) && (ofsw <=? (Ptrofs.modulus - largest_size_chunk)) + && (0 <=? ofsr) && (ofsr <=? (Ptrofs.modulus - largest_size_chunk)) + && ((ofsw + size_chunk chunkw <=? ofsr) || + (ofsr + size_chunk chunkr <=? ofsw)). + +Definition may_overlap chunk addr args chunk' addr' args' := + match addr, addr', args, args' with + | (Aindexed ofs), (Aindexed ofs'), + (base :: nil), (base' :: nil) => + if peq base base' + then negb (can_swap_accesses_ofs (Ptrofs.unsigned ofs') chunk' (Ptrofs.unsigned ofs) chunk) + else true | _, _, _, _ => true + end. diff --git a/mppa_k1c/CSE2depsproof.v b/mppa_k1c/CSE2depsproof.v new file mode 100644 index 00000000..a3811e78 --- /dev/null +++ b/mppa_k1c/CSE2depsproof.v @@ -0,0 +1,127 @@ +Require Import Coqlib Maps Errors Integers Floats Lattice Kildall. +Require Import AST Linking. +Require Import Memory Registers Op RTL Maps. + +Require Import Globalenvs Values. +Require Import Linking Values Memory Globalenvs Events Smallstep. +Require Import Registers Op RTL. +Require Import CSE2 CSE2deps. +Require Import Lia. + +Lemma ptrofs_size : + Ptrofs.wordsize = (if Archi.ptr64 then 64 else 32)%nat. +Proof. + unfold Ptrofs.wordsize. + unfold Wordsize_Ptrofs.wordsize. + trivial. +Qed. + +Lemma ptrofs_modulus : + Ptrofs.modulus = if Archi.ptr64 then 18446744073709551616 else 4294967296. +Proof. + unfold Ptrofs.modulus. + rewrite ptrofs_size. + destruct Archi.ptr64; reflexivity. +Qed. + +Section SOUNDNESS. + Variable F V : Type. + Variable genv: Genv.t F V. + Variable sp : val. + +Section MEMORY_WRITE. + Variable m m2 : mem. + Variable chunkw chunkr : memory_chunk. + Variable base : val. + + Variable addrw addrr valw : val. + Hypothesis STORE : Mem.storev chunkw m addrw valw = Some m2. + + Section INDEXED_AWAY. + Variable ofsw ofsr : ptrofs. + Hypothesis ADDRW : eval_addressing genv sp + (Aindexed ofsw) (base :: nil) = Some addrw. + Hypothesis ADDRR : eval_addressing genv sp + (Aindexed ofsr) (base :: nil) = Some addrr. + + Lemma load_store_away1 : + forall RANGEW : 0 <= Ptrofs.unsigned ofsw <= Ptrofs.modulus - largest_size_chunk, + forall RANGER : 0 <= Ptrofs.unsigned ofsr <= Ptrofs.modulus - largest_size_chunk, + forall SWAPPABLE : Ptrofs.unsigned ofsw + size_chunk chunkw <= Ptrofs.unsigned ofsr + \/ Ptrofs.unsigned ofsr + size_chunk chunkr <= Ptrofs.unsigned ofsw, + Mem.loadv chunkr m2 addrr = Mem.loadv chunkr m addrr. + + Proof. + intros. + + pose proof (max_size_chunk chunkr) as size_chunkr_bounded. + pose proof (max_size_chunk chunkw) as size_chunkw_bounded. + unfold largest_size_chunk in *. + + rewrite ptrofs_modulus in *. + simpl in *. + inv ADDRR. + inv ADDRW. + destruct base; try discriminate. + eapply Mem.load_store_other with (chunk := chunkw) (v := valw) (b := b). + exact STORE. + right. + + all: try (destruct (Ptrofs.unsigned_add_either i ofsr) as [OFSR | OFSR]; + rewrite OFSR). + all: try (destruct (Ptrofs.unsigned_add_either i ofsw) as [OFSW | OFSW]; + rewrite OFSW). + all: try rewrite ptrofs_modulus in *. + all: destruct Archi.ptr64. + + all: intuition lia. + Qed. + + Theorem load_store_away : + can_swap_accesses_ofs (Ptrofs.unsigned ofsr) chunkr (Ptrofs.unsigned ofsw) chunkw = true -> + Mem.loadv chunkr m2 addrr = Mem.loadv chunkr m addrr. + Proof. + intro SWAP. + unfold can_swap_accesses_ofs in SWAP. + repeat rewrite andb_true_iff in SWAP. + repeat rewrite orb_true_iff in SWAP. + repeat rewrite Z.leb_le in SWAP. + apply load_store_away1. + all: tauto. + Qed. + End INDEXED_AWAY. +End MEMORY_WRITE. +End SOUNDNESS. + + +Section SOUNDNESS. + Variable F V : Type. + Variable genv: Genv.t F V. + Variable sp : val. + +Lemma may_overlap_sound: + forall m m' : mem, + forall chunk addr args chunk' addr' args' v a a' rs, + (eval_addressing genv sp addr (rs ## args)) = Some a -> + (eval_addressing genv sp addr' (rs ## args')) = Some a' -> + (may_overlap chunk addr args chunk' addr' args') = false -> + (Mem.storev chunk m a v) = Some m' -> + (Mem.loadv chunk' m' a') = (Mem.loadv chunk' m a'). +Proof. + intros until rs. + intros ADDR ADDR' OVERLAP STORE. + destruct addr; destruct addr'; try discriminate. + { (* Aindexed / Aindexed *) + destruct args as [ | base [ | ]]. 1,3: discriminate. + destruct args' as [ | base' [ | ]]. 1,3: discriminate. + simpl in OVERLAP. + destruct (peq base base'). 2: discriminate. + subst base'. + destruct (can_swap_accesses_ofs (Ptrofs.unsigned i0) chunk' (Ptrofs.unsigned i) chunk) eqn:SWAP. + 2: discriminate. + simpl in *. + eapply load_store_away with (F:=F) (V:=V) (genv:=genv) (sp:=sp); eassumption. + } +Qed. + +End SOUNDNESS. -- cgit From 51db43fa5ea3f0bfcb42f68b59df1c39842c6486 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 3 Mar 2020 18:15:49 +0100 Subject: try to get it to compile --- Makefile | 1 + 1 file changed, 1 insertion(+) diff --git a/Makefile b/Makefile index 5566cf57..2cd40800 100644 --- a/Makefile +++ b/Makefile @@ -86,6 +86,7 @@ BACKEND=\ ValueDomain.v ValueAOp.v ValueAnalysis.v \ ConstpropOp.v Constprop.v ConstpropOpproof.v Constpropproof.v \ CSEdomain.v CombineOp.v CSE.v CombineOpproof.v CSEproof.v \ + CSE2deps.v CSE2depsproof.v \ CSE2.v CSE2proof.v \ NeedDomain.v NeedOp.v Deadcode.v Deadcodeproof.v \ Unusedglob.v Unusedglobproof.v \ -- cgit From 690fa3a3969f3e1294f8b381f6b8d9c051b264d3 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Tue, 3 Mar 2020 18:27:24 +0100 Subject: Linearize: Dependencies computing to decide which sequence to put first --- backend/Linearizeaux.ml | 163 +++++++++++++++++++++++++++++++++++++++--------- 1 file changed, 132 insertions(+), 31 deletions(-) diff --git a/backend/Linearizeaux.ml b/backend/Linearizeaux.ml index 3ef86344..58d7558b 100644 --- a/backend/Linearizeaux.ml +++ b/backend/Linearizeaux.ml @@ -140,33 +140,6 @@ let rec last_element = function | e :: [] -> e | e' :: e :: l -> last_element (e::l) -(** old version -let dfs code entrypoint = - let visited = ref (PTree.map (fun n i -> false) code) in - let rec dfs_list code = function - | [] -> [] - | node :: ln -> - let node_dfs = - if not (get_some @@ PTree.get node !visited) then begin - visited := PTree.set node true !visited; - match PTree.get node code with - | None -> failwith "No such node" - | Some bb -> [node] @ match (last_element bb) with - | Lop _ | Lload _ | Lgetstack _ | Lsetstack _ | Lstore _ | Lcall _ - | Lbuiltin _ -> assert false - | Ltailcall _ | Lreturn -> [] - | Lbranch n -> dfs_list code [n] - | Lcond (_, _, ifso, ifnot) -> dfs_list code [ifnot; ifso] - | Ljumptable(_, ln) -> dfs_list code ln - end - else [] - in node_dfs @ (dfs_list code ln) - in dfs_list code [entrypoint] - -let enumerate_aux_trace f reach = dfs f.fn_code f.fn_entrypoint -*) - - let forward_sequences code entry = let visited = ref (PTree.map (fun n i -> false) code) in (* returns the list of traversed nodes, and a list of nodes to start traversing next *) @@ -273,12 +246,140 @@ let try_merge code (fs: (BinNums.positive list) list) = done; !seqs -let order_sequences fs = fs +(** Code adapted from Duplicateaux.get_loop_headers + * + * Getting loop branches with a DFS visit : + * Each node is either Unvisited, Visited, or Processed + * pre-order: node becomes Processed + * post-order: node becomes Visited + * + * If we come accross an edge to a Processed node, it's a loop! + *) +type pos = BinNums.positive + +module PP = struct + type t = pos * pos + let compare a b = + let ax, ay = a in + let bx, by = b in + let dx = compare ax bx in + if (dx == 0) then compare ay by + else dx +end + +module PPMap = Map.Make(PP) + +type vstate = Unvisited | Processed | Visited + +let get_loop_edges code entry = + let visited = ref (PTree.map (fun n i -> Unvisited) code) in + let is_loop_edge = ref PPMap.empty + in let rec dfs_visit code from = function + | [] -> () + | node :: ln -> + match (get_some @@ PTree.get node !visited) with + | Visited -> () + | Processed -> begin + let from_node = get_some from in + is_loop_edge := PPMap.add (from_node, node) true !is_loop_edge; + visited := PTree.set node Visited !visited + end + | Unvisited -> begin + visited := PTree.set node Processed !visited; + let bb = get_some @@ PTree.get node code in + let next_visits = (match (last_element bb) with + | Lop _ | Lload _ | Lgetstack _ | Lsetstack _ | Lstore _ | Lcall _ + | Lbuiltin _ -> assert false + | Ltailcall _ | Lreturn -> [] + | Lbranch n -> [n] + | Lcond (_, _, ifso, ifnot) -> [ifso; ifnot] + | Ljumptable(_, ln) -> ln + ) in dfs_visit code (Some node) next_visits; + visited := PTree.set node Visited !visited; + dfs_visit code from ln + end + in begin + dfs_visit code None [entry]; + !is_loop_edge + end + +let ppmap_is_true pp ppmap = PPMap.mem pp ppmap && PPMap.find pp ppmap + +module Int = struct + type t = int + let compare x y = compare x y +end + +module ISet = Set.Make(Int) + +let construct_depmap code entry fs = + let is_loop_edge = get_loop_edges code entry in + let visited = ref (PTree.map (fun n i -> false) code) in + let depmap = Array.map (fun e -> ISet.empty) fs in + let find_index_of_node n = + let index = ref 0 in + begin + Array.iteri (fun i s -> + match List.find_opt (fun e -> e == n) s with + | Some _ -> index := i + | None -> () + ) fs; + !index + end + in let rec dfs_visit code = function + | [] -> () + | node :: ln -> + match (get_some @@ PTree.get node !visited) with + | true -> () + | false -> begin + visited := PTree.set node true !visited; + let bb = get_some @@ PTree.get node code in + let next_visits = + match (last_element bb) with + | Ltailcall _ | Lreturn -> [] + | Lbranch n -> [n] + | Lcond (_, _, ifso, ifnot) -> begin + (if not (ppmap_is_true (node, ifso) is_loop_edge) then + let in_index_fs = find_index_of_node node in + let out_index_fs = find_index_of_node ifso in + depmap.(out_index_fs) <- ISet.add in_index_fs depmap.(out_index_fs) + else + ()); + [ifso; ifnot] + end + | Ljumptable(_, ln) -> begin + let in_index_fs = find_index_of_node node in + List.iter (fun n -> + if not (ppmap_is_true (node, n) is_loop_edge) then + let out_index_fs = find_index_of_node n in + depmap.(out_index_fs) <- ISet.add in_index_fs depmap.(out_index_fs) + else + () + ) ln; + ln + end + (* end of bblocks should not be another value than one of the above *) + | _ -> failwith "last_element gave an invalid output" + in dfs_visit code next_visits + end + in begin + dfs_visit code [entry]; + depmap + end + +let order_sequences code entry fs = + let fs_a = Array.of_list fs in + let depmap = construct_depmap code entry fs_a in + Array.iter (fun _ -> ()) depmap; + (* algo *) + fs let enumerate_aux_trace f reach = - let fs = forward_sequences f.fn_code f.fn_entrypoint - in let ofs = order_sequences fs - in List.flatten ofs + let code = f.fn_code in + let entry = f.fn_entrypoint in + let fs = forward_sequences code entry in + let ofs = order_sequences code entry fs in + List.flatten ofs let enumerate_aux f reach = if !Clflags.option_ftracelinearize then enumerate_aux_trace f reach -- cgit From 86d593820f481b893c7ca00d39b2ac73a6e73aa0 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 3 Mar 2020 19:01:25 +0100 Subject: same version as in dm-cse2 --- Makefile | 1 + backend/CSE2proof.v | 48 ++++++++++++++++++++++++------------------------ 2 files changed, 25 insertions(+), 24 deletions(-) diff --git a/Makefile b/Makefile index 5566cf57..2cd40800 100644 --- a/Makefile +++ b/Makefile @@ -86,6 +86,7 @@ BACKEND=\ ValueDomain.v ValueAOp.v ValueAnalysis.v \ ConstpropOp.v Constprop.v ConstpropOpproof.v Constpropproof.v \ CSEdomain.v CombineOp.v CSE.v CombineOpproof.v CSEproof.v \ + CSE2deps.v CSE2depsproof.v \ CSE2.v CSE2proof.v \ NeedDomain.v NeedOp.v Deadcode.v Deadcodeproof.v \ Unusedglob.v Unusedglobproof.v \ diff --git a/backend/CSE2proof.v b/backend/CSE2proof.v index 7e1dd430..9f55846d 100644 --- a/backend/CSE2proof.v +++ b/backend/CSE2proof.v @@ -55,9 +55,9 @@ Definition sem_sym_val sym rs (v : option val) : Prop := match eval_addressing genv sp addr rs##args with | Some a => match Mem.loadv chunk m a with | Some dat => v = Some dat - | None => v = None \/ v = Some (default_notrap_load_value chunk) + | None => v = None \/ v = Some Vundef end - | None => v = None \/ v = Some (default_notrap_load_value chunk) + | None => v = None \/ v = Some Vundef end end. @@ -404,9 +404,9 @@ Lemma find_load_sound : match eval_addressing genv sp addr rs##args with | Some a => match Mem.loadv chunk m a with | Some dat => rs#src = dat - | None => rs#src = default_notrap_load_value chunk + | None => rs#src = Vundef end - | None => rs#src = default_notrap_load_value chunk + | None => rs#src = Vundef end. Proof. intros until rs. @@ -421,9 +421,9 @@ Proof. match eval_addressing genv sp addr rs##args with | Some a => match Mem.loadv chunk m a with | Some dat => rs#src = dat - | None => rs#src = default_notrap_load_value chunk + | None => rs#src = Vundef end - | None => rs#src = default_notrap_load_value chunk + | None => rs#src = Vundef end end -> fold_left @@ -433,9 +433,9 @@ Proof. match eval_addressing genv sp addr rs##args with | Some a => match Mem.loadv chunk m a with | Some dat => rs#src = dat - | None => rs#src = default_notrap_load_value chunk + | None => rs#src = Vundef end - | None => rs#src = default_notrap_load_value chunk + | None => rs#src = Vundef end) as REC. { @@ -523,7 +523,7 @@ Lemma find_load_notrap1_sound' : sem_rel rel rs -> find_load rel chunk addr args = Some src -> eval_addressing genv sp addr rs##args = None -> - rs # src = (default_notrap_load_value chunk). + rs # src = Vundef. Proof. intros until rs. intros REL FINDLOAD ADDR. pose proof (find_load_sound rel chunk addr src args rs REL FINDLOAD) as Z. @@ -543,7 +543,7 @@ Lemma find_load_notrap2_sound' : find_load rel chunk addr args = Some src -> eval_addressing genv sp addr rs##args = Some a -> Mem.loadv chunk m a = None -> - rs # src = (default_notrap_load_value chunk). + rs # src = Vundef. Proof. intros until a. intros REL FINDLOAD ADDR LOAD. pose proof (find_load_sound rel chunk addr src args rs REL FINDLOAD) as Z. @@ -689,11 +689,11 @@ Lemma load2_notrap1_sound : sem_rel rel rs -> not (In dst args) -> eval_addressing genv sp addr (rs ## args) = None -> - sem_rel (load2 chunk addr dst args rel) (rs # dst <- (default_notrap_load_value chunk)). + sem_rel (load2 chunk addr dst args rel) (rs # dst <- Vundef). Proof. intros until rs. intros REL NOT_IN ADDR x. - pose proof (kill_reg_sound rel dst rs (default_notrap_load_value chunk) REL x) as KILL. + pose proof (kill_reg_sound rel dst rs Vundef REL x) as KILL. unfold load2. destruct (peq x dst). { @@ -726,11 +726,11 @@ Lemma load2_notrap2_sound : not (In dst args) -> eval_addressing genv sp addr (rs ## args) = Some a -> Mem.loadv chunk m a = None -> - sem_rel (load2 chunk addr dst args rel) (rs # dst <- (default_notrap_load_value chunk)). + sem_rel (load2 chunk addr dst args rel) (rs # dst <- Vundef). Proof. intros until a. intros REL NOT_IN ADDR LOAD x. - pose proof (kill_reg_sound rel dst rs (default_notrap_load_value chunk) REL x) as KILL. + pose proof (kill_reg_sound rel dst rs Vundef REL x) as KILL. unfold load2. destruct (peq x dst). { @@ -784,7 +784,7 @@ Lemma load1_notrap1_sound : forall rs : regset, sem_rel rel rs -> eval_addressing genv sp addr (rs ## args) = None -> - sem_rel (load1 chunk addr dst args rel) (rs # dst <- (default_notrap_load_value chunk)). + sem_rel (load1 chunk addr dst args rel) (rs # dst <- Vundef). Proof. intros until rs. intros REL ADDR LOAD. @@ -807,7 +807,7 @@ Lemma load1_notrap2_sound : sem_rel rel rs -> eval_addressing genv sp addr (rs ## args) = Some a -> Mem.loadv chunk m a = None -> - sem_rel (load1 chunk addr dst args rel) (rs # dst <- (default_notrap_load_value chunk)). + sem_rel (load1 chunk addr dst args rel) (rs # dst <- Vundef). Proof. intros until a. intros REL ADDR LOAD. @@ -841,9 +841,9 @@ Proof. assert (match eval_addressing genv sp addr rs## (map (forward_move rel) args) with | Some a => match Mem.loadv chunk m a with | Some dat => rs#src = dat - | None => rs#src = default_notrap_load_value chunk + | None => rs#src = Vundef end - | None => rs#src = default_notrap_load_value chunk + | None => rs#src = Vundef end) as FIND_LOAD. { apply (find_load_sound rel); trivial. @@ -869,7 +869,7 @@ Lemma load_notrap1_sound : forall rs : regset, sem_rel rel rs -> eval_addressing genv sp addr (rs ## args) = None -> - sem_rel (load chunk addr dst args rel) (rs # dst <- (default_notrap_load_value chunk)). + sem_rel (load chunk addr dst args rel) (rs # dst <- Vundef). Proof. intros until rs. intros REL ADDR. @@ -879,9 +879,9 @@ Proof. assert (match eval_addressing genv sp addr rs## (map (forward_move rel) args) with | Some a => match Mem.loadv chunk m a with | Some dat => rs#src = dat - | None => rs#src = default_notrap_load_value chunk + | None => rs#src = Vundef end - | None => rs#src = default_notrap_load_value chunk + | None => rs#src = Vundef end) as FIND_LOAD. { apply (find_load_sound rel); trivial. @@ -906,7 +906,7 @@ Lemma load_notrap2_sound : sem_rel rel rs -> eval_addressing genv sp addr (rs ## args) = Some a -> Mem.loadv chunk m a = None -> - sem_rel (load chunk addr dst args rel) (rs # dst <- (default_notrap_load_value chunk)). + sem_rel (load chunk addr dst args rel) (rs # dst <- Vundef). Proof. intros until a. intros REL ADDR. @@ -916,9 +916,9 @@ Proof. assert (match eval_addressing genv sp addr rs## (map (forward_move rel) args) with | Some a => match Mem.loadv chunk m a with | Some dat => rs#src = dat - | None => rs#src = default_notrap_load_value chunk + | None => rs#src = Vundef end - | None => rs#src = default_notrap_load_value chunk + | None => rs#src = Vundef end) as FIND_LOAD. { apply (find_load_sound rel); trivial. -- cgit From 668912983cd68f5f233bfd3af280f911a8522a84 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 3 Mar 2020 19:18:44 +0100 Subject: fix for ppc --- powerpc/CSE2depsproof.v | 36 ++++++++++++++++++++++-------------- 1 file changed, 22 insertions(+), 14 deletions(-) diff --git a/powerpc/CSE2depsproof.v b/powerpc/CSE2depsproof.v index a047b02a..fdded9b6 100644 --- a/powerpc/CSE2depsproof.v +++ b/powerpc/CSE2depsproof.v @@ -24,6 +24,14 @@ Proof. destruct Archi.ptr64; reflexivity. Qed. +Lemma ptrofs_max_unsigned : + Ptrofs.max_unsigned = if Archi.ptr64 then 18446744073709551615 else 4294967295. +Proof. + unfold Ptrofs.max_unsigned. + rewrite ptrofs_modulus. + destruct Archi.ptr64; reflexivity. +Qed. + Section SOUNDNESS. Variable F V : Type. Variable genv: Genv.t F V. @@ -38,17 +46,17 @@ Section MEMORY_WRITE. Hypothesis STORE : Mem.storev chunkw m addrw valw = Some m2. Section INDEXED_AWAY. - Variable ofsw ofsr : ptrofs. + Variable ofsw ofsr : int. Hypothesis ADDRW : eval_addressing genv sp (Aindexed ofsw) (base :: nil) = Some addrw. Hypothesis ADDRR : eval_addressing genv sp (Aindexed ofsr) (base :: nil) = Some addrr. Lemma load_store_away1 : - forall RANGEW : 0 <= Ptrofs.unsigned ofsw <= Ptrofs.modulus - largest_size_chunk, - forall RANGER : 0 <= Ptrofs.unsigned ofsr <= Ptrofs.modulus - largest_size_chunk, - forall SWAPPABLE : Ptrofs.unsigned ofsw + size_chunk chunkw <= Ptrofs.unsigned ofsr - \/ Ptrofs.unsigned ofsr + size_chunk chunkr <= Ptrofs.unsigned ofsw, + forall RANGEW : 0 <= Int.unsigned ofsw <= Ptrofs.modulus - largest_size_chunk, + forall RANGER : 0 <= Int.unsigned ofsr <= Ptrofs.modulus - largest_size_chunk, + forall SWAPPABLE : Int.unsigned ofsw + size_chunk chunkw <= Int.unsigned ofsr + \/ Int.unsigned ofsr + size_chunk chunkr <= Int.unsigned ofsw, Mem.loadv chunkr m2 addrr = Mem.loadv chunkr m addrr. Proof. intros. @@ -66,19 +74,19 @@ Section MEMORY_WRITE. exact STORE. right. - all: try (destruct (Ptrofs.unsigned_add_either i ofsr) as [OFSR | OFSR]; + all: try (destruct (Ptrofs.unsigned_add_either i (Ptrofs.of_int ofsr)) as [OFSR | OFSR]; rewrite OFSR). - all: try (destruct (Ptrofs.unsigned_add_either i ofsw) as [OFSW | OFSW]; + all: try (destruct (Ptrofs.unsigned_add_either i (Ptrofs.of_int ofsw)) as [OFSW | OFSW]; rewrite OFSW). - - all: try rewrite ptrofs_modulus in *. - - all: repeat rewrite Ptrofs.unsigned_repr by (change Ptrofs.max_unsigned with 4294967295; lia). - all: intuition lia. + all: unfold Ptrofs.of_int. + + all: repeat rewrite Ptrofs.unsigned_repr by (unfold Ptrofs.max_unsigned; rewrite ptrofs_modulus; destruct Archi.ptr64; lia). + all: repeat rewrite ptrofs_modulus. + all: destruct Archi.ptr64; intuition lia. Qed. Theorem load_store_away : - can_swap_accesses_ofs (Ptrofs.unsigned ofsr) chunkr (Ptrofs.unsigned ofsw) chunkw = true -> + can_swap_accesses_ofs (Int.unsigned ofsr) chunkr (Int.unsigned ofsw) chunkw = true -> Mem.loadv chunkr m2 addrr = Mem.loadv chunkr m addrr. Proof. intro SWAP. @@ -117,7 +125,7 @@ Proof. simpl in OVERLAP. destruct (peq base base'). 2: discriminate. subst base'. - destruct (can_swap_accesses_ofs (Ptrofs.unsigned i0) chunk' (Ptrofs.unsigned i) chunk) eqn:SWAP. + destruct (can_swap_accesses_ofs (Int.unsigned i0) chunk' (Int.unsigned i) chunk) eqn:SWAP. 2: discriminate. simpl in *. eapply load_store_away with (F:=F) (V:=V) (genv:=genv) (sp:=sp); eassumption. -- cgit From 039b532ae972292ec2f726505422afd49569b738 Mon Sep 17 00:00:00 2001 From: Bernhard Schommer Date: Wed, 4 Mar 2020 18:18:40 +0100 Subject: Include typedef name in error message (#228) In case of redefinition of a typedef name with a different type. --- cparser/Elab.ml | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/cparser/Elab.ml b/cparser/Elab.ml index f60e15a3..9e17cb7e 100644 --- a/cparser/Elab.ml +++ b/cparser/Elab.ml @@ -2428,8 +2428,8 @@ let enter_typedef loc env sto (s, ty, init) = env end else begin - error loc "typedef redefinition with different types (%a vs %a)" - (print_typ env) ty (print_typ env) ty'; + error loc "redefinition of typedef '%s' with different type (%a vs %a)" + s (print_typ env) ty (print_typ env) ty'; env end | _ -> -- cgit From 5eb64cdae7deeaef774fdead6f3ce6e4108a3256 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Fri, 6 Mar 2020 15:59:44 +0100 Subject: [UNTESTED] Sequence ordering --- backend/Linearizeaux.ml | 30 +++++++++++++++++++++++++++--- 1 file changed, 27 insertions(+), 3 deletions(-) diff --git a/backend/Linearizeaux.ml b/backend/Linearizeaux.ml index 58d7558b..ce518dbb 100644 --- a/backend/Linearizeaux.ml +++ b/backend/Linearizeaux.ml @@ -370,9 +370,33 @@ let construct_depmap code entry fs = let order_sequences code entry fs = let fs_a = Array.of_list fs in let depmap = construct_depmap code entry fs_a in - Array.iter (fun _ -> ()) depmap; - (* algo *) - fs + let fs_evaluated = Array.map (fun e -> false) fs_a in + let ordered_fs = ref [] in + let evaluate s_id = + begin + assert (not fs_evaluated.(s_id)); + ordered_fs := fs_a.(s_id) :: !ordered_fs; + fs_evaluated.(s_id) <- true; + Array.iteri (fun i deps -> + depmap.(i) <- ISet.remove s_id deps + ) depmap + end + in let select_next () = + let selected_id = ref (-1) in + begin + Array.iteri (fun i deps -> + if !selected_id == -1 && deps == ISet.empty && not fs_evaluated.(i) + then selected_id := i + ) depmap; + !selected_id + end + in begin + while List.length !ordered_fs != List.length fs do + let next_id = select_next () in + evaluate next_id + done; + !ordered_fs + end let enumerate_aux_trace f reach = let code = f.fn_code in -- cgit From 8300321b348a6b416a8e0498ecebf944697d0641 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Fri, 6 Mar 2020 16:20:22 +0100 Subject: Adding debug info in Linearizeaux --- backend/Linearizeaux.ml | 12 ++++++++++++ 1 file changed, 12 insertions(+) diff --git a/backend/Linearizeaux.ml b/backend/Linearizeaux.ml index ce518dbb..b609b57a 100644 --- a/backend/Linearizeaux.ml +++ b/backend/Linearizeaux.ml @@ -367,6 +367,16 @@ let construct_depmap code entry fs = depmap end +let print_sequence s = + Printf.printf "["; + List.iter (fun n -> Printf.printf "%d, " (P.to_int n)) s; + Printf.printf "]\n" + +let print_ssequence ofs = + Printf.printf "["; + List.iter (fun s -> print_sequence s) ofs; + Printf.printf "]\n" + let order_sequences code entry fs = let fs_a = Array.of_list fs in let depmap = construct_depmap code entry fs_a in @@ -391,10 +401,12 @@ let order_sequences code entry fs = !selected_id end in begin + print_ssequence fs; while List.length !ordered_fs != List.length fs do let next_id = select_next () in evaluate next_id done; + print_ssequence !ordered_fs; !ordered_fs end -- cgit From 7b85e3b00e500c5d65cf2df1adeae8ecd7d3e88d Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Sat, 7 Mar 2020 06:50:50 +0100 Subject: removing warnings on hints in core --- mppa_k1c/Asmblockdeps.v | 6 +++--- mppa_k1c/Asmblockgenproof1.v | 4 ++-- mppa_k1c/PostpassSchedulingproof.v | 2 +- mppa_k1c/abstractbb/AbstractBasicBlocksDef.v | 2 +- mppa_k1c/abstractbb/ImpSimuTest.v | 14 +++++++------- mppa_k1c/abstractbb/Impure/ImpHCons.v | 4 ++-- mppa_k1c/abstractbb/Parallelizability.v | 8 ++++---- mppa_k1c/abstractbb/SeqSimuTheory.v | 11 ++++------- mppa_k1c/lib/Asmblockgenproof0.v | 4 ++-- mppa_k1c/lib/ForwardSimulationBlock.v | 6 +++--- mppa_k1c/lib/Machblockgenproof.v | 20 ++++++++++---------- 11 files changed, 39 insertions(+), 42 deletions(-) diff --git a/mppa_k1c/Asmblockdeps.v b/mppa_k1c/Asmblockdeps.v index 02f9141b..bc9f2584 100644 --- a/mppa_k1c/Asmblockdeps.v +++ b/mppa_k1c/Asmblockdeps.v @@ -1005,7 +1005,7 @@ Theorem bisimu_par_wio_basic ge fn rsr rsw mr mw sr sw bi: Proof. (* a little tactic to automate reasoning on preg_eq *) -Local Hint Resolve not_eq_sym ppos_pmem_discr ppos_discr. +Local Hint Resolve not_eq_sym ppos_pmem_discr ppos_discr: core. Local Ltac preg_eq_discr r rd := destruct (preg_eq r rd); try (subst r; rewrite assign_eq, Pregmap.gss; auto); rewrite (assign_diff _ (#rd) (#r) _); auto; @@ -1053,7 +1053,7 @@ Local Ltac preg_eq_discr r rd := preg_eq_discr r rd0. } (* Load Octuple word *) - + Local Hint Resolve not_eq_sym ppos_pmem_discr ppos_discr. + + Local Hint Resolve not_eq_sym ppos_pmem_discr ppos_discr: core. unfold parexec_load_o_offset. destruct (gpreg_o_expand rd) as [[[rd0 rd1] rd2] rd3]; destruct Ge; simpl. rewrite H0, H. @@ -1423,7 +1423,7 @@ Section SECT_BBLOCK_EQUIV. Variable Ge: genv. -Local Hint Resolve trans_state_match. +Local Hint Resolve trans_state_match: core. Lemma bblock_simu_reduce: forall p1 p2 ge fn, diff --git a/mppa_k1c/Asmblockgenproof1.v b/mppa_k1c/Asmblockgenproof1.v index ecb4629b..d3c2008f 100644 --- a/mppa_k1c/Asmblockgenproof1.v +++ b/mppa_k1c/Asmblockgenproof1.v @@ -859,7 +859,7 @@ Proof. destruct cmp; discriminate. Qed. -Local Hint Resolve Val_cmpu_bool_correct Val_cmplu_bool_correct. +Local Hint Resolve Val_cmpu_bool_correct Val_cmplu_bool_correct: core. Lemma transl_cbranch_correct_1: forall cond args lbl k c m ms b sp rs m' tbb, @@ -1163,7 +1163,7 @@ Proof. split; intros; Simpl. Qed. -Local Hint Resolve Val_cmpu_correct Val_cmplu_correct. +Local Hint Resolve Val_cmpu_correct Val_cmplu_correct: core. Lemma transl_condimm_int32u_correct: forall cmp rd r1 n k rs m, diff --git a/mppa_k1c/PostpassSchedulingproof.v b/mppa_k1c/PostpassSchedulingproof.v index fbb06c9b..3b123c75 100644 --- a/mppa_k1c/PostpassSchedulingproof.v +++ b/mppa_k1c/PostpassSchedulingproof.v @@ -567,7 +567,7 @@ Proof. unfold builtin_alone in H0. erewrite H0; eauto. Qed. -Local Hint Resolve verified_schedule_nob_checks_alls_bundles. +Local Hint Resolve verified_schedule_nob_checks_alls_bundles: core. Lemma verified_schedule_checks_alls_bundles bb lb bundle: verified_schedule bb = OK lb -> diff --git a/mppa_k1c/abstractbb/AbstractBasicBlocksDef.v b/mppa_k1c/abstractbb/AbstractBasicBlocksDef.v index 5c94d435..cf46072f 100644 --- a/mppa_k1c/abstractbb/AbstractBasicBlocksDef.v +++ b/mppa_k1c/abstractbb/AbstractBasicBlocksDef.v @@ -403,7 +403,7 @@ Proof. * eapply H2; eauto. intros; eapply H0; eauto. rewrite rev_append_rev, in_app_iff; auto. * intros; eapply H0; eauto. rewrite rev_append_rev, in_app_iff, <- in_rev; auto. Qed. -Local Hint Resolve app_fail_allvalid_correct. +Local Hint Resolve app_fail_allvalid_correct: core. Lemma app_fail_correct l pt t1 t2: match_pt t1 pt -> diff --git a/mppa_k1c/abstractbb/ImpSimuTest.v b/mppa_k1c/abstractbb/ImpSimuTest.v index ea55b735..7a77ec15 100644 --- a/mppa_k1c/abstractbb/ImpSimuTest.v +++ b/mppa_k1c/abstractbb/ImpSimuTest.v @@ -304,12 +304,12 @@ Proof. rewrite <- EQT; eauto. + exploit smem_valid_set_decompose_1; eauto. - clear DM0. unfold hsmem_post_eval, hsmem_post_eval in * |- *; simpl. - Local Hint Resolve smem_valid_set_decompose_1. + Local Hint Resolve smem_valid_set_decompose_1: core. intros; case (R.eq_dec x x0). + intros; subst; rewrite !Dict.set_spec_eq; simpl; eauto. + intros; rewrite !Dict.set_spec_diff; simpl; eauto. Qed. -Local Hint Resolve naive_set_correct. +Local Hint Resolve naive_set_correct: core. Definition equiv_hsmem ge (hd1 hd2: hsmem) := (forall m, allvalid ge hd1.(hpre) m <-> allvalid ge hd2.(hpre) m) @@ -523,7 +523,7 @@ Lemma hinst_smem_correct i: forall hd hod, WHEN hinst_smem i hd hod ~> hd' THEN forall ge od d, smem_model ge od hod -> smem_model ge d hd -> (forall m, smem_valid ge d m -> smem_valid ge od m) -> smem_model ge (inst_smem i d od) hd'. Proof. - Local Hint Resolve smem_valid_set_proof. + Local Hint Resolve smem_valid_set_proof: core. induction i; simpl; wlp_simplify; eauto 15 with wlp. Qed. Global Opaque hinst_smem. @@ -563,7 +563,7 @@ Definition bblock_hsmem: bblock -> ?? hsmem Lemma bblock_hsmem_correct p: WHEN bblock_hsmem p ~> hd THEN forall ge, smem_model ge (bblock_smem p) hd. Proof. - Local Hint Resolve hsmem_empty_correct. + Local Hint Resolve hsmem_empty_correct: core. wlp_simplify. Qed. Global Opaque bblock_hsmem. @@ -775,7 +775,7 @@ Proof. intro H; erewrite <- list_term_eval_set_hid; rewrite H. apply list_term_eval_set_hid. Qed. -Local Hint Resolve term_eval_set_hid_equiv list_term_eval_set_hid_equiv. +Local Hint Resolve term_eval_set_hid_equiv list_term_eval_set_hid_equiv: core. Program Definition bblock_simu_test (p1 p2: bblock): ?? bool := DO log <~ count_logger ();; @@ -802,7 +802,7 @@ Obligation 2. wlp_simplify. Qed. -Local Hint Resolve g_bblock_simu_test_correct. +Local Hint Resolve g_bblock_simu_test_correct: core. Theorem bblock_simu_test_correct p1 p2: WHEN bblock_simu_test p1 p2 ~> b THEN b=true -> forall ge, bblock_simu ge p1 p2. @@ -1123,7 +1123,7 @@ Definition get {A} (d:t A) (x:R.t): option A Definition set {A} (d:t A) (x:R.t) (v:A): t A := PositiveMap.add x v d. -Local Hint Unfold PositiveMap.E.eq. +Local Hint Unfold PositiveMap.E.eq: core. Lemma set_spec_eq A d x (v: A): get (set d x v) x = Some v. diff --git a/mppa_k1c/abstractbb/Impure/ImpHCons.v b/mppa_k1c/abstractbb/Impure/ImpHCons.v index d8002375..637116cc 100644 --- a/mppa_k1c/abstractbb/Impure/ImpHCons.v +++ b/mppa_k1c/abstractbb/Impure/ImpHCons.v @@ -95,7 +95,7 @@ Proof. wlp_simplify. Qed. Global Opaque assert_list_incl. -Hint Resolve assert_list_incl_correct. +Hint Resolve assert_list_incl_correct: wlp. End Sets. @@ -165,7 +165,7 @@ Lemma hConsV_correct A (hasheq: A -> A -> ?? bool): (forall x y, WHEN hasheq x y ~> b THEN b=true -> x=y) -> forall x, WHEN hco.(hC) x ~> x' THEN x.(hdata).(data)=x'.(data). Proof. - Local Hint Resolve f_equal2. + Local Hint Resolve f_equal2: core. wlp_simplify. exploit H; eauto. + wlp_simplify. diff --git a/mppa_k1c/abstractbb/Parallelizability.v b/mppa_k1c/abstractbb/Parallelizability.v index 22809095..30904b5d 100644 --- a/mppa_k1c/abstractbb/Parallelizability.v +++ b/mppa_k1c/abstractbb/Parallelizability.v @@ -332,7 +332,7 @@ Fixpoint bblock_wframe(p:bblock): list R.t := | i::p' => (inst_wframe i)++(bblock_wframe p') end. -Local Hint Resolve Permutation_app_head Permutation_app_tail Permutation_app_comm. +Local Hint Resolve Permutation_app_head Permutation_app_tail Permutation_app_comm: core. Lemma bblock_wframe_Permutation p p': Permutation p p' -> Permutation (bblock_wframe p) (bblock_wframe p'). @@ -620,7 +620,7 @@ Include ParallelizablityChecking L. Section PARALLEL2. Variable ge: genv. -Local Hint Resolve S.empty_match_frame S.add_match_frame S.union_match_frame S.is_disjoint_match_frame. +Local Hint Resolve S.empty_match_frame S.add_match_frame S.union_match_frame S.is_disjoint_match_frame: core. (** Now, refinement of each operation toward parallelizable *) @@ -659,14 +659,14 @@ Fixpoint inst_sframe (i: inst): S.t := | a::i' => S.add (fst a) (S.union (exp_sframe (snd a)) (inst_sframe i')) end. -Local Hint Resolve exp_sframe_correct. +Local Hint Resolve exp_sframe_correct: core. Lemma inst_sframe_correct i: S.match_frame (inst_sframe i) (inst_frame i). Proof. induction i as [|[y e] i']; simpl; auto. Qed. -Local Hint Resolve inst_wsframe_correct inst_sframe_correct. +Local Hint Resolve inst_wsframe_correct inst_sframe_correct: core. Fixpoint is_pararec (p: bblock) (wsframe: S.t): bool := match p with diff --git a/mppa_k1c/abstractbb/SeqSimuTheory.v b/mppa_k1c/abstractbb/SeqSimuTheory.v index 649dd083..e234883f 100644 --- a/mppa_k1c/abstractbb/SeqSimuTheory.v +++ b/mppa_k1c/abstractbb/SeqSimuTheory.v @@ -102,9 +102,6 @@ Fixpoint bblock_smem_rec (p: bblock) (d: smem): smem := let d':=inst_smem i d d in bblock_smem_rec p' d' end. -(* -Local Hint Resolve smem_eval_empty. -*) Definition bblock_smem: bblock -> smem := fun p => bblock_smem_rec p smem_empty. @@ -124,7 +121,7 @@ Proof. intros d a H; eapply inst_smem_pre_monotonic; eauto. Qed. -Local Hint Resolve inst_smem_pre_monotonic bblock_smem_pre_monotonic. +Local Hint Resolve inst_smem_pre_monotonic bblock_smem_pre_monotonic: core. Lemma term_eval_exp e (od:smem) m0 old: (forall x, term_eval ge (od x) m0 = Some (old x)) -> @@ -185,7 +182,7 @@ Lemma bblocks_smem_rec_Some_correct1 p m0: forall (m1 m2: mem) (d: smem), (forall x, term_eval ge (d x) m0 = Some (m1 x)) -> forall x, term_eval ge (bblock_smem_rec p d x) m0 = Some (m2 x). Proof. - Local Hint Resolve inst_smem_Some_correct1. + Local Hint Resolve inst_smem_Some_correct1: core. induction p as [ | i p]; simpl; intros m1 m2 d H. - inversion_clear H; eauto. - intros H0 x0. @@ -299,7 +296,7 @@ Lemma block_smem_rec_valid p m0: forall (m1 m2: mem) (d:smem), (forall x, term_eval ge (d x) m0 = Some (m1 x)) -> pre (bblock_smem_rec p d) ge m0. Proof. - Local Hint Resolve inst_valid. + Local Hint Resolve inst_valid: core. induction p as [ | i p]; simpl; intros m1 d H; auto. intros H0 H1. destruct (inst_run ge i m1 m1) eqn: Heqov; eauto. @@ -326,7 +323,7 @@ Theorem bblock_smem_simu p1 p2: smem_simu (bblock_smem p1) (bblock_smem p2) -> bblock_simu ge p1 p2. Proof. - Local Hint Resolve bblock_smem_valid bblock_smem_Some_correct1. + Local Hint Resolve bblock_smem_valid bblock_smem_Some_correct1: core. intros (INCL & EQUIV) m DONTFAIL; unfold smem_valid in * |-. destruct (run ge p1 m) as [m1|] eqn: RUN1; simpl; try congruence. assert (X: forall x, term_eval ge (bblock_smem p1 x) m = Some (m1 x)); eauto. diff --git a/mppa_k1c/lib/Asmblockgenproof0.v b/mppa_k1c/lib/Asmblockgenproof0.v index 940c6563..58455ada 100644 --- a/mppa_k1c/lib/Asmblockgenproof0.v +++ b/mppa_k1c/lib/Asmblockgenproof0.v @@ -414,7 +414,7 @@ Proof. Qed. -Local Hint Resolve code_tail_0 code_tail_S. +Local Hint Resolve code_tail_0 code_tail_S: core. Lemma code_tail_next: forall fn ofs c0, @@ -458,7 +458,7 @@ Proof. omega. Qed. -Local Hint Resolve code_tail_next. +Local Hint Resolve code_tail_next: core. Lemma code_tail_next_int: forall fn ofs bi c, diff --git a/mppa_k1c/lib/ForwardSimulationBlock.v b/mppa_k1c/lib/ForwardSimulationBlock.v index 39dd2234..224eda0a 100644 --- a/mppa_k1c/lib/ForwardSimulationBlock.v +++ b/mppa_k1c/lib/ForwardSimulationBlock.v @@ -21,7 +21,7 @@ Section starN_lemma. Variable L: semantics. -Local Hint Resolve starN_refl starN_step Eapp_assoc. +Local Hint Resolve starN_refl starN_step Eapp_assoc: core. Lemma starN_split n s t s': starN (step L) (globalenv L) n s t s' -> @@ -93,7 +93,7 @@ Hypothesis simu_end_block: (** Introduction d'une sémantique par bloc sur L1 appelée "memoL1" *) -Local Hint Resolve starN_refl starN_step. +Local Hint Resolve starN_refl starN_step: core. Definition follows_in_block (head current: state L1): Prop := dist_end_block head >= dist_end_block current @@ -164,7 +164,7 @@ Inductive is_well_memorized (s s': memostate): Prop := memorized s' = None -> is_well_memorized s s'. -Local Hint Resolve StartBloc MidBloc ExitBloc. +Local Hint Resolve StartBloc MidBloc ExitBloc: core. Definition memoL1 := {| state := memostate; diff --git a/mppa_k1c/lib/Machblockgenproof.v b/mppa_k1c/lib/Machblockgenproof.v index 91be5e2e..0de2df52 100644 --- a/mppa_k1c/lib/Machblockgenproof.v +++ b/mppa_k1c/lib/Machblockgenproof.v @@ -72,7 +72,7 @@ Proof. apply match_states_trans_state. Qed. -Local Hint Resolve match_states_trans_state. +Local Hint Resolve match_states_trans_state: core. Lemma symbols_preserved: forall (s: ident), Genv.find_symbol tge s = Genv.find_symbol ge s. @@ -284,7 +284,7 @@ Proof. Qed. Local Hint Resolve symbols_preserved senv_preserved init_mem_preserved prog_main_preserved functions_translated - parent_sp_preserved. + parent_sp_preserved: core. Definition dist_end_block_code (c: Mach.code) := @@ -299,8 +299,8 @@ Definition dist_end_block (s: Mach.state): nat := | _ => 0 end. -Local Hint Resolve exec_nil_body exec_cons_body. -Local Hint Resolve exec_MBgetstack exec_MBsetstack exec_MBgetparam exec_MBop exec_MBload exec_MBstore. +Local Hint Resolve exec_nil_body exec_cons_body: core. +Local Hint Resolve exec_MBgetstack exec_MBsetstack exec_MBgetparam exec_MBop exec_MBload exec_MBstore: core. Lemma size_add_label l bh: size (add_label l bh) = size bh + 1. Proof. @@ -336,7 +336,7 @@ Proof. omega. Qed. -Local Hint Resolve dist_end_block_code_simu_mid_block. +Local Hint Resolve dist_end_block_code_simu_mid_block: core. Lemma size_nonzero c b bl: @@ -392,8 +392,8 @@ destruct i; congruence. Qed. -Local Hint Resolve Mlabel_is_not_cfi. -Local Hint Resolve MBbasic_is_not_cfi. +Local Hint Resolve Mlabel_is_not_cfi: core. +Local Hint Resolve MBbasic_is_not_cfi: core. Lemma add_to_new_block_is_label i: header (add_to_new_bblock (trans_inst i)) <> nil -> exists l, i = Mlabel l. @@ -408,7 +408,7 @@ Proof. + unfold cfi_bblock in H; simpl in H; congruence. Qed. -Local Hint Resolve Mlabel_is_not_basic. +Local Hint Resolve Mlabel_is_not_basic: core. Lemma trans_code_decompose c: forall b bl, is_trans_code c (b::bl) -> @@ -510,8 +510,8 @@ Proof. rewrite Hs2, Hb2; eauto. Qed. -Local Hint Resolve exec_MBcall exec_MBtailcall exec_MBbuiltin exec_MBgoto exec_MBcond_true exec_MBcond_false exec_MBjumptable exec_MBreturn exec_Some_exit exec_None_exit. -Local Hint Resolve eval_builtin_args_preserved external_call_symbols_preserved find_funct_ptr_same. +Local Hint Resolve exec_MBcall exec_MBtailcall exec_MBbuiltin exec_MBgoto exec_MBcond_true exec_MBcond_false exec_MBjumptable exec_MBreturn exec_Some_exit exec_None_exit: core. +Local Hint Resolve eval_builtin_args_preserved external_call_symbols_preserved find_funct_ptr_same: core. Lemma match_states_concat_trans_code st f sp c rs m h: -- cgit From 1df2fadbf5ab0687d2aac52f3a83bbe071c25139 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Mon, 9 Mar 2020 08:07:04 +0100 Subject: removing some coqc 8.10 warnings --- backend/Duplicateproof.v | 2 +- mppa_k1c/lib/Machblockgen.v | 8 ++++---- 2 files changed, 5 insertions(+), 5 deletions(-) diff --git a/backend/Duplicateproof.v b/backend/Duplicateproof.v index a8e9b16b..466b4b75 100644 --- a/backend/Duplicateproof.v +++ b/backend/Duplicateproof.v @@ -378,7 +378,7 @@ Theorem step_simulation: step tge s2 t s2' /\ match_states s1' s2'. Proof. - Local Hint Resolve transf_fundef_correct. + Local Hint Resolve transf_fundef_correct: core. induction 1; intros; inv MS. (* Inop *) - eapply dupmap_correct in DUPLIC; eauto. diff --git a/mppa_k1c/lib/Machblockgen.v b/mppa_k1c/lib/Machblockgen.v index a65b218f..2ba42814 100644 --- a/mppa_k1c/lib/Machblockgen.v +++ b/mppa_k1c/lib/Machblockgen.v @@ -105,7 +105,7 @@ Inductive is_end_block: Machblock_inst -> code -> Prop := | End_basic bi bh bl: header bh <> nil -> is_end_block (MB_basic bi) (bh::bl) | End_cfi cfi bl: bl <> nil -> is_end_block (MB_cfi cfi) bl. -Local Hint Resolve End_empty End_basic End_cfi. +Local Hint Resolve End_empty End_basic End_cfi: core. Inductive is_trans_code: Mach.code -> code -> Prop := | Tr_nil: is_trans_code nil nil @@ -123,7 +123,7 @@ Inductive is_trans_code: Mach.code -> code -> Prop := header bh = nil -> is_trans_code (i::c) (add_basic bi bh::bl). -Local Hint Resolve Tr_nil Tr_end_block. +Local Hint Resolve Tr_nil Tr_end_block: core. Lemma add_to_code_is_trans_code i c bl: is_trans_code c bl -> @@ -145,7 +145,7 @@ Proof. rewrite <- Heqti. eapply End_cfi. congruence. Qed. -Local Hint Resolve add_to_code_is_trans_code. +Local Hint Resolve add_to_code_is_trans_code: core. Lemma trans_code_is_trans_code_rev c1: forall c2 mbi, is_trans_code c2 mbi -> @@ -185,7 +185,7 @@ Proof. exists mbi1. split; congruence. Qed. -Local Hint Resolve trans_code_is_trans_code. +Local Hint Resolve trans_code_is_trans_code: core. Theorem is_trans_code_inv c bl: is_trans_code c bl <-> bl=(trans_code c). Proof. -- cgit From f2a5f59fca7be2c9b31a18e31c66cd21819fce56 Mon Sep 17 00:00:00 2001 From: Sylvain Boulmé Date: Mon, 9 Mar 2020 08:25:59 +0100 Subject: removing more coq8.10 warnings --- mppa_k1c/Asmblockdeps.v | 2 +- mppa_k1c/Asmblockgen.v | 2 ++ mppa_k1c/Asmblockgenproof1.v | 2 ++ mppa_k1c/Asmvliw.v | 6 +++++- 4 files changed, 10 insertions(+), 2 deletions(-) diff --git a/mppa_k1c/Asmblockdeps.v b/mppa_k1c/Asmblockdeps.v index bc9f2584..01eda623 100644 --- a/mppa_k1c/Asmblockdeps.v +++ b/mppa_k1c/Asmblockdeps.v @@ -339,7 +339,7 @@ Proof. } destruct (Mem.load _ m1 _ _) in *; destruct (Mem.load _ m0 _ _) in *; congruence. Qed. - + Definition goto_label_deps (f: function) (lbl: label) (vpc: val) := match label_pos lbl 0 (fn_blocks f) with | None => None diff --git a/mppa_k1c/Asmblockgen.v b/mppa_k1c/Asmblockgen.v index 50637723..36269954 100644 --- a/mppa_k1c/Asmblockgen.v +++ b/mppa_k1c/Asmblockgen.v @@ -28,6 +28,8 @@ Require Import Chunks. Local Open Scope string_scope. Local Open Scope error_monad_scope. +Import PArithCoercions. + (** The code generation functions take advantage of several characteristics of the [Mach] code generated by earlier passes of the compiler, mostly that argument and result registers are of the correct diff --git a/mppa_k1c/Asmblockgenproof1.v b/mppa_k1c/Asmblockgenproof1.v index d3c2008f..5b44ddaa 100644 --- a/mppa_k1c/Asmblockgenproof1.v +++ b/mppa_k1c/Asmblockgenproof1.v @@ -23,6 +23,8 @@ Require Import Op Locations Machblock Conventions. Require Import Asmblock Asmblockgen Asmblockgenproof0 Asmblockprops. Require Import Chunks. +Import PArithCoercions. + (** Decomposition of integer constants. *) Lemma make_immed32_sound: diff --git a/mppa_k1c/Asmvliw.v b/mppa_k1c/Asmvliw.v index e042d95a..946007c1 100644 --- a/mppa_k1c/Asmvliw.v +++ b/mppa_k1c/Asmvliw.v @@ -555,6 +555,8 @@ Inductive ar_instruction : Type := | PArithARRI64 (i: arith_name_arri64) (rd rs: ireg) (imm: int64) . +Module PArithCoercions. + Coercion PArithR: arith_name_r >-> Funclass. Coercion PArithRR: arith_name_rr >-> Funclass. Coercion PArithRI32: arith_name_ri32 >-> Funclass. @@ -569,6 +571,8 @@ Coercion PArithARR: arith_name_arr >-> Funclass. Coercion PArithARRI32: arith_name_arri32 >-> Funclass. Coercion PArithARRI64: arith_name_arri64 >-> Funclass. +End PArithCoercions. + Inductive basic : Type := | PArith (i: ar_instruction) | PLoad (i: ld_instruction) @@ -1709,7 +1713,7 @@ Proof. Qed. -Local Hint Resolve parexec_bblock_write_in_order. +Local Hint Resolve parexec_bblock_write_in_order: core. Lemma det_parexec_write_in_order f b rs m rs' m': det_parexec f b rs m rs' m' -> parexec_wio_bblock f b rs m = Next rs' m'. -- cgit From 7794bebc14750c5d8116f54cabe143231ef60308 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Mon, 9 Mar 2020 11:27:38 +0100 Subject: Linearizeaux: Fixed bug where the output list was in reverse order --- backend/Linearizeaux.ml | 24 ++++++++++++++++++++---- 1 file changed, 20 insertions(+), 4 deletions(-) diff --git a/backend/Linearizeaux.ml b/backend/Linearizeaux.ml index b609b57a..7aed5936 100644 --- a/backend/Linearizeaux.ml +++ b/backend/Linearizeaux.ml @@ -312,6 +312,18 @@ end module ISet = Set.Make(Int) +let print_iset s = begin + Printf.printf "{"; + ISet.iter (fun e -> Printf.printf "%d, " e) s; + Printf.printf "}" +end + +let print_depmap dm = begin + Printf.printf "[|"; + Array.iter (fun s -> print_iset s; Printf.printf ", ") dm; + Printf.printf "|]\n" +end + let construct_depmap code entry fs = let is_loop_edge = get_loop_edges code entry in let visited = ref (PTree.map (fun n i -> false) code) in @@ -395,19 +407,23 @@ let order_sequences code entry fs = let selected_id = ref (-1) in begin Array.iteri (fun i deps -> - if !selected_id == -1 && deps == ISet.empty && not fs_evaluated.(i) - then selected_id := i + begin + (* Printf.printf "Deps: "; print_iset deps; Printf.printf "\n"; *) + if !selected_id == -1 && deps == ISet.empty && not fs_evaluated.(i) + then selected_id := i + end ) depmap; !selected_id end in begin + (* Printf.printf "depmap: "; print_depmap depmap; *) print_ssequence fs; while List.length !ordered_fs != List.length fs do let next_id = select_next () in evaluate next_id done; - print_ssequence !ordered_fs; - !ordered_fs + (* print_ssequence (List.rev (!ordered_fs)); *) + List.rev (!ordered_fs) end let enumerate_aux_trace f reach = -- cgit From 510923fcea8ededcd71fc81ae0fb1981bf8b9223 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 9 Mar 2020 12:43:36 +0100 Subject: cycles.h for ARMv7 --- test/monniaux/cycles.h | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/test/monniaux/cycles.h b/test/monniaux/cycles.h index 21541145..4a87299b 100644 --- a/test/monniaux/cycles.h +++ b/test/monniaux/cycles.h @@ -38,6 +38,13 @@ static inline cycle_t get_cycle(void) { return cycles; } +#elif defined (__ARM_ARCH) && (__ARM_ARCH >= 6) +static inline cycle_t get_cycle(void) { + cycle_t cycles; + __asm__ volatile ("mrc p15, 0, %0, c9, c13, 0":"=r" (cycles)); + return cycles; +} + #else static inline cycle_t get_cycle(void) { return 0; } #endif -- cgit From 611d7bec0f35fa5fb017ecf36a17f8967425548e Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Mon, 9 Mar 2020 14:11:48 +0100 Subject: Removing prints in Duplicateaux.ml --- backend/Duplicateaux.ml | 85 ++++++++++++------------------------------------- 1 file changed, 20 insertions(+), 65 deletions(-) diff --git a/backend/Duplicateaux.ml b/backend/Duplicateaux.ml index d0b7129e..e2b5647b 100644 --- a/backend/Duplicateaux.ml +++ b/backend/Duplicateaux.ml @@ -95,52 +95,6 @@ let print_intset s = Printf.printf "}" end -(* FIXME - dominators not working well because the order of dataflow update isn't right *) - (* -let get_dominators code entrypoint = - let bfs_order = bfs code entrypoint - and predecessors = get_predecessors_rtl code - in let doms = ref (PTree.map (fun n i -> PSet.of_list bfs_order) code) - in begin - Printf.printf "BFS: "; - print_intlist bfs_order; - Printf.printf "\n"; - List.iter (fun n -> - let preds = get_some @@ PTree.get n predecessors - and single = PSet.singleton n - in match preds with - | [] -> doms := PTree.set n single !doms - | p::lp -> - let set_p = get_some @@ PTree.get p !doms - and set_lp = List.map (fun p -> get_some @@ PTree.get p !doms) lp - in let inter = List.fold_left PSet.inter set_p set_lp - in let union = PSet.union inter single - in begin - Printf.printf "----------------------------------------\n"; - Printf.printf "n = %d\n" (P.to_int n); - Printf.printf "set_p = "; print_intset set_p; Printf.printf "\n"; - Printf.printf "set_lp = ["; List.iter (fun s -> print_intset s; Printf.printf ", ") set_lp; Printf.printf "]\n"; - Printf.printf "=> inter = "; print_intset inter; Printf.printf "\n"; - Printf.printf "=> union = "; print_intset union; Printf.printf "\n"; - doms := PTree.set n union !doms - end - ) bfs_order; - !doms - end -*) - -let print_dominators dominators = - let domlist = PTree.elements dominators - in begin - Printf.printf "{\n"; - List.iter (fun (n, doms) -> - Printf.printf "\t"; - Printf.printf "%d:" (P.to_int n); - print_intset doms; - Printf.printf "\n" - ) domlist - end - type vstate = Unvisited | Processed | Visited (** Getting loop branches with a DFS visit : @@ -253,30 +207,31 @@ let get_directions code entrypoint = and is_loop_header = get_loop_headers code entrypoint and directions = ref (PTree.map (fun n i -> false) code) (* false <=> fallthru *) in begin - Printf.printf "Loop headers: "; - ptree_printbool is_loop_header; - Printf.printf "\n"; + (* Printf.printf "Loop headers: "; *) + (* ptree_printbool is_loop_header; *) + (* Printf.printf "\n"; *) List.iter (fun n -> match (get_some @@ PTree.get n code) with | Icond (cond, lr, ifso, ifnot) -> - Printf.printf "Analyzing %d.." (P.to_int n); + (* Printf.printf "Analyzing %d.." (P.to_int n); *) let preferred = ref false in (try - Printf.printf " call.."; + (* Printf.printf " call.."; *) do_call_heuristic code ifso ifnot is_loop_header preferred; - Printf.printf " opcode.."; + (* Printf.printf " opcode.."; *) do_opcode_heuristic code cond ifso ifnot preferred; - Printf.printf " return.."; + (* Printf.printf " return.."; *) do_return_heuristic code ifso ifnot is_loop_header preferred; - Printf.printf " store.."; + (* Printf.printf " store.."; *) do_store_heuristic code ifso ifnot is_loop_header preferred; - Printf.printf " loop.."; + (* Printf.printf " loop.."; *) do_loop_heuristic code ifso ifnot is_loop_header preferred; - Printf.printf "Random choice for %d\n" (P.to_int n); + (* Printf.printf "Random choice for %d\n" (P.to_int n); *) preferred := Random.bool () - with HeuristicSucceeded | DuplicateOpcodeHeuristic.HeuristicSucceeded - -> Printf.printf " %s\n" (match !preferred with true -> "BRANCH" - | false -> "FALLTHROUGH") + with HeuristicSucceeded | DuplicateOpcodeHeuristic.HeuristicSucceeded -> begin + (* Printf.printf " %s\n" (match !preferred with true -> "BRANCH" | false -> "FALLTHROUGH"); *) + () + end ); directions := PTree.set n !preferred !directions | _ -> () ) bfs_order; @@ -306,9 +261,9 @@ let rec to_ttl_code_rec directions = function let to_ttl_code code entrypoint = let directions = get_directions code entrypoint in begin - Printf.printf "Ifso directions: "; + (* Printf.printf "Ifso directions: "; ptree_printbool directions; - Printf.printf "\n"; + Printf.printf "\n"; *) Random.init(0); (* using same seed to make it deterministic *) to_ttl_code_rec directions (PTree.elements code) end @@ -423,7 +378,7 @@ let select_traces code entrypoint = end end done; - Printf.printf "DFS: \t"; print_intlist order; Printf.printf "\n"; + (* Printf.printf "DFS: \t"; print_intlist order; Printf.printf "\n"; *) !traces end @@ -471,7 +426,7 @@ let rec change_pointers code n n' = function * n': the integer which should contain the duplicate of n * returns: new code, new ptree *) let duplicate code ptree parent n preds n' = - Printf.printf "Duplicating node %d into %d..\n" (P.to_int n) (P.to_int n'); + (* Printf.printf "Duplicating node %d into %d..\n" (P.to_int n) (P.to_int n'); *) match PTree.get n' code with | Some _ -> failwith "The PTree already has a node n'" | None -> @@ -548,7 +503,7 @@ let rec invert_iconds_trace code = function | Icond (c, lr, ifso, ifnot) -> assert (n' == ifso || n' == ifnot); if (n' == ifso) then ( - Printf.printf "Reversing ifso/ifnot for node %d\n" (P.to_int n); + (* Printf.printf "Reversing ifso/ifnot for node %d\n" (P.to_int n); *) PTree.set n (Icond (Op.negate_condition c, lr, ifnot, ifso)) code ) else code | _ -> code @@ -568,5 +523,5 @@ let duplicate_aux f = let traces = select_traces (to_ttl_code code entrypoint) entrypoint in let icond_code = invert_iconds code traces in let preds = get_predecessors_rtl icond_code in - let (new_code, pTreeId) = (print_traces traces; superblockify_traces icond_code preds traces) in + let (new_code, pTreeId) = ((* print_traces traces; *) superblockify_traces icond_code preds traces) in ((new_code, f.fn_entrypoint), pTreeId) -- cgit From 7c403eb8d50a0292c741b25ff967d2a170637258 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 9 Mar 2020 14:27:41 +0100 Subject: cycles for aarch64 --- test/monniaux/cycles.h | 16 ++++++++++++++++ 1 file changed, 16 insertions(+) diff --git a/test/monniaux/cycles.h b/test/monniaux/cycles.h index 4a87299b..aed9941a 100644 --- a/test/monniaux/cycles.h +++ b/test/monniaux/cycles.h @@ -1,5 +1,6 @@ #include #include + typedef unsigned long cycle_t; #ifdef MAX_MEASURES @@ -39,12 +40,27 @@ static inline cycle_t get_cycle(void) { } #elif defined (__ARM_ARCH) && (__ARM_ARCH >= 6) +#if (__ARM_ARCH < 8) +/* need this kernel module +https://github.com/zertyz/MTL/tree/master/cpp/time/kernel/arm */ static inline cycle_t get_cycle(void) { cycle_t cycles; __asm__ volatile ("mrc p15, 0, %0, c9, c13, 0":"=r" (cycles)); return cycles; } +#else +/* need this kernel module: +https://github.com/jerinjacobk/armv8_pmu_cycle_counter_el0 +on 5+ kernels, remove first argument of access_ok macro */ + +static inline cycle_t get_cycle(void) +{ + uint64_t val; + __asm__ volatile("mrs %0, pmccntr_el0" : "=r"(val)); + return val; +} +#endif #else static inline cycle_t get_cycle(void) { return 0; } #endif -- cgit From 4226a49dccaafe0ecd4b591eaab932679712d58b Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Mon, 9 Mar 2020 14:44:35 +0100 Subject: Duplicate: getting rid of the annoying exception-based code --- backend/Duplicateaux.ml | 75 ++++++++++++++---------------------- mppa_k1c/DuplicateOpcodeHeuristic.ml | 9 +---- 2 files changed, 31 insertions(+), 53 deletions(-) diff --git a/backend/Duplicateaux.ml b/backend/Duplicateaux.ml index e2b5647b..05b8ddb8 100644 --- a/backend/Duplicateaux.ml +++ b/backend/Duplicateaux.ml @@ -160,47 +160,37 @@ let rec look_ahead code node is_loop_header predicate = else look_ahead code n is_loop_header predicate ) -exception HeuristicSucceeded - -let do_call_heuristic code ifso ifnot is_loop_header preferred = +let do_call_heuristic code cond ifso ifnot is_loop_header = let predicate n = (function | Icall _ -> true | _ -> false) @@ get_some @@ PTree.get n code - in if (look_ahead code ifso is_loop_header predicate) then - (preferred := false; raise HeuristicSucceeded) - else if (look_ahead code ifnot is_loop_header predicate) then - (preferred := true; raise HeuristicSucceeded) - else () + in if (look_ahead code ifso is_loop_header predicate) then Some false + else if (look_ahead code ifnot is_loop_header predicate) then Some true + else None -let do_opcode_heuristic code cond ifso ifnot preferred = DuplicateOpcodeHeuristic.opcode_heuristic code cond ifso ifnot preferred +let do_opcode_heuristic code cond ifso ifnot is_loop_header = DuplicateOpcodeHeuristic.opcode_heuristic code cond ifso ifnot is_loop_header -let do_return_heuristic code ifso ifnot is_loop_header preferred = +let do_return_heuristic code cond ifso ifnot is_loop_header = let predicate n = (function | Ireturn _ -> true | _ -> false) @@ get_some @@ PTree.get n code - in if (look_ahead code ifso is_loop_header predicate) then - (preferred := false; raise HeuristicSucceeded) - else if (look_ahead code ifnot is_loop_header predicate) then - (preferred := true; raise HeuristicSucceeded) - else () + in if (look_ahead code ifso is_loop_header predicate) then Some false + else if (look_ahead code ifnot is_loop_header predicate) then Some true + else None -let do_store_heuristic code ifso ifnot is_loop_header preferred = +let do_store_heuristic code cond ifso ifnot is_loop_header = let predicate n = (function | Istore _ -> true | _ -> false) @@ get_some @@ PTree.get n code - in if (look_ahead code ifso is_loop_header predicate) then - (preferred := false; raise HeuristicSucceeded) - else if (look_ahead code ifnot is_loop_header predicate) then - (preferred := true; raise HeuristicSucceeded) - else () + in if (look_ahead code ifso is_loop_header predicate) then Some false + else if (look_ahead code ifnot is_loop_header predicate) then Some true + else None -let do_loop_heuristic code ifso ifnot is_loop_header preferred = +let do_loop_heuristic code cond ifso ifnot is_loop_header = let predicate n = get_some @@ PTree.get n is_loop_header - in if (look_ahead code ifso is_loop_header predicate) then - (preferred := true; raise HeuristicSucceeded) - else if (look_ahead code ifnot is_loop_header predicate) then - (preferred := false; raise HeuristicSucceeded) - else () + in if (look_ahead code ifso is_loop_header predicate) then Some true + else if (look_ahead code ifnot is_loop_header predicate) then Some false + else None let get_directions code entrypoint = let bfs_order = bfs code entrypoint @@ -214,25 +204,18 @@ let get_directions code entrypoint = match (get_some @@ PTree.get n code) with | Icond (cond, lr, ifso, ifnot) -> (* Printf.printf "Analyzing %d.." (P.to_int n); *) - let preferred = ref false - in (try - (* Printf.printf " call.."; *) - do_call_heuristic code ifso ifnot is_loop_header preferred; - (* Printf.printf " opcode.."; *) - do_opcode_heuristic code cond ifso ifnot preferred; - (* Printf.printf " return.."; *) - do_return_heuristic code ifso ifnot is_loop_header preferred; - (* Printf.printf " store.."; *) - do_store_heuristic code ifso ifnot is_loop_header preferred; - (* Printf.printf " loop.."; *) - do_loop_heuristic code ifso ifnot is_loop_header preferred; - (* Printf.printf "Random choice for %d\n" (P.to_int n); *) - preferred := Random.bool () - with HeuristicSucceeded | DuplicateOpcodeHeuristic.HeuristicSucceeded -> begin - (* Printf.printf " %s\n" (match !preferred with true -> "BRANCH" | false -> "FALLTHROUGH"); *) - () - end - ); directions := PTree.set n !preferred !directions + let heuristics = [ do_call_heuristic; do_opcode_heuristic; + do_return_heuristic; do_store_heuristic; do_loop_heuristic ] in + let preferred = ref None in + begin + List.iter (fun do_heur -> + match !preferred with + | None -> preferred := do_heur code cond ifso ifnot is_loop_header + | Some _ -> () + ) heuristics; + match !preferred with None -> preferred := Some (Random.bool ()) | Some _ -> (); + directions := PTree.set n (get_some !preferred) !directions + end | _ -> () ) bfs_order; !directions diff --git a/mppa_k1c/DuplicateOpcodeHeuristic.ml b/mppa_k1c/DuplicateOpcodeHeuristic.ml index 690553ce..2ec314c1 100644 --- a/mppa_k1c/DuplicateOpcodeHeuristic.ml +++ b/mppa_k1c/DuplicateOpcodeHeuristic.ml @@ -2,10 +2,8 @@ open Op open Integers -exception HeuristicSucceeded - -let opcode_heuristic code cond ifso ifnot preferred = - let decision = match cond with +let opcode_heuristic code cond ifso ifnot is_loop_header = + match cond with | Ccompimm (c, n) | Ccompuimm (c, n) -> if n == Integers.Int.zero then (match c with | Clt | Cle -> Some false | Cgt | Cge -> Some true @@ -27,6 +25,3 @@ let opcode_heuristic code cond ifso ifnot preferred = | _ -> None ) | _ -> None - in match decision with - | Some b -> (preferred := b; raise HeuristicSucceeded) - | None -> () -- cgit From ec0d767ba602c35e320ee77f2ccd6f513adeb7b6 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Mon, 9 Mar 2020 14:57:14 +0100 Subject: Linearizeaux: forgotten print --- backend/Linearizeaux.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/backend/Linearizeaux.ml b/backend/Linearizeaux.ml index 7aed5936..22db25e0 100644 --- a/backend/Linearizeaux.ml +++ b/backend/Linearizeaux.ml @@ -417,7 +417,7 @@ let order_sequences code entry fs = end in begin (* Printf.printf "depmap: "; print_depmap depmap; *) - print_ssequence fs; + (* print_ssequence fs; *) while List.length !ordered_fs != List.length fs do let next_id = select_next () in evaluate next_id -- cgit From b016de5a1a8230b5a6c51d8e7cd8829d39a4c781 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Mon, 9 Mar 2020 15:16:08 +0100 Subject: [BROKEN] Replacing the boolean -fduplicate option by an integer To control the threshold for duplication --- backend/Duplicateaux.ml | 10 ++++++---- driver/Clflags.ml | 4 ++-- driver/Compiler.v | 10 +++++----- driver/Driver.ml | 4 ++-- extraction/extraction.v | 2 -- 5 files changed, 15 insertions(+), 15 deletions(-) diff --git a/backend/Duplicateaux.ml b/backend/Duplicateaux.ml index 05b8ddb8..636a8d8e 100644 --- a/backend/Duplicateaux.ml +++ b/backend/Duplicateaux.ml @@ -467,7 +467,7 @@ let tail_duplicate code preds ptree trace = in (new_code, new_ptree, !nb_duplicated) let superblockify_traces code preds traces = - let max_nb_duplicated = 1 (* FIXME - should be architecture dependent *) + let max_nb_duplicated = !Clflags.option_fduplicate (* FIXME - should be architecture dependent *) in let ptree = make_identity_ptree code in let rec f code ptree = function | [] -> (code, ptree, 0) @@ -499,12 +499,14 @@ let rec invert_iconds code = function else code in invert_iconds code' ts -(* For now, identity function *) let duplicate_aux f = let entrypoint = f.fn_entrypoint in let code = f.fn_code in let traces = select_traces (to_ttl_code code entrypoint) entrypoint in let icond_code = invert_iconds code traces in let preds = get_predecessors_rtl icond_code in - let (new_code, pTreeId) = ((* print_traces traces; *) superblockify_traces icond_code preds traces) in - ((new_code, f.fn_entrypoint), pTreeId) + if !Clflags.option_fduplicate >= 1 then + let (new_code, pTreeId) = ((* print_traces traces; *) superblockify_traces icond_code preds traces) in + ((new_code, f.fn_entrypoint), pTreeId) + else + ((icond_code, entrypoint), make_identity_ptree code) diff --git a/driver/Clflags.ml b/driver/Clflags.ml index 6d6f1df4..79c0bce0 100644 --- a/driver/Clflags.ml +++ b/driver/Clflags.ml @@ -28,8 +28,8 @@ let option_fconstprop = ref true let option_fcse = ref true let option_fcse2 = ref true let option_fredundancy = ref true -let option_fduplicate = ref false -let option_finvertcond = ref true (* only active if option_fduplicate is also true *) +let option_fduplicate = ref 0 +let option_finvertcond = ref true let option_ftracelinearize = ref false let option_fpostpass = ref true let option_fpostpass_sched = ref "list" diff --git a/driver/Compiler.v b/driver/Compiler.v index 499feff2..da19a0b9 100644 --- a/driver/Compiler.v +++ b/driver/Compiler.v @@ -134,7 +134,7 @@ Definition transf_rtl_program (f: RTL.program) : res Asm.program := @@ print (print_RTL 2) @@ time "Renumbering" Renumber.transf_program @@ print (print_RTL 3) - @@@ partial_if Compopts.optim_duplicate (time "Tail-duplicating" Duplicate.transf_program) + @@@ time "Tail-duplicating" Duplicate.transf_program @@ print (print_RTL 4) @@ total_if Compopts.optim_constprop (time "Constant propagation" Constprop.transf_program) @@ print (print_RTL 5) @@ -254,7 +254,7 @@ Definition CompCert's_passes := ::: mkpass (match_if Compopts.optim_tailcalls Tailcallproof.match_prog) ::: mkpass Inliningproof.match_prog ::: mkpass Renumberproof.match_prog - ::: mkpass (match_if Compopts.optim_duplicate Duplicateproof.match_prog) + ::: mkpass Duplicateproof.match_prog ::: mkpass (match_if Compopts.optim_constprop Constpropproof.match_prog) ::: mkpass (match_if Compopts.optim_constprop Renumberproof.match_prog) ::: mkpass (match_if Compopts.optim_CSE CSEproof.match_prog) @@ -301,7 +301,7 @@ Proof. set (p7 := total_if optim_tailcalls Tailcall.transf_program p6) in *. destruct (Inlining.transf_program p7) as [p8|e] eqn:P8; simpl in T; try discriminate. set (p9 := Renumber.transf_program p8) in *. - destruct (partial_if optim_duplicate Duplicate.transf_program p9) as [p10|e] eqn:P10; simpl in T; try discriminate. + destruct (Duplicate.transf_program p9) as [p10|e] eqn:P10; simpl in T; try discriminate. set (p11 := total_if optim_constprop Constprop.transf_program p10) in *. set (p12 := total_if optim_constprop Renumber.transf_program p11) in *. destruct (partial_if optim_CSE CSE.transf_program p12) as [p13|e] eqn:P13; simpl in T; try discriminate. @@ -326,7 +326,7 @@ Proof. exists p7; split. apply total_if_match. apply Tailcallproof.transf_program_match. exists p8; split. apply Inliningproof.transf_program_match; auto. exists p9; split. apply Renumberproof.transf_program_match; auto. - exists p10; split. eapply partial_if_match; eauto. apply Duplicateproof.transf_program_match; auto. + exists p10; split. apply Duplicateproof.transf_program_match; auto. exists p11; split. apply total_if_match. apply Constpropproof.transf_program_match. exists p12; split. apply total_if_match. apply Renumberproof.transf_program_match. exists p13; split. eapply partial_if_match; eauto. apply CSEproof.transf_program_match. @@ -412,7 +412,7 @@ Ltac DestructM := eapply Inliningproof.transf_program_correct; eassumption. eapply compose_forward_simulations. eapply Renumberproof.transf_program_correct; eassumption. eapply compose_forward_simulations. - eapply match_if_simulation. eassumption. exact Duplicateproof.transf_program_correct. + eapply Duplicateproof.transf_program_correct; eassumption. eapply compose_forward_simulations. eapply match_if_simulation. eassumption. exact Constpropproof.transf_program_correct. eapply compose_forward_simulations. diff --git a/driver/Driver.ml b/driver/Driver.ml index db71aef9..dd357423 100644 --- a/driver/Driver.ml +++ b/driver/Driver.ml @@ -204,7 +204,7 @@ Processing options: -finvertcond Invert conditions based on predicted paths (to prefer fallthrough). Requires -fduplicate to be also activated [on] -ftracelinearize Linearizes based on the traces identified by duplicate phase - It is recommended to also activate -fduplicate with this pass [off] + It is heavily recommended to activate -finvertcond with this pass [off] -fforward-moves Forward moves after CSE -finline Perform inlining of functions [on] -finline-functions-called-once Integrate functions only required by their @@ -393,7 +393,7 @@ let cmdline_actions = @ f_opt "cse2" option_fcse2 @ f_opt "redundancy" option_fredundancy @ f_opt "postpass" option_fpostpass - @ f_opt "duplicate" option_fduplicate + @ [ Exact "-fduplicate", Integer (fun n -> option_fduplicate := n) ] @ f_opt "invertcond" option_finvertcond @ f_opt "tracelinearize" option_ftracelinearize @ f_opt_str "postpass" option_fpostpass option_fpostpass_sched diff --git a/extraction/extraction.v b/extraction/extraction.v index 929c21e0..ba6b080b 100644 --- a/extraction/extraction.v +++ b/extraction/extraction.v @@ -105,8 +105,6 @@ Extract Constant Compopts.generate_float_constants => "fun _ -> !Clflags.option_ffloatconstprop >= 2". Extract Constant Compopts.optim_tailcalls => "fun _ -> !Clflags.option_ftailcalls". -Extract Constant Compopts.optim_duplicate => - "fun _ -> !Clflags.option_fduplicate". Extract Constant Compopts.optim_constprop => "fun _ -> !Clflags.option_fconstprop". Extract Constant Compopts.optim_CSE => -- cgit From deaf767b7f60cc3b3a0d3314e763a682571a00fa Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Mon, 9 Mar 2020 15:16:44 +0100 Subject: more portable cycles.h --- test/monniaux/clock.c | 4 ++-- test/monniaux/cycles.h | 36 +++++++++++++++++++++++++-------- test/monniaux/quicksort/quicksort_run.c | 2 +- 3 files changed, 31 insertions(+), 11 deletions(-) diff --git a/test/monniaux/clock.c b/test/monniaux/clock.c index fb636667..4ec679f6 100644 --- a/test/monniaux/clock.c +++ b/test/monniaux/clock.c @@ -24,9 +24,9 @@ cycle_t get_current_cycle(void) { } void print_total_clock(void) { - printf("time cycles: %lu\n", total_clock); + printf("time cycles: %" PRcycle "\n", total_clock); } void printerr_total_clock(void) { - fprintf(stderr, "time cycles: %lu\n", total_clock); + fprintf(stderr, "time cycles: %" PRcycle "\n", total_clock); } diff --git a/test/monniaux/cycles.h b/test/monniaux/cycles.h index aed9941a..c7dc582b 100644 --- a/test/monniaux/cycles.h +++ b/test/monniaux/cycles.h @@ -1,14 +1,11 @@ +#include #include #include -typedef unsigned long cycle_t; - -#ifdef MAX_MEASURES - static cycle_t _last_stop[MAX_MEASURES] = {0}; - static cycle_t _total_cycles[MAX_MEASURES] = {0}; -#endif - #ifdef __K1C__ +typedef uint64_t cycle_t; +#define PRcycle PRId64 + #include <../../k1-cos/include/hal/cos_registers.h> static inline void cycle_count_config(void) @@ -28,11 +25,20 @@ static inline cycle_t get_cycle(void) #else // not K1C static inline void cycle_count_config(void) { } -#ifdef __x86_64__ +#if defined(__i386__) || defined( __x86_64__) +#define PRcycle PRId64 +typedef uint64_t cycle_t; #include static inline cycle_t get_cycle(void) { return __rdtsc(); } #elif __riscv +#ifdef __riscv32 +#define PRcycle PRId32 +typedef uint32_t cycle_t; +#else +#define PRcycle PRId64 +typedef uint64_t cycle_t; +#endif static inline cycle_t get_cycle(void) { cycle_t cycles; asm volatile ("rdcycle %0" : "=r" (cycles)); @@ -41,6 +47,9 @@ static inline cycle_t get_cycle(void) { #elif defined (__ARM_ARCH) && (__ARM_ARCH >= 6) #if (__ARM_ARCH < 8) +typedef uint32_t cycle_t; +#define PRcycle PRId32 + /* need this kernel module https://github.com/zertyz/MTL/tree/master/cpp/time/kernel/arm */ static inline cycle_t get_cycle(void) { @@ -49,6 +58,8 @@ static inline cycle_t get_cycle(void) { return cycles; } #else +#define PRcycle PRId64 +typedef uint64_t cycle_t; /* need this kernel module: https://github.com/jerinjacobk/armv8_pmu_cycle_counter_el0 @@ -61,7 +72,10 @@ static inline cycle_t get_cycle(void) return val; } #endif + #else +#define PRcycle PRId32 +typedef uint32_t cycle_t; static inline cycle_t get_cycle(void) { return 0; } #endif #endif @@ -71,3 +85,9 @@ static inline cycle_t get_cycle(void) { return 0; } #define TIMESTOP(i) {cycle_t cur = get_cycle(); _total_cycles[i] += cur - _last_stop[i]; _last_stop[i] = cur;} #define TIMEPRINT(n) { for (int i = 0; i <= n; i++) printf("%d cycles: %" PRIu64 "\n", i, _total_cycles[i]); } #endif + + +#ifdef MAX_MEASURES + static cycle_t _last_stop[MAX_MEASURES] = {0}; + static cycle_t _total_cycles[MAX_MEASURES] = {0}; +#endif diff --git a/test/monniaux/quicksort/quicksort_run.c b/test/monniaux/quicksort/quicksort_run.c index c35d0752..3c640b24 100644 --- a/test/monniaux/quicksort/quicksort_run.c +++ b/test/monniaux/quicksort/quicksort_run.c @@ -13,7 +13,7 @@ int main (void) { quicksort(vec, len); quicksort_time = get_cycle() - quicksort_time; printf("sorted=%s\n" - "time cycles:%" PRIu64 "\n", + "time cycles:%" PRcycle "\n", data_vec_is_sorted(vec, len)?"true":"false", quicksort_time); free(vec); -- cgit From 103083dfcef7a71a57fd6c05af276db1f034ac75 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Tue, 10 Mar 2020 11:18:06 +0100 Subject: Fixing build --- driver/Driver.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/driver/Driver.ml b/driver/Driver.ml index dd357423..43aedf50 100644 --- a/driver/Driver.ml +++ b/driver/Driver.ml @@ -318,7 +318,7 @@ let cmdline_actions = [ Exact "-O0", Unit (unset_all optimization_options); Exact "-O", Unit (set_all optimization_options); - _Regexp "-O1", Self (fun _ -> set_all optimization_options (); option_fpostpass := false; option_fduplicate := false); + _Regexp "-O1", Self (fun _ -> set_all optimization_options (); option_fpostpass := false); _Regexp "-O[123]$", Unit (set_all optimization_options); Exact "-Os", Set option_Osize; Exact "-Obranchless", Set option_Obranchless; -- cgit From 22bfb4389571e9b2779b6e5b9f48b8a0e5c35867 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Tue, 10 Mar 2020 14:16:23 +0100 Subject: Bug fix in ftracelinearize --- backend/Linearizeaux.ml | 19 +++++++++++++++---- 1 file changed, 15 insertions(+), 4 deletions(-) diff --git a/backend/Linearizeaux.ml b/backend/Linearizeaux.ml index 22db25e0..23d06075 100644 --- a/backend/Linearizeaux.ml +++ b/backend/Linearizeaux.ml @@ -140,10 +140,21 @@ let rec last_element = function | e :: [] -> e | e' :: e :: l -> last_element (e::l) +let print_plist l = + let rec f = function + | [] -> () + | n :: l -> Printf.printf "%d, " (P.to_int n); f l + in begin + Printf.printf "["; + f l; + Printf.printf "]" + end + let forward_sequences code entry = let visited = ref (PTree.map (fun n i -> false) code) in (* returns the list of traversed nodes, and a list of nodes to start traversing next *) let rec traverse_fallthrough code node = + (* Printf.printf "Traversing %d..\n" (P.to_int node); *) if not (get_some @@ PTree.get node !visited) then begin visited := PTree.set node true !visited; match PTree.get node code with @@ -164,8 +175,8 @@ let forward_sequences code entry = in let rec f code = function | [] -> [] | node :: ln -> - let fs, rem = traverse_fallthrough code node - in [fs] @ (f code rem) + let fs, rem_from_node = traverse_fallthrough code node + in [fs] @ ((f code rem_from_node) @ (f code ln)) in (f code [entry]) module PInt = struct @@ -417,12 +428,12 @@ let order_sequences code entry fs = end in begin (* Printf.printf "depmap: "; print_depmap depmap; *) - (* print_ssequence fs; *) + (* Printf.printf "forward sequences identified: "; print_ssequence fs; *) while List.length !ordered_fs != List.length fs do let next_id = select_next () in evaluate next_id done; - (* print_ssequence (List.rev (!ordered_fs)); *) + (* Printf.printf "forward sequences ordered: "; print_ssequence (List.rev (!ordered_fs)); *) List.rev (!ordered_fs) end -- cgit From 273c48f412d018e5d5649db266b282fc272a0af8 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Tue, 10 Mar 2020 15:23:41 +0100 Subject: Linearize: More helpful message when tracelinearize fails --- backend/Linearizeaux.ml | 16 ++++++++++++---- 1 file changed, 12 insertions(+), 4 deletions(-) diff --git a/backend/Linearizeaux.ml b/backend/Linearizeaux.ml index 23d06075..bd8f747e 100644 --- a/backend/Linearizeaux.ml +++ b/backend/Linearizeaux.ml @@ -415,16 +415,24 @@ let order_sequences code entry fs = ) depmap end in let select_next () = - let selected_id = ref (-1) in + let selected_id = ref None in begin Array.iteri (fun i deps -> begin (* Printf.printf "Deps: "; print_iset deps; Printf.printf "\n"; *) - if !selected_id == -1 && deps == ISet.empty && not fs_evaluated.(i) - then selected_id := i + match !selected_id with + | None -> if (deps == ISet.empty && not fs_evaluated.(i)) then selected_id := Some i + | Some id -> () end ) depmap; - !selected_id + match !selected_id with + | Some id -> id + | None -> begin + Printf.printf "original fs: "; print_ssequence fs; + Printf.printf "depmap: "; print_depmap depmap; + Printf.printf "current ordered fs: "; print_ssequence @@ List.rev (!ordered_fs); + failwith "Could not find a next schedulable trace. Are the dependencies alright?" + end end in begin (* Printf.printf "depmap: "; print_depmap depmap; *) -- cgit From 9a86085b2226905e8cf14b600f2069202c5d12bd Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Tue, 10 Mar 2020 15:24:13 +0100 Subject: Some dependencies were not taken into account in tracelinearize --- backend/Linearizeaux.ml | 27 ++++++++++++--------------- 1 file changed, 12 insertions(+), 15 deletions(-) diff --git a/backend/Linearizeaux.ml b/backend/Linearizeaux.ml index bd8f747e..3602cb91 100644 --- a/backend/Linearizeaux.ml +++ b/backend/Linearizeaux.ml @@ -349,6 +349,14 @@ let construct_depmap code entry fs = ) fs; !index end + in let check_and_update_depmap from target = + if not (ppmap_is_true (from, target) is_loop_edge) then + let in_index_fs = find_index_of_node from in + let out_index_fs = find_index_of_node target in + if out_index_fs != in_index_fs then + depmap.(out_index_fs) <- ISet.add in_index_fs depmap.(out_index_fs) + else () + else () in let rec dfs_visit code = function | [] -> () | node :: ln -> @@ -360,25 +368,14 @@ let construct_depmap code entry fs = let next_visits = match (last_element bb) with | Ltailcall _ | Lreturn -> [] - | Lbranch n -> [n] + | Lbranch n -> (check_and_update_depmap node n; [n]) | Lcond (_, _, ifso, ifnot) -> begin - (if not (ppmap_is_true (node, ifso) is_loop_edge) then - let in_index_fs = find_index_of_node node in - let out_index_fs = find_index_of_node ifso in - depmap.(out_index_fs) <- ISet.add in_index_fs depmap.(out_index_fs) - else - ()); + check_and_update_depmap node ifso; + check_and_update_depmap node ifnot; [ifso; ifnot] end | Ljumptable(_, ln) -> begin - let in_index_fs = find_index_of_node node in - List.iter (fun n -> - if not (ppmap_is_true (node, n) is_loop_edge) then - let out_index_fs = find_index_of_node n in - depmap.(out_index_fs) <- ISet.add in_index_fs depmap.(out_index_fs) - else - () - ) ln; + List.iter (fun n -> check_and_update_depmap node n) ln; ln end (* end of bblocks should not be another value than one of the above *) -- cgit From a7098538edfda9fdbd95bc7c6ba6e380811230fa Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Tue, 10 Mar 2020 15:41:05 +0100 Subject: Linearizeaux, forgot to visit the rest of the nodes in dfs_visit --- backend/Linearizeaux.ml | 48 ++++++++++++++++++++++++++---------------------- 1 file changed, 26 insertions(+), 22 deletions(-) diff --git a/backend/Linearizeaux.ml b/backend/Linearizeaux.ml index 3602cb91..5bdeeb8f 100644 --- a/backend/Linearizeaux.ml +++ b/backend/Linearizeaux.ml @@ -350,6 +350,7 @@ let construct_depmap code entry fs = !index end in let check_and_update_depmap from target = + (* Printf.printf "From %d to %d\n" (P.to_int from) (P.to_int target); *) if not (ppmap_is_true (from, target) is_loop_edge) then let in_index_fs = find_index_of_node from in let out_index_fs = find_index_of_node target in @@ -360,28 +361,31 @@ let construct_depmap code entry fs = in let rec dfs_visit code = function | [] -> () | node :: ln -> - match (get_some @@ PTree.get node !visited) with - | true -> () - | false -> begin - visited := PTree.set node true !visited; - let bb = get_some @@ PTree.get node code in - let next_visits = - match (last_element bb) with - | Ltailcall _ | Lreturn -> [] - | Lbranch n -> (check_and_update_depmap node n; [n]) - | Lcond (_, _, ifso, ifnot) -> begin - check_and_update_depmap node ifso; - check_and_update_depmap node ifnot; - [ifso; ifnot] - end - | Ljumptable(_, ln) -> begin - List.iter (fun n -> check_and_update_depmap node n) ln; - ln - end - (* end of bblocks should not be another value than one of the above *) - | _ -> failwith "last_element gave an invalid output" - in dfs_visit code next_visits - end + begin + match (get_some @@ PTree.get node !visited) with + | true -> () + | false -> begin + visited := PTree.set node true !visited; + let bb = get_some @@ PTree.get node code in + let next_visits = + match (last_element bb) with + | Ltailcall _ | Lreturn -> [] + | Lbranch n -> (check_and_update_depmap node n; [n]) + | Lcond (_, _, ifso, ifnot) -> begin + check_and_update_depmap node ifso; + check_and_update_depmap node ifnot; + [ifso; ifnot] + end + | Ljumptable(_, ln) -> begin + List.iter (fun n -> check_and_update_depmap node n) ln; + ln + end + (* end of bblocks should not be another value than one of the above *) + | _ -> failwith "last_element gave an invalid output" + in dfs_visit code next_visits + end; + dfs_visit code ln + end in begin dfs_visit code [entry]; depmap -- cgit From 530d30cf71661419f54e175dd6bdb7d3f68f7f5c Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Tue, 10 Mar 2020 15:48:47 +0100 Subject: Linearizeaux: dumb selector when cycling dependencies are found --- backend/Linearizeaux.ml | 10 ++++++---- 1 file changed, 6 insertions(+), 4 deletions(-) diff --git a/backend/Linearizeaux.ml b/backend/Linearizeaux.ml index 5bdeeb8f..a813ac96 100644 --- a/backend/Linearizeaux.ml +++ b/backend/Linearizeaux.ml @@ -429,10 +429,12 @@ let order_sequences code entry fs = match !selected_id with | Some id -> id | None -> begin - Printf.printf "original fs: "; print_ssequence fs; - Printf.printf "depmap: "; print_depmap depmap; - Printf.printf "current ordered fs: "; print_ssequence @@ List.rev (!ordered_fs); - failwith "Could not find a next schedulable trace. Are the dependencies alright?" + Array.iteri (fun i deps -> + match !selected_id with + | None -> if not fs_evaluated.(i) then selected_id := Some i + | Some id -> () + ) depmap; + get_some !selected_id end end in begin -- cgit From e63e318c720c678d44cbb27d940ebfa076a7f8b4 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 11 Mar 2020 13:34:35 +0100 Subject: remet is_trivial_op dans CSE2 --- backend/CSE2.v | 2 +- backend/CSE2proof.v | 5 ++++- 2 files changed, 5 insertions(+), 2 deletions(-) diff --git a/backend/CSE2.v b/backend/CSE2.v index be72405b..d5444e3b 100644 --- a/backend/CSE2.v +++ b/backend/CSE2.v @@ -469,7 +469,7 @@ Definition transf_instr (fmap : option (PMap.t RB.t)) match instr with | Iop op args dst s => let args' := subst_args fmap pc args in - match find_op_in_fmap fmap pc op args' with + match (if is_trivial_op op then None else find_op_in_fmap fmap pc op args') with | None => Iop op args' dst s | Some src => Iop Omove (src::nil) dst s end diff --git a/backend/CSE2proof.v b/backend/CSE2proof.v index 7e1dd430..6368e585 100644 --- a/backend/CSE2proof.v +++ b/backend/CSE2proof.v @@ -1200,8 +1200,11 @@ Proof. reflexivity. - (* op *) unfold transf_instr in *. - destruct find_op_in_fmap eqn:FIND_OP. + destruct (if is_trivial_op op then None else find_op_in_fmap (forward_map f) pc op + (subst_args (forward_map f) pc args)) eqn:FIND_OP. { + destruct (is_trivial_op op). + discriminate. unfold find_op_in_fmap, fmap_sem', fmap_sem in *. destruct (forward_map f) as [map |] eqn:MAP. 2: discriminate. -- cgit From 9e706de1eb25d6d6dbeb1eb2ced71e48523a499f Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 11 Mar 2020 15:43:26 +0100 Subject: Fixed stupid typo bug preventing the prediction update for the RANDOM predictor --- backend/Duplicateaux.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/backend/Duplicateaux.ml b/backend/Duplicateaux.ml index 636a8d8e..209527b9 100644 --- a/backend/Duplicateaux.ml +++ b/backend/Duplicateaux.ml @@ -213,7 +213,7 @@ let get_directions code entrypoint = | None -> preferred := do_heur code cond ifso ifnot is_loop_header | Some _ -> () ) heuristics; - match !preferred with None -> preferred := Some (Random.bool ()) | Some _ -> (); + (match !preferred with None -> preferred := Some (Random.bool ()) | Some _ -> ()); directions := PTree.set n (get_some !preferred) !directions end | _ -> () -- cgit From 1b00a75796a8ace42cc480efadaad948407f5a31 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 11 Mar 2020 16:16:18 +0100 Subject: More debug info on Linearize and Duplicate --- backend/Duplicateaux.ml | 71 ++++++++++++++++++++++++++++++++----------------- backend/Linearizeaux.ml | 2 +- 2 files changed, 47 insertions(+), 26 deletions(-) diff --git a/backend/Duplicateaux.ml b/backend/Duplicateaux.ml index 209527b9..d3036b9a 100644 --- a/backend/Duplicateaux.ml +++ b/backend/Duplicateaux.ml @@ -161,36 +161,52 @@ let rec look_ahead code node is_loop_header predicate = ) let do_call_heuristic code cond ifso ifnot is_loop_header = - let predicate n = (function - | Icall _ -> true - | _ -> false) @@ get_some @@ PTree.get n code - in if (look_ahead code ifso is_loop_header predicate) then Some false - else if (look_ahead code ifnot is_loop_header predicate) then Some true - else None + begin + Printf.printf "\tCall heuristic..\n"; + let predicate n = (function + | Icall _ -> true + | _ -> false) @@ get_some @@ PTree.get n code + in if (look_ahead code ifso is_loop_header predicate) then Some false + else if (look_ahead code ifnot is_loop_header predicate) then Some true + else None + end -let do_opcode_heuristic code cond ifso ifnot is_loop_header = DuplicateOpcodeHeuristic.opcode_heuristic code cond ifso ifnot is_loop_header +let do_opcode_heuristic code cond ifso ifnot is_loop_header = + begin + Printf.printf "\tOpcode heuristic..\n"; + DuplicateOpcodeHeuristic.opcode_heuristic code cond ifso ifnot is_loop_header + end let do_return_heuristic code cond ifso ifnot is_loop_header = - let predicate n = (function - | Ireturn _ -> true - | _ -> false) @@ get_some @@ PTree.get n code - in if (look_ahead code ifso is_loop_header predicate) then Some false - else if (look_ahead code ifnot is_loop_header predicate) then Some true - else None + begin + Printf.printf "\tReturn heuristic..\n"; + let predicate n = (function + | Ireturn _ -> true + | _ -> false) @@ get_some @@ PTree.get n code + in if (look_ahead code ifso is_loop_header predicate) then Some false + else if (look_ahead code ifnot is_loop_header predicate) then Some true + else None + end let do_store_heuristic code cond ifso ifnot is_loop_header = - let predicate n = (function - | Istore _ -> true - | _ -> false) @@ get_some @@ PTree.get n code - in if (look_ahead code ifso is_loop_header predicate) then Some false - else if (look_ahead code ifnot is_loop_header predicate) then Some true - else None + begin + Printf.printf "\tStore heuristic..\n"; + let predicate n = (function + | Istore _ -> true + | _ -> false) @@ get_some @@ PTree.get n code + in if (look_ahead code ifso is_loop_header predicate) then Some false + else if (look_ahead code ifnot is_loop_header predicate) then Some true + else None + end let do_loop_heuristic code cond ifso ifnot is_loop_header = - let predicate n = get_some @@ PTree.get n is_loop_header - in if (look_ahead code ifso is_loop_header predicate) then Some true - else if (look_ahead code ifnot is_loop_header predicate) then Some false - else None + begin + Printf.printf "\tLoop heuristic..\n"; + let predicate n = get_some @@ PTree.get n is_loop_header + in if (look_ahead code ifso is_loop_header predicate) then Some true + else if (look_ahead code ifnot is_loop_header predicate) then Some false + else None + end let get_directions code entrypoint = let bfs_order = bfs code entrypoint @@ -208,13 +224,18 @@ let get_directions code entrypoint = do_return_heuristic; do_store_heuristic; do_loop_heuristic ] in let preferred = ref None in begin + Printf.printf "Deciding condition for RTL node %d\n" (P.to_int n); List.iter (fun do_heur -> match !preferred with | None -> preferred := do_heur code cond ifso ifnot is_loop_header | Some _ -> () ) heuristics; - (match !preferred with None -> preferred := Some (Random.bool ()) | Some _ -> ()); - directions := PTree.set n (get_some !preferred) !directions + (match !preferred with None -> (Printf.printf "\tRANDOM\n"; preferred := Some (Random.bool ())) | Some _ -> ()); + directions := PTree.set n (get_some !preferred) !directions; + (match !preferred with | Some false -> Printf.printf "\tFALLTHROUGH\n" + | Some true -> Printf.printf "\tBRANCH\n" + | None -> ()); + Printf.printf "---------------------------------------\n" end | _ -> () ) bfs_order; diff --git a/backend/Linearizeaux.ml b/backend/Linearizeaux.ml index a813ac96..eed58f8d 100644 --- a/backend/Linearizeaux.ml +++ b/backend/Linearizeaux.ml @@ -444,7 +444,7 @@ let order_sequences code entry fs = let next_id = select_next () in evaluate next_id done; - (* Printf.printf "forward sequences ordered: "; print_ssequence (List.rev (!ordered_fs)); *) + Printf.printf "forward sequences ordered: "; print_ssequence (List.rev (!ordered_fs)); List.rev (!ordered_fs) end -- cgit From 3fef5e1d45820a775a7c941851af6f0bf3f1537d Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 11 Mar 2020 17:00:48 +0100 Subject: Adding info field for branching in RTL, LTL, XTL and all associated passes --- backend/Allocation.v | 4 ++-- backend/Allocproof.v | 6 +++--- backend/CSE.v | 6 +++--- backend/CSE2.v | 6 +++--- backend/Constprop.v | 6 +++--- backend/Constpropproof.v | 4 ++-- backend/Deadcode.v | 4 ++-- backend/Duplicate.v | 4 ++-- backend/Duplicateaux.ml | 27 ++++++++++++++------------- backend/Duplicateproof.v | 8 ++++---- backend/ForwardMoves.v | 6 +++--- backend/Inlining.v | 4 ++-- backend/Inliningspec.v | 6 +++--- backend/LTL.v | 8 ++++---- backend/Linearize.v | 2 +- backend/Linearizeaux.ml | 10 +++++----- backend/Liveness.v | 2 +- backend/PrintLTL.ml | 2 +- backend/PrintRTL.ml | 2 +- backend/PrintXTL.ml | 2 +- backend/RTL.v | 19 ++++++++++--------- backend/RTLgen.v | 2 +- backend/RTLgenspec.v | 4 ++-- backend/RTLtyping.v | 6 +++--- backend/Regalloc.ml | 20 ++++++++++---------- backend/Renumber.v | 2 +- backend/Splitting.ml | 4 ++-- backend/Tunneling.v | 4 ++-- backend/Unusedglob.v | 2 +- backend/ValueAnalysis.v | 2 +- backend/XTL.ml | 6 +++--- backend/XTL.mli | 2 +- 32 files changed, 97 insertions(+), 95 deletions(-) diff --git a/backend/Allocation.v b/backend/Allocation.v index d18b07a9..2323c050 100644 --- a/backend/Allocation.v +++ b/backend/Allocation.v @@ -314,10 +314,10 @@ Definition pair_instr_block Some(BSbuiltin ef args res mv1 args' res' mv2 s) | _ => None end - | Icond cond args s1 s2 => + | Icond cond args s1 s2 i => let (mv1, b1) := extract_moves nil b in match b1 with - | Lcond cond' args' s1' s2' :: b2 => + | Lcond cond' args' s1' s2' i' :: b2 => assertion (eq_condition cond cond'); assertion (peq s1 s1'); assertion (peq s2 s2'); diff --git a/backend/Allocproof.v b/backend/Allocproof.v index b6880860..3c7df58a 100644 --- a/backend/Allocproof.v +++ b/backend/Allocproof.v @@ -169,11 +169,11 @@ Inductive expand_block_shape: block_shape -> RTL.instruction -> LTL.bblock -> Pr (Ibuiltin ef args res s) (expand_moves mv1 (Lbuiltin ef args' res' :: expand_moves mv2 (Lbranch s :: k))) - | ebs_cond: forall cond args mv args' s1 s2 k, + | ebs_cond: forall cond args mv args' s1 s2 k i i', wf_moves mv -> expand_block_shape (BScond cond args mv args' s1 s2) - (Icond cond args s1 s2) - (expand_moves mv (Lcond cond args' s1 s2 :: k)) + (Icond cond args s1 s2 i) + (expand_moves mv (Lcond cond args' s1 s2 i' :: k)) | ebs_jumptable: forall arg mv arg' tbl k, wf_moves mv -> expand_block_shape (BSjumptable arg mv arg' tbl) diff --git a/backend/CSE.v b/backend/CSE.v index 2827161d..1936d4e4 100644 --- a/backend/CSE.v +++ b/backend/CSE.v @@ -496,7 +496,7 @@ Definition transfer (f: function) (approx: PMap.t VA.t) (pc: node) (before: numb | EF_vload _ | EF_annot _ _ _ | EF_annot_val _ _ _ | EF_debug _ _ _ => set_res_unknown before res end - | Icond cond args ifso ifnot => + | Icond cond args ifso ifnot _ => before | Ijumptable arg tbl => before @@ -549,10 +549,10 @@ Definition transf_instr (n: numbering) (instr: instruction) := let (n1, vl) := valnum_regs n args in let (addr', args') := reduce _ combine_addr n1 addr args vl in Istore chunk addr' args' src s - | Icond cond args s1 s2 => + | Icond cond args s1 s2 i => let (n1, vl) := valnum_regs n args in let (cond', args') := reduce _ combine_cond n1 cond args vl in - Icond cond' args' s1 s2 + Icond cond' args' s1 s2 i | _ => instr end. diff --git a/backend/CSE2.v b/backend/CSE2.v index d5444e3b..900a7517 100644 --- a/backend/CSE2.v +++ b/backend/CSE2.v @@ -405,7 +405,7 @@ Qed. Definition apply_instr instr (rel : RELATION.t) : RB.t := match instr with | Inop _ - | Icond _ _ _ _ + | Icond _ _ _ _ _ | Ijumptable _ _ => Some rel | Istore chunk addr args _ _ => Some (kill_store chunk addr args rel) | Iop op args dst _ => Some (gen_oper op dst args rel) @@ -485,8 +485,8 @@ Definition transf_instr (fmap : option (PMap.t RB.t)) Icall sig ros (subst_args fmap pc args) dst s | Itailcall sig ros args => Itailcall sig ros (subst_args fmap pc args) - | Icond cond args s1 s2 => - Icond cond (subst_args fmap pc args) s1 s2 + | Icond cond args s1 s2 i => + Icond cond (subst_args fmap pc args) s1 s2 i | Ijumptable arg tbl => Ijumptable (subst_arg fmap pc arg) tbl | Ireturn (Some arg) => diff --git a/backend/Constprop.v b/backend/Constprop.v index eda41b39..0be9438c 100644 --- a/backend/Constprop.v +++ b/backend/Constprop.v @@ -69,7 +69,7 @@ Fixpoint successor_rec (n: nat) (f: function) (ae: AE.t) (pc: node) : node := match f.(fn_code)!pc with | Some (Inop s) => successor_rec n' f ae s - | Some (Icond cond args s1 s2) => + | Some (Icond cond args s1 s2 _) => match resolve_branch (eval_static_condition cond (aregs ae args)) with | Some b => successor_rec n' f ae (if b then s1 else s2) | None => pc @@ -217,14 +217,14 @@ Definition transf_instr (f: function) (an: PMap.t VA.t) (rm: romem) end | _, _ => dfl end - | Icond cond args s1 s2 => + | Icond cond args s1 s2 i => let aargs := aregs ae args in match resolve_branch (eval_static_condition cond aargs) with | Some b => if b then Inop s1 else Inop s2 | None => let (cond', args') := cond_strength_reduction cond args aargs in - Icond cond' args' s1 s2 + Icond cond' args' s1 s2 i end | Ijumptable arg tbl => match areg ae arg with diff --git a/backend/Constpropproof.v b/backend/Constpropproof.v index 63cfee24..60663503 100644 --- a/backend/Constpropproof.v +++ b/backend/Constpropproof.v @@ -142,8 +142,8 @@ Inductive match_pc (f: function) (rs: regset) (m: mem): nat -> node -> node -> P f.(fn_code)!pc = Some (Inop s) -> match_pc f rs m n s pcx -> match_pc f rs m (S n) pc pcx - | match_pc_cond: forall n pc cond args s1 s2 pcx, - f.(fn_code)!pc = Some (Icond cond args s1 s2) -> + | match_pc_cond: forall n pc cond args s1 s2 pcx i, + f.(fn_code)!pc = Some (Icond cond args s1 s2 i) -> (forall b, eval_condition cond rs##args m = Some b -> match_pc f rs m n (if b then s1 else s2) pcx) -> diff --git a/backend/Deadcode.v b/backend/Deadcode.v index 1f208a91..3412a6fa 100644 --- a/backend/Deadcode.v +++ b/backend/Deadcode.v @@ -142,7 +142,7 @@ Definition transfer (f: function) (approx: PMap.t VA.t) nmem_dead_stack f.(fn_stacksize)) | Some(Ibuiltin ef args res s) => transfer_builtin approx!!pc ef args res ne nm - | Some(Icond cond args s1 s2) => + | Some(Icond cond args s1 s2 _) => if peq s1 s2 then after else (add_needs args (needs_of_condition cond) ne, nm) | Some(Ijumptable arg tbl) => @@ -192,7 +192,7 @@ Definition transf_instr (approx: PMap.t VA.t) (an: PMap.t NA.t) if nmem_contains (snd an!!pc) (aaddr_arg approx!!pc dst) sz then instr else Inop s - | Icond cond args s1 s2 => + | Icond cond args s1 s2 _ => if peq s1 s2 then Inop s1 else instr | _ => instr diff --git a/backend/Duplicate.v b/backend/Duplicate.v index 82c17367..af85efe4 100644 --- a/backend/Duplicate.v +++ b/backend/Duplicate.v @@ -134,8 +134,8 @@ Definition verify_match_inst dupmap inst tinst := else Error (msg "Different ef in Ibuiltin") | _ => Error (msg "verify_match_inst Ibuiltin") end - | Icond cond lr n1 n2 => match tinst with - | Icond cond' lr' n1' n2' => + | Icond cond lr n1 n2 i => match tinst with + | Icond cond' lr' n1' n2' i' => if (list_eq_dec Pos.eq_dec lr lr') then if (eq_condition cond cond') then do u1 <- verify_is_copy dupmap n1 n1'; diff --git a/backend/Duplicateaux.ml b/backend/Duplicateaux.ml index d3036b9a..86bc06c9 100644 --- a/backend/Duplicateaux.ml +++ b/backend/Duplicateaux.ml @@ -39,7 +39,7 @@ let bfs code entrypoint = | Ibuiltin(_, _, _, n) -> Queue.add n to_visit | Ijumptable(_, ln) -> List.iter (fun n -> Queue.add n to_visit) ln | Itailcall _ | Ireturn _ -> () - | Icond (_, _, n1, n2) -> Queue.add n1 to_visit; Queue.add n2 to_visit + | Icond (_, _, n1, n2, _) -> Queue.add n1 to_visit; Queue.add n2 to_visit | Inop n | Iop (_, _, _, n) | Iload (_, _, _, _, _, n) | Istore (_, _, _, _, n) -> Queue.add n to_visit end done; @@ -56,7 +56,7 @@ let get_predecessors_rtl code = let succ = match i with | Inop n | Iop (_,_,_,n) | Iload (_, _,_,_,_,n) | Istore (_,_,_,_,n) | Icall (_,_,_,_,n) | Ibuiltin (_, _, _, n) -> [n] - | Icond (_,_,n1,n2) -> [n1;n2] + | Icond (_,_,n1,n2,_) -> [n1;n2] | Ijumptable (_,ln) -> ln | Itailcall _ | Ireturn _ -> [] in List.iter (fun s -> @@ -123,7 +123,7 @@ let get_loop_headers code entrypoint = | Some i -> let next_visits = (match i with | Icall (_, _, _, _, n) | Ibuiltin (_, _, _, n) | Inop n | Iop (_, _, _, n) | Iload (_, _, _, _, _, n) | Istore (_, _, _, _, n) -> [n] - | Icond (_, _, n1, n2) -> [n1; n2] + | Icond (_, _, n1, n2, _) -> [n1; n2] | Itailcall _ | Ireturn _ -> [] | Ijumptable (_, ln) -> ln ) in dfs_visit code next_visits; @@ -218,7 +218,7 @@ let get_directions code entrypoint = (* Printf.printf "\n"; *) List.iter (fun n -> match (get_some @@ PTree.get n code) with - | Icond (cond, lr, ifso, ifnot) -> + | Icond (cond, lr, ifso, ifnot, _) -> (* Printf.printf "Analyzing %d.." (P.to_int n); *) let heuristics = [ do_call_heuristic; do_opcode_heuristic; do_return_heuristic; do_store_heuristic; do_loop_heuristic ] in @@ -251,9 +251,9 @@ let to_ttl_inst direction = function | Icall (s, ri, lr, r, n) -> Tleaf (Icall(s, ri, lr, r, n)) | Itailcall (s, ri, lr) -> Tleaf (Itailcall(s, ri, lr)) | Ibuiltin (ef, lbr, br, n) -> Tleaf (Ibuiltin(ef, lbr, br, n)) -| Icond (cond, lr, n, n') -> (match direction with - | false -> Tnext (n', Icond(cond, lr, n, n')) - | true -> Tnext (n, Icond(cond, lr, n, n'))) +| Icond (cond, lr, n, n', i) -> (match direction with + | false -> Tnext (n', Icond(cond, lr, n, n', i)) + | true -> Tnext (n, Icond(cond, lr, n, n', i))) | Ijumptable (r, ln) -> Tleaf (Ijumptable(r, ln)) let rec to_ttl_code_rec directions = function @@ -299,7 +299,7 @@ let dfs code entrypoint = | Itailcall _ | Ireturn _ -> [] | _ -> failwith "Tleaf case not handled in dfs" ) | Tnext (n,i) -> (dfs_list code [n]) @ match i with - | Icond (_, _, n1, n2) -> dfs_list code [n1; n2] + | Icond (_, _, n1, n2, _) -> dfs_list code [n1; n2] | Inop _ | Iop _ | Iload _ | Istore _ -> [] | _ -> failwith "Tnext case not handled in dfs" end @@ -314,7 +314,7 @@ let get_predecessors_ttl code = | Tnext (_, i) -> let succ = match i with | Inop n | Iop (_,_,_,n) | Iload (_, _,_,_,_,n) | Istore (_,_,_,_,n) | Icall (_,_,_,_,n) | Ibuiltin (_, _, _, n) -> [n] - | Icond (_,_,n1,n2) -> [n1;n2] + | Icond (_,_,n1,n2,_) -> [n1;n2] | Ijumptable (_,ln) -> ln | _ -> [] in List.iter (fun s -> preds := PTree.set s (node::(get_some @@ PTree.get s !preds)) !preds) succ @@ -413,10 +413,10 @@ let rec change_pointers code n n' = function | Ibuiltin(a, b, c, n0) -> assert (n0 == n); Ibuiltin(a, b, c, n') | Ijumptable(a, ln) -> assert (optbool @@ List.find_opt (fun e -> e == n) ln); Ijumptable(a, List.map (fun e -> if (e == n) then n' else e) ln) - | Icond(a, b, n1, n2) -> assert (n1 == n || n2 == n); + | Icond(a, b, n1, n2, i) -> assert (n1 == n || n2 == n); let n1' = if (n1 == n) then n' else n1 in let n2' = if (n2 == n) then n' else n2 - in Icond(a, b, n1', n2') + in Icond(a, b, n1', n2', i) | Inop n0 -> assert (n0 == n); Inop n' | Iop (a, b, c, n0) -> assert (n0 == n); Iop (a, b, c, n') | Iload (a, b, c, d, e, n0) -> assert (n0 == n); Iload (a, b, c, d, e, n') @@ -504,11 +504,12 @@ let rec invert_iconds_trace code = function | n::[] -> code | n :: n' :: t -> let code' = match ptree_get_some n code with - | Icond (c, lr, ifso, ifnot) -> + | Icond (c, lr, ifso, ifnot, i) -> assert (n' == ifso || n' == ifnot); if (n' == ifso) then ( (* Printf.printf "Reversing ifso/ifnot for node %d\n" (P.to_int n); *) - PTree.set n (Icond (Op.negate_condition c, lr, ifnot, ifso)) code ) + let i' = match i with None -> None | Some b -> Some (not b) in + PTree.set n (Icond (Op.negate_condition c, lr, ifnot, ifso, i')) code ) else code | _ -> code in invert_iconds_trace code' (n'::t) diff --git a/backend/Duplicateproof.v b/backend/Duplicateproof.v index 466b4b75..6b598dc7 100644 --- a/backend/Duplicateproof.v +++ b/backend/Duplicateproof.v @@ -23,12 +23,12 @@ Inductive match_inst (dupmap: PTree.t node): instruction -> instruction -> Prop match_inst dupmap (Itailcall s ri lr) (Itailcall s ri lr) | match_inst_builtin: forall n n' ef la br, dupmap!n' = (Some n) -> match_inst dupmap (Ibuiltin ef la br n) (Ibuiltin ef la br n') - | match_inst_cond: forall ifso ifso' ifnot ifnot' c lr, + | match_inst_cond: forall ifso ifso' ifnot ifnot' c lr i i', dupmap!ifso' = (Some ifso) -> dupmap!ifnot' = (Some ifnot) -> - match_inst dupmap (Icond c lr ifso ifnot) (Icond c lr ifso' ifnot') - | match_inst_revcond: forall ifso ifso' ifnot ifnot' c lr, + match_inst dupmap (Icond c lr ifso ifnot i) (Icond c lr ifso' ifnot' i') + | match_inst_revcond: forall ifso ifso' ifnot ifnot' c lr i i', dupmap!ifso' = (Some ifso) -> dupmap!ifnot' = (Some ifnot) -> - match_inst dupmap (Icond c lr ifso ifnot) (Icond (negate_condition c) lr ifnot' ifso') + match_inst dupmap (Icond c lr ifso ifnot i) (Icond (negate_condition c) lr ifnot' ifso' i') | match_inst_jumptable: forall ln ln' r, list_forall2 (fun n n' => (dupmap!n' = (Some n))) ln ln' -> match_inst dupmap (Ijumptable r ln) (Ijumptable r ln') diff --git a/backend/ForwardMoves.v b/backend/ForwardMoves.v index c73b0213..7cfd411f 100644 --- a/backend/ForwardMoves.v +++ b/backend/ForwardMoves.v @@ -250,7 +250,7 @@ Fixpoint kill_builtin_res (res : builtin_res reg) (rel : RELATION.t) := Definition apply_instr instr x := match instr with | Inop _ - | Icond _ _ _ _ + | Icond _ _ _ _ _ | Ijumptable _ _ | Istore _ _ _ _ _ => Some x | Iop Omove (src :: nil) dst _ => Some (move src dst x) @@ -309,8 +309,8 @@ Definition transf_instr (fmap : option (PMap.t RB.t)) Icall sig ros (subst_args fmap pc args) dst s | Itailcall sig ros args => Itailcall sig ros (subst_args fmap pc args) - | Icond cond args s1 s2 => - Icond cond (subst_args fmap pc args) s1 s2 + | Icond cond args s1 s2 i => + Icond cond (subst_args fmap pc args) s1 s2 i | Ijumptable arg tbl => Ijumptable (subst_arg fmap pc arg) tbl | Ireturn (Some arg) => diff --git a/backend/Inlining.v b/backend/Inlining.v index 9cf535b9..8c7e1898 100644 --- a/backend/Inlining.v +++ b/backend/Inlining.v @@ -397,9 +397,9 @@ Definition expand_instr (ctx: context) (pc: node) (i: instruction): mon unit := | Ibuiltin ef args res s => set_instr (spc ctx pc) (Ibuiltin ef (map (sbuiltinarg ctx) args) (sbuiltinres ctx res) (spc ctx s)) - | Icond cond args s1 s2 => + | Icond cond args s1 s2 info => set_instr (spc ctx pc) - (Icond cond (sregs ctx args) (spc ctx s1) (spc ctx s2)) + (Icond cond (sregs ctx args) (spc ctx s1) (spc ctx s2) info) | Ijumptable r tbl => set_instr (spc ctx pc) (Ijumptable (sreg ctx r) (List.map (spc ctx) tbl)) diff --git a/backend/Inliningspec.v b/backend/Inliningspec.v index e20fb373..eba026ec 100644 --- a/backend/Inliningspec.v +++ b/backend/Inliningspec.v @@ -312,9 +312,9 @@ Inductive tr_instr: context -> node -> instruction -> code -> Prop := match res with BR r => Ple r ctx.(mreg) | _ => True end -> c!(spc ctx pc) = Some (Ibuiltin ef (map (sbuiltinarg ctx) args) (sbuiltinres ctx res) (spc ctx s)) -> tr_instr ctx pc (Ibuiltin ef args res s) c - | tr_cond: forall ctx pc cond args s1 s2 c, - c!(spc ctx pc) = Some (Icond cond (sregs ctx args) (spc ctx s1) (spc ctx s2)) -> - tr_instr ctx pc (Icond cond args s1 s2) c + | tr_cond: forall ctx pc cond args s1 s2 c i, + c!(spc ctx pc) = Some (Icond cond (sregs ctx args) (spc ctx s1) (spc ctx s2) i) -> + tr_instr ctx pc (Icond cond args s1 s2 i) c | tr_jumptable: forall ctx pc r tbl c, c!(spc ctx pc) = Some (Ijumptable (sreg ctx r) (List.map (spc ctx) tbl)) -> tr_instr ctx pc (Ijumptable r tbl) c diff --git a/backend/LTL.v b/backend/LTL.v index ee8b4826..3edd60a2 100644 --- a/backend/LTL.v +++ b/backend/LTL.v @@ -37,7 +37,7 @@ Inductive instruction: Type := | Ltailcall (sg: signature) (ros: mreg + ident) | Lbuiltin (ef: external_function) (args: list (builtin_arg loc)) (res: builtin_res mreg) | Lbranch (s: node) - | Lcond (cond: condition) (args: list mreg) (s1 s2: node) + | Lcond (cond: condition) (args: list mreg) (s1 s2: node) (info: option bool) | Ljumptable (arg: mreg) (tbl: list node) | Lreturn. @@ -263,11 +263,11 @@ Inductive step: state -> trace -> state -> Prop := | exec_Lbranch: forall s f sp pc bb rs m, step (Block s f sp (Lbranch pc :: bb) rs m) E0 (State s f sp pc rs m) - | exec_Lcond: forall s f sp cond args pc1 pc2 bb rs b pc rs' m, + | exec_Lcond: forall s f sp cond args pc1 pc2 bb rs b pc rs' m i, eval_condition cond (reglist rs args) m = Some b -> pc = (if b then pc1 else pc2) -> rs' = undef_regs (destroyed_by_cond cond) rs -> - step (Block s f sp (Lcond cond args pc1 pc2 :: bb) rs m) + step (Block s f sp (Lcond cond args pc1 pc2 i :: bb) rs m) E0 (State s f sp pc rs' m) | exec_Ljumptable: forall s f sp arg tbl bb rs m n pc rs', rs (R arg) = Vint n -> @@ -328,7 +328,7 @@ Fixpoint successors_block (b: bblock) : list node := | nil => nil (**r should never happen *) | Ltailcall _ _ :: _ => nil | Lbranch s :: _ => s :: nil - | Lcond _ _ s1 s2 :: _ => s1 :: s2 :: nil + | Lcond _ _ s1 s2 _ :: _ => s1 :: s2 :: nil | Ljumptable _ tbl :: _ => tbl | Lreturn :: _ => nil | instr :: b' => successors_block b' diff --git a/backend/Linearize.v b/backend/Linearize.v index 4216958c..66b36428 100644 --- a/backend/Linearize.v +++ b/backend/Linearize.v @@ -179,7 +179,7 @@ Fixpoint linearize_block (b: LTL.bblock) (k: code) : code := Lbuiltin ef args res :: linearize_block b' k | LTL.Lbranch s :: b' => add_branch s k - | LTL.Lcond cond args s1 s2 :: b' => + | LTL.Lcond cond args s1 s2 _ :: b' => if starts_with s1 k then Lcond (negate_condition cond) args s2 :: add_branch s1 k else diff --git a/backend/Linearizeaux.ml b/backend/Linearizeaux.ml index eed58f8d..c9a5d620 100644 --- a/backend/Linearizeaux.ml +++ b/backend/Linearizeaux.ml @@ -81,7 +81,7 @@ let basic_blocks f joins = | [] -> assert false | Lbranch s :: _ -> next_in_block blk minpc s | Ltailcall (sig0, ros) :: _ -> end_block blk minpc - | Lcond (cond, args, ifso, ifnot) :: _ -> + | Lcond (cond, args, ifso, ifnot, _) :: _ -> end_block blk minpc; start_block ifso; start_block ifnot | Ljumptable(arg, tbl) :: _ -> end_block blk minpc; List.iter start_block tbl @@ -165,7 +165,7 @@ let forward_sequences code entry = | Lbuiltin _ -> assert false | Ltailcall _ | Lreturn -> ([], []) | Lbranch n -> let ln, rem = traverse_fallthrough code n in (ln, rem) - | Lcond (_, _, ifso, ifnot) -> let ln, rem = traverse_fallthrough code ifnot in (ln, [ifso] @ rem) + | Lcond (_, _, ifso, ifnot, _) -> let ln, rem = traverse_fallthrough code ifnot in (ln, [ifso] @ rem) | Ljumptable(_, ln) -> match ln with | [] -> ([], []) | n :: ln -> let lln, rem = traverse_fallthrough code n in (lln, ln @ rem) @@ -219,7 +219,7 @@ let can_be_merged code s s' = | Lop _ | Lload _ | Lgetstack _ | Lsetstack _ | Lstore _ | Lcall _ | Lbuiltin _ | Ltailcall _ | Lreturn -> false | Lbranch n -> n == first_s' - | Lcond (_, _, ifso, ifnot) -> ifnot == first_s' + | Lcond (_, _, ifso, ifnot, _) -> ifnot == first_s' | Ljumptable (_, ln) -> match ln with | [] -> false @@ -303,7 +303,7 @@ let get_loop_edges code entry = | Lbuiltin _ -> assert false | Ltailcall _ | Lreturn -> [] | Lbranch n -> [n] - | Lcond (_, _, ifso, ifnot) -> [ifso; ifnot] + | Lcond (_, _, ifso, ifnot, _) -> [ifso; ifnot] | Ljumptable(_, ln) -> ln ) in dfs_visit code (Some node) next_visits; visited := PTree.set node Visited !visited; @@ -371,7 +371,7 @@ let construct_depmap code entry fs = match (last_element bb) with | Ltailcall _ | Lreturn -> [] | Lbranch n -> (check_and_update_depmap node n; [n]) - | Lcond (_, _, ifso, ifnot) -> begin + | Lcond (_, _, ifso, ifnot, _) -> begin check_and_update_depmap node ifso; check_and_update_depmap node ifnot; [ifso; ifnot] diff --git a/backend/Liveness.v b/backend/Liveness.v index afe11ae6..9652b363 100644 --- a/backend/Liveness.v +++ b/backend/Liveness.v @@ -94,7 +94,7 @@ Definition transfer | Ibuiltin ef args res s => reg_list_live (params_of_builtin_args args) (reg_list_dead (params_of_builtin_res res) after) - | Icond cond args ifso ifnot => + | Icond cond args ifso ifnot _ => reg_list_live args after | Ijumptable arg tbl => reg_live arg after diff --git a/backend/PrintLTL.ml b/backend/PrintLTL.ml index b309a9f2..f173e374 100644 --- a/backend/PrintLTL.ml +++ b/backend/PrintLTL.ml @@ -83,7 +83,7 @@ let print_instruction pp succ = function (print_builtin_args loc) args | Lbranch s -> print_succ pp s succ - | Lcond(cond, args, s1, s2) -> + | Lcond(cond, args, s1, s2, _) -> fprintf pp "if (%a) goto %d else goto %d" (print_condition mreg) (cond, args) (P.to_int s1) (P.to_int s2) diff --git a/backend/PrintRTL.ml b/backend/PrintRTL.ml index c25773e5..5eab9901 100644 --- a/backend/PrintRTL.ml +++ b/backend/PrintRTL.ml @@ -75,7 +75,7 @@ let print_instruction pp (pc, i) = (name_of_external ef) (print_builtin_args reg) args; print_succ pp s (pc - 1) - | Icond(cond, args, s1, s2) -> + | Icond(cond, args, s1, s2, _) -> fprintf pp "if (%a) goto %d else goto %d\n" (PrintOp.print_condition reg) (cond, args) (P.to_int s1) (P.to_int s2) diff --git a/backend/PrintXTL.ml b/backend/PrintXTL.ml index 1c7655fb..d1b79623 100644 --- a/backend/PrintXTL.ml +++ b/backend/PrintXTL.ml @@ -104,7 +104,7 @@ let print_instruction pp succ = function (print_builtin_args var) args | Xbranch s -> print_succ pp s succ - | Xcond(cond, args, s1, s2) -> + | Xcond(cond, args, s1, s2, _) -> fprintf pp "if (%a) goto %d else goto %d" (print_condition var) (cond, args) (P.to_int s1) (P.to_int s2) diff --git a/backend/RTL.v b/backend/RTL.v index 29a49311..dec59ca2 100644 --- a/backend/RTL.v +++ b/backend/RTL.v @@ -67,11 +67,12 @@ Inductive instruction: Type := (** [Ibuiltin ef args dest succ] calls the built-in function identified by [ef], giving it the values of [args] as arguments. It stores the return value in [dest] and branches to [succ]. *) - | Icond: condition -> list reg -> node -> node -> instruction - (** [Icond cond args ifso ifnot] evaluates the boolean condition + | Icond: condition -> list reg -> node -> node -> option bool -> instruction + (** [Icond cond args ifso ifnot info] evaluates the boolean condition [cond] over the values of registers [args]. If the condition is true, it transitions to [ifso]. If the condition is false, - it transitions to [ifnot]. *) + it transitions to [ifnot]. [info] is a ghost field there to provide + information relative to branch prediction. *) | Ijumptable: reg -> list node -> instruction (** [Ijumptable arg tbl] transitions to the node that is the [n]-th element of the list [tbl], where [n] is the unsigned integer @@ -262,8 +263,8 @@ Inductive step: state -> trace -> state -> Prop := step (State s f sp pc rs m) t (State s f sp pc' (regmap_setres res vres rs) m') | exec_Icond: - forall s f sp pc rs m cond args ifso ifnot b pc', - (fn_code f)!pc = Some(Icond cond args ifso ifnot) -> + forall s f sp pc rs m cond args ifso ifnot b pc' predb, + (fn_code f)!pc = Some(Icond cond args ifso ifnot predb) -> eval_condition cond rs##args m = Some b -> pc' = (if b then ifso else ifnot) -> step (State s f sp pc rs m) @@ -403,7 +404,7 @@ Definition successors_instr (i: instruction) : list node := | Icall sig ros args res s => s :: nil | Itailcall sig ros args => nil | Ibuiltin ef args res s => s :: nil - | Icond cond args ifso ifnot => ifso :: ifnot :: nil + | Icond cond args ifso ifnot _ => ifso :: ifnot :: nil | Ijumptable arg tbl => tbl | Ireturn optarg => nil end. @@ -424,7 +425,7 @@ Definition instr_uses (i: instruction) : list reg := | Itailcall sig (inl r) args => r :: args | Itailcall sig (inr id) args => args | Ibuiltin ef args res s => params_of_builtin_args args - | Icond cond args ifso ifnot => args + | Icond cond args ifso ifnot _ => args | Ijumptable arg tbl => arg :: nil | Ireturn None => nil | Ireturn (Some arg) => arg :: nil @@ -442,7 +443,7 @@ Definition instr_defs (i: instruction) : option reg := | Itailcall sig ros args => None | Ibuiltin ef args res s => match res with BR r => Some r | _ => None end - | Icond cond args ifso ifnot => None + | Icond cond args ifso ifnot _ => None | Ijumptable arg tbl => None | Ireturn optarg => None end. @@ -485,7 +486,7 @@ Definition max_reg_instr (m: positive) (pc: node) (i: instruction) := | Ibuiltin ef args res s => fold_left Pos.max (params_of_builtin_args args) (fold_left Pos.max (params_of_builtin_res res) m) - | Icond cond args ifso ifnot => fold_left Pos.max args m + | Icond cond args ifso ifnot _ => fold_left Pos.max args m | Ijumptable arg tbl => Pos.max arg m | Ireturn None => m | Ireturn (Some arg) => Pos.max arg m diff --git a/backend/RTLgen.v b/backend/RTLgen.v index 2c27944a..ac98f3a1 100644 --- a/backend/RTLgen.v +++ b/backend/RTLgen.v @@ -479,7 +479,7 @@ with transl_condexpr (map: mapping) (a: condexpr) (ntrue nfalse: node) match a with | CEcond c al => do rl <- alloc_regs map al; - do nt <- add_instr (Icond c rl ntrue nfalse); + do nt <- add_instr (Icond c rl ntrue nfalse None); transl_exprlist map al rl nt | CEcondition a b c => do nc <- transl_condexpr map c ntrue nfalse; diff --git a/backend/RTLgenspec.v b/backend/RTLgenspec.v index 92b48e2b..30ad7d82 100644 --- a/backend/RTLgenspec.v +++ b/backend/RTLgenspec.v @@ -744,9 +744,9 @@ Inductive tr_expr (c: code): with tr_condition (c: code): mapping -> list reg -> condexpr -> node -> node -> node -> Prop := - | tr_CEcond: forall map pr cond bl ns ntrue nfalse n1 rl, + | tr_CEcond: forall map pr cond bl ns ntrue nfalse n1 rl i, tr_exprlist c map pr bl ns n1 rl -> - c!n1 = Some (Icond cond rl ntrue nfalse) -> + c!n1 = Some (Icond cond rl ntrue nfalse i) -> tr_condition c map pr (CEcond cond bl) ns ntrue nfalse | tr_CEcondition: forall map pr a1 a2 a3 ns ntrue nfalse n2 n3, tr_condition c map pr a1 ns n2 n3 -> diff --git a/backend/RTLtyping.v b/backend/RTLtyping.v index 857f2211..15ed6d8a 100644 --- a/backend/RTLtyping.v +++ b/backend/RTLtyping.v @@ -139,11 +139,11 @@ Inductive wt_instr : instruction -> Prop := valid_successor s -> wt_instr (Ibuiltin ef args res s) | wt_Icond: - forall cond args s1 s2, + forall cond args s1 s2 i, map env args = type_of_condition cond -> valid_successor s1 -> valid_successor s2 -> - wt_instr (Icond cond args s1 s2) + wt_instr (Icond cond args s1 s2 i) | wt_Ijumptable: forall arg tbl, env arg = Tint -> @@ -313,7 +313,7 @@ Definition type_instr (e: S.typenv) (i: instruction) : res S.typenv := | _ => type_builtin_args e args sig.(sig_args) end; type_builtin_res e1 res (proj_sig_res sig) - | Icond cond args s1 s2 => + | Icond cond args s1 s2 _ => do x1 <- check_successor s1; do x2 <- check_successor s2; S.set_list e args (type_of_condition cond) diff --git a/backend/Regalloc.ml b/backend/Regalloc.ml index f2658b04..ffe26933 100644 --- a/backend/Regalloc.ml +++ b/backend/Regalloc.ml @@ -295,8 +295,8 @@ let block_of_RTL_instr funsig tyenv = function (Xbuiltin(ef, args2, res2) :: movelist (params_of_builtin_res res2) (params_of_builtin_res res1) [Xbranch s]) - | RTL.Icond(cond, args, s1, s2) -> - [Xcond(cond, vregs tyenv args, s1, s2)] + | RTL.Icond(cond, args, s1, s2, i) -> + [Xcond(cond, vregs tyenv args, s1, s2, i)] | RTL.Ijumptable(arg, tbl) -> [Xjumptable(vreg tyenv arg, tbl)] | RTL.Ireturn None -> @@ -380,7 +380,7 @@ let live_before instr after = vset_addargs args (vset_removeres res after) | Xbranch s -> after - | Xcond(cond, args, s1, s2) -> + | Xcond(cond, args, s1, s2, _) -> List.fold_right VSet.add args after | Xjumptable(arg, tbl) -> VSet.add arg after @@ -575,7 +575,7 @@ let spill_costs f = charge_list 10 1 (params_of_builtin_res res) end | Xbranch _ -> () - | Xcond(cond, args, _, _) -> + | Xcond(cond, args, _, _, _) -> charge_list 10 1 args | Xjumptable(arg, _) -> charge 10 1 arg @@ -718,7 +718,7 @@ let add_interfs_instr g instr live = end | Xbranch s -> () - | Xcond(cond, args, s1, s2) -> + | Xcond(cond, args, s1, s2, _) -> add_interfs_destroyed g live (destroyed_by_cond cond) | Xjumptable(arg, tbl) -> add_interfs_destroyed g live destroyed_by_jumptable @@ -797,7 +797,7 @@ let tospill_instr alloc instr ts = (addlist_tospill alloc (params_of_builtin_res res) ts) | Xbranch s -> ts - | Xcond(cond, args, s1, s2) -> + | Xcond(cond, args, s1, s2, _) -> addlist_tospill alloc args ts | Xjumptable(arg, tbl) -> add_tospill alloc arg ts @@ -990,9 +990,9 @@ let spill_instr tospill eqs instr = (c1 @ Xbuiltin(ef, args', res') :: c2, eqs2) | Xbranch s -> ([instr], eqs) - | Xcond(cond, args, s1, s2) -> + | Xcond(cond, args, s1, s2, i) -> let (args', c1, eqs1) = reload_vars tospill eqs args in - (c1 @ [Xcond(cond, args', s1, s2)], eqs1) + (c1 @ [Xcond(cond, args', s1, s2, i)], eqs1) | Xjumptable(arg, tbl) -> let (arg', c1, eqs1) = reload_var tospill eqs arg in (c1 @ [Xjumptable(arg', tbl)], eqs1) @@ -1128,8 +1128,8 @@ let transl_instr alloc instr k = AST.map_builtin_res (mreg_of alloc) res) :: k | Xbranch s -> LTL.Lbranch s :: [] - | Xcond(cond, args, s1, s2) -> - LTL.Lcond(cond, mregs_of alloc args, s1, s2) :: [] + | Xcond(cond, args, s1, s2, i) -> + LTL.Lcond(cond, mregs_of alloc args, s1, s2, i) :: [] | Xjumptable(arg, tbl) -> LTL.Ljumptable(mreg_of alloc arg, tbl) :: [] | Xreturn optarg -> diff --git a/backend/Renumber.v b/backend/Renumber.v index 7ba16658..2727b979 100644 --- a/backend/Renumber.v +++ b/backend/Renumber.v @@ -48,7 +48,7 @@ Definition renum_instr (i: instruction) : instruction := | Icall sg ros args res s => Icall sg ros args res (renum_pc s) | Itailcall sg ros args => i | Ibuiltin ef args res s => Ibuiltin ef args res (renum_pc s) - | Icond cond args s1 s2 => Icond cond args (renum_pc s1) (renum_pc s2) + | Icond cond args s1 s2 info => Icond cond args (renum_pc s1) (renum_pc s2) info | Ijumptable arg tbl => Ijumptable arg (List.map renum_pc tbl) | Ireturn or => i end. diff --git a/backend/Splitting.ml b/backend/Splitting.ml index 78eb66a5..3ca45c3b 100644 --- a/backend/Splitting.ml +++ b/backend/Splitting.ml @@ -162,8 +162,8 @@ let ren_instr f maps pc i = | Ibuiltin(ef, args, res, s) -> Ibuiltin(ef, List.map (AST.map_builtin_arg (ren_reg before)) args, AST.map_builtin_res (ren_reg after) res, s) - | Icond(cond, args, s1, s2) -> - Icond(cond, ren_regs before args, s1, s2) + | Icond(cond, args, s1, s2, i) -> + Icond(cond, ren_regs before args, s1, s2, i) | Ijumptable(arg, tbl) -> Ijumptable(ren_reg before arg, tbl) | Ireturn optarg -> diff --git a/backend/Tunneling.v b/backend/Tunneling.v index da1ce45a..a4c4a195 100644 --- a/backend/Tunneling.v +++ b/backend/Tunneling.v @@ -78,11 +78,11 @@ Definition record_gotos (f: LTL.function) : U.t := Definition tunnel_instr (uf: U.t) (i: instruction) : instruction := match i with | Lbranch s => Lbranch (U.repr uf s) - | Lcond cond args s1 s2 => + | Lcond cond args s1 s2 info => let s1' := U.repr uf s1 in let s2' := U.repr uf s2 in if peq s1' s2' then Lbranch s1' - else Lcond cond args s1' s2' + else Lcond cond args s1' s2' info | Ljumptable arg tbl => Ljumptable arg (List.map (U.repr uf) tbl) | _ => i end. diff --git a/backend/Unusedglob.v b/backend/Unusedglob.v index 1b5f2547..93ca7af4 100644 --- a/backend/Unusedglob.v +++ b/backend/Unusedglob.v @@ -53,7 +53,7 @@ Definition ref_instruction (i: instruction) : list ident := | Itailcall _ (inl r) _ => nil | Itailcall _ (inr id) _ => id :: nil | Ibuiltin _ args _ _ => globals_of_builtin_args args - | Icond cond _ _ _ => nil + | Icond cond _ _ _ _ => nil | Ijumptable _ _ => nil | Ireturn _ => nil end. diff --git a/backend/ValueAnalysis.v b/backend/ValueAnalysis.v index e25d2e5f..2e79d1a9 100644 --- a/backend/ValueAnalysis.v +++ b/backend/ValueAnalysis.v @@ -156,7 +156,7 @@ Definition transfer (f: function) (rm: romem) (pc: node) (ae: aenv) (am: amem) : VA.Bot | Some(Ibuiltin ef args res s) => transfer_builtin ae am rm ef args res - | Some(Icond cond args s1 s2) => + | Some(Icond cond args s1 s2 _) => VA.State ae am | Some(Ijumptable arg tbl) => VA.State ae am diff --git a/backend/XTL.ml b/backend/XTL.ml index c496fafb..1d8e89c0 100644 --- a/backend/XTL.ml +++ b/backend/XTL.ml @@ -36,7 +36,7 @@ type instruction = | Xtailcall of signature * (var, ident) sum * var list | Xbuiltin of external_function * var builtin_arg list * var builtin_res | Xbranch of node - | Xcond of condition * var list * node * node + | Xcond of condition * var list * node * node * bool option | Xjumptable of var * node list | Xreturn of var list @@ -105,7 +105,7 @@ let twin_reg r = let rec successors_block = function | Xbranch s :: _ -> [s] | Xtailcall(sg, vos, args) :: _ -> [] - | Xcond(cond, args, s1, s2) :: _ -> [s1; s2] + | Xcond(cond, args, s1, s2, _) :: _ -> [s1; s2] | Xjumptable(arg, tbl) :: _ -> tbl | Xreturn _:: _ -> [] | instr :: blk -> successors_block blk @@ -179,7 +179,7 @@ let type_instr = function type_builtin_res res (proj_sig_res sg) | Xbranch s -> () - | Xcond(cond, args, s1, s2) -> + | Xcond(cond, args, s1, s2, _) -> set_vars_type args (type_of_condition cond) | Xjumptable(arg, tbl) -> set_var_type arg Tint diff --git a/backend/XTL.mli b/backend/XTL.mli index b4b77fab..7b7f7186 100644 --- a/backend/XTL.mli +++ b/backend/XTL.mli @@ -37,7 +37,7 @@ type instruction = | Xtailcall of signature * (var, ident) sum * var list | Xbuiltin of external_function * var builtin_arg list * var builtin_res | Xbranch of node - | Xcond of condition * var list * node * node + | Xcond of condition * var list * node * node * bool option | Xjumptable of var * node list | Xreturn of var list -- cgit From 4fe7ec168a9ce2c8c6e04d7f56729fd7a5758ce1 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 11 Mar 2020 17:30:37 +0100 Subject: [BROKEN] Started to change the trace selection --- backend/Duplicateaux.ml | 92 ++++++++++++++++++++----------------------------- 1 file changed, 38 insertions(+), 54 deletions(-) diff --git a/backend/Duplicateaux.ml b/backend/Duplicateaux.ml index 86bc06c9..c379faf3 100644 --- a/backend/Duplicateaux.ml +++ b/backend/Duplicateaux.ml @@ -2,19 +2,6 @@ open RTL open Maps open Camlcoq -(* TTL : IR emphasizing the preferred next node *) -module TTL = struct - type instruction = - | Tleaf of RTL.instruction - | Tnext of node * RTL.instruction - - type code = instruction PTree.t -end;; - -open TTL - -(** RTL to TTL *) - let get_some = function | None -> failwith "Did not get some" | Some thing -> thing @@ -242,37 +229,26 @@ let get_directions code entrypoint = !directions end -let to_ttl_inst direction = function -| Ireturn o -> Tleaf (Ireturn o) -| Inop n -> Tnext (n, Inop n) -| Iop (op, lr, r, n) -> Tnext (n, Iop(op, lr, r, n)) -| Iload (tm, m, a, lr, r, n) -> Tnext (n, Iload(tm, m, a, lr, r, n)) -| Istore (m, a, lr, r, n) -> Tnext (n, Istore(m, a, lr, r, n)) -| Icall (s, ri, lr, r, n) -> Tleaf (Icall(s, ri, lr, r, n)) -| Itailcall (s, ri, lr) -> Tleaf (Itailcall(s, ri, lr)) -| Ibuiltin (ef, lbr, br, n) -> Tleaf (Ibuiltin(ef, lbr, br, n)) -| Icond (cond, lr, n, n', i) -> (match direction with - | false -> Tnext (n', Icond(cond, lr, n, n', i)) - | true -> Tnext (n, Icond(cond, lr, n, n', i))) -| Ijumptable (r, ln) -> Tleaf (Ijumptable(r, ln)) - -let rec to_ttl_code_rec directions = function +let update_direction direction = function +| Icond (cond, lr, n, n', _) -> Icond (cond, lr, n, n', Some direction) +| i -> i + +let rec update_direction_rec directions = function | [] -> PTree.empty | m::lm -> let (n, i) = m in let direction = get_some @@ PTree.get n directions - in PTree.set n (to_ttl_inst direction i) (to_ttl_code_rec directions lm) + in PTree.set n (update_direction direction i) (update_direction_rec directions lm) -let to_ttl_code code entrypoint = +let update_directions code entrypoint = let directions = get_directions code entrypoint in begin (* Printf.printf "Ifso directions: "; ptree_printbool directions; Printf.printf "\n"; *) - Random.init(0); (* using same seed to make it deterministic *) - to_ttl_code_rec directions (PTree.elements code) + update_direction_rec directions (PTree.elements code) end -(** Trace selection on TTL *) +(** Trace selection *) let rec exists_false_rec = function | [] -> false @@ -280,7 +256,7 @@ let rec exists_false_rec = function let exists_false boolmap = exists_false_rec (PTree.elements boolmap) -(* DFS on TTL to guide the exploration *) +(* DFS using prediction info to guide the exploration *) let dfs code entrypoint = let visited = ref (PTree.map (fun n i -> false) code) in let rec dfs_list code = function @@ -291,22 +267,21 @@ let dfs code entrypoint = visited := PTree.set node true !visited; match PTree.get node code with | None -> failwith "No such node" - | Some ti -> [node] @ match ti with - | Tleaf i -> (match i with - | Icall(_, _, _, _, n) -> dfs_list code [n] - | Ibuiltin(_, _, _, n) -> dfs_list code [n] - | Ijumptable(_, ln) -> dfs_list code ln - | Itailcall _ | Ireturn _ -> [] - | _ -> failwith "Tleaf case not handled in dfs" ) - | Tnext (n,i) -> (dfs_list code [n]) @ match i with - | Icond (_, _, n1, n2, _) -> dfs_list code [n1; n2] - | Inop _ | Iop _ | Iload _ | Istore _ -> [] - | _ -> failwith "Tnext case not handled in dfs" + | Some i -> [node] @ match i with + | Icall(_, _, _, _, n) -> dfs_list code [n] + | Ibuiltin(_, _, _, n) -> dfs_list code [n] + | Ijumptable(_, ln) -> dfs_list code ln + | Itailcall _ | Ireturn _ -> [] + | Inop _ | Iop _ | Iload _ | Istore _ -> [] + | Icond (_, _, n1, n2, info) -> match info with + | Some false -> dfs_list code [n2; n1] + | _ -> dfs_list code [n1; n2] end else [] in node_dfs @ (dfs_list code ln) in dfs_list code [entrypoint] +(* let get_predecessors_ttl code = let preds = ref (PTree.map (fun n i -> []) code) in let process_inst (node, ti) = match ti with @@ -322,8 +297,7 @@ let get_predecessors_ttl code = List.iter process_inst (PTree.elements code); !preds end - -let rtl_proj code = PTree.map (fun n ti -> match ti with Tleaf i | Tnext(_, i) -> i) code +*) let rec select_unvisited_node is_visited = function | [] -> failwith "Empty list" @@ -332,11 +306,21 @@ let rec select_unvisited_node is_visited = function let best_successor_of node code is_visited = match (PTree.get node code) with | None -> failwith "No such node in the code" - | Some ti -> match ti with - | Tleaf _ -> None - | Tnext (n,_) -> if not (ptree_get_some n is_visited) then Some n - else None - + | Some i -> + let next_node = match i with + | Inop n -> Some n + | Iop (_, _, _, n) -> Some n + | Iload (_, _, _, _, _, n) -> Some n + | Istore (_, _, _, _, n) -> Some n + | Icall (_, _, _, _, n) -> Some n + | Ibuiltin (_, _, _, n) -> Some n + | Icond (_, _, n1, n2, ob) -> (match ob with None -> None | Some false -> Some n2 | Some true -> Some n1) + | _ -> None + in match next_node with + | None -> None + | Some n -> if not (ptree_get_some n is_visited) then Some n else None + +(* FIXME - could be improved by selecting in priority the predicted paths *) let best_predecessor_of node predecessors order is_visited = match (PTree.get node predecessors) with | None -> failwith "No predecessor list found" @@ -347,7 +331,7 @@ let best_predecessor_of node predecessors order is_visited = * "Trace Selection for Compiling Large C Application Programs to Microcode" *) let select_traces code entrypoint = let order = dfs code entrypoint in - let predecessors = get_predecessors_ttl code in + let predecessors = get_predecessors_rtl code in let traces = ref [] in let is_visited = ref (PTree.map (fun n i -> false) code) in begin (* mark all nodes visited *) while exists_false !is_visited do (* while (there are unvisited nodes) *) @@ -524,7 +508,7 @@ let rec invert_iconds code = function let duplicate_aux f = let entrypoint = f.fn_entrypoint in let code = f.fn_code in - let traces = select_traces (to_ttl_code code entrypoint) entrypoint in + let traces = select_traces code entrypoint in let icond_code = invert_iconds code traces in let preds = get_predecessors_rtl icond_code in if !Clflags.option_fduplicate >= 1 then -- cgit From c77d6412f132bf6c09189e5f2d3c8799440f1977 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Thu, 12 Mar 2020 14:25:30 +0100 Subject: Fixed typo in Duplicate: dfs --- backend/Duplicateaux.ml | 5 ++++- 1 file changed, 4 insertions(+), 1 deletion(-) diff --git a/backend/Duplicateaux.ml b/backend/Duplicateaux.ml index c379faf3..cabcf1fd 100644 --- a/backend/Duplicateaux.ml +++ b/backend/Duplicateaux.ml @@ -270,9 +270,12 @@ let dfs code entrypoint = | Some i -> [node] @ match i with | Icall(_, _, _, _, n) -> dfs_list code [n] | Ibuiltin(_, _, _, n) -> dfs_list code [n] + | Iop (_, _, _, n) -> dfs_list code [n] + | Iload (_, _, _, _, _, n) -> dfs_list code [n] + | Istore (_, _, _, _, n) -> dfs_list code [n] + | Inop n -> dfs_list code [n] | Ijumptable(_, ln) -> dfs_list code ln | Itailcall _ | Ireturn _ -> [] - | Inop _ | Iop _ | Iload _ | Istore _ -> [] | Icond (_, _, n1, n2, info) -> match info with | Some false -> dfs_list code [n2; n1] | _ -> dfs_list code [n1; n2] -- cgit From 786ada1bd193e3995b948e2bd11d6285654a5c6a Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Thu, 12 Mar 2020 16:08:47 +0100 Subject: Correcting a few bugs in trace selection and expansion --- backend/Duplicateaux.ml | 15 +++++++-------- backend/Linearizeaux.ml | 20 +++++++++++++------- 2 files changed, 20 insertions(+), 15 deletions(-) diff --git a/backend/Duplicateaux.ml b/backend/Duplicateaux.ml index cabcf1fd..37647714 100644 --- a/backend/Duplicateaux.ml +++ b/backend/Duplicateaux.ml @@ -198,9 +198,8 @@ let do_loop_heuristic code cond ifso ifnot is_loop_header = let get_directions code entrypoint = let bfs_order = bfs code entrypoint and is_loop_header = get_loop_headers code entrypoint - and directions = ref (PTree.map (fun n i -> false) code) (* false <=> fallthru *) + and directions = ref (PTree.map (fun n i -> None) code) (* None <=> no predicted direction *) in begin - (* Printf.printf "Loop headers: "; *) (* ptree_printbool is_loop_header; *) (* Printf.printf "\n"; *) List.iter (fun n -> @@ -217,11 +216,10 @@ let get_directions code entrypoint = | None -> preferred := do_heur code cond ifso ifnot is_loop_header | Some _ -> () ) heuristics; - (match !preferred with None -> (Printf.printf "\tRANDOM\n"; preferred := Some (Random.bool ())) | Some _ -> ()); - directions := PTree.set n (get_some !preferred) !directions; + directions := PTree.set n !preferred !directions; (match !preferred with | Some false -> Printf.printf "\tFALLTHROUGH\n" - | Some true -> Printf.printf "\tBRANCH\n" - | None -> ()); + | Some true -> Printf.printf "\tBRANCH\n" + | None -> Printf.printf "\tUNSURE\n"); Printf.printf "---------------------------------------\n" end | _ -> () @@ -230,7 +228,7 @@ let get_directions code entrypoint = end let update_direction direction = function -| Icond (cond, lr, n, n', _) -> Icond (cond, lr, n, n', Some direction) +| Icond (cond, lr, n, n', _) -> Icond (cond, lr, n, n', direction) | i -> i let rec update_direction_rec directions = function @@ -239,6 +237,7 @@ let rec update_direction_rec directions = function in let direction = get_some @@ PTree.get n directions in PTree.set n (update_direction direction i) (update_direction_rec directions lm) +(* Uses branch prediction to write prediction annotations in Icond *) let update_directions code entrypoint = let directions = get_directions code entrypoint in begin @@ -510,7 +509,7 @@ let rec invert_iconds code = function let duplicate_aux f = let entrypoint = f.fn_entrypoint in - let code = f.fn_code in + let code = update_directions (f.fn_code) entrypoint in let traces = select_traces code entrypoint in let icond_code = invert_iconds code traces in let preds = get_predecessors_rtl icond_code in diff --git a/backend/Linearizeaux.ml b/backend/Linearizeaux.ml index c9a5d620..5b3384f2 100644 --- a/backend/Linearizeaux.ml +++ b/backend/Linearizeaux.ml @@ -154,7 +154,7 @@ let forward_sequences code entry = let visited = ref (PTree.map (fun n i -> false) code) in (* returns the list of traversed nodes, and a list of nodes to start traversing next *) let rec traverse_fallthrough code node = - (* Printf.printf "Traversing %d..\n" (P.to_int node); *) + Printf.printf "Traversing %d..\n" (P.to_int node); if not (get_some @@ PTree.get node !visited) then begin visited := PTree.set node true !visited; match PTree.get node code with @@ -163,12 +163,13 @@ let forward_sequences code entry = let ln, rem = match (last_element bb) with | Lop _ | Lload _ | Lgetstack _ | Lsetstack _ | Lstore _ | Lcall _ | Lbuiltin _ -> assert false - | Ltailcall _ | Lreturn -> ([], []) + | Ltailcall _ | Lreturn -> begin Printf.printf "STOP tailcall/return\n"; ([], []) end | Lbranch n -> let ln, rem = traverse_fallthrough code n in (ln, rem) - | Lcond (_, _, ifso, ifnot, _) -> let ln, rem = traverse_fallthrough code ifnot in (ln, [ifso] @ rem) - | Ljumptable(_, ln) -> match ln with - | [] -> ([], []) - | n :: ln -> let lln, rem = traverse_fallthrough code n in (lln, ln @ rem) + | Lcond (_, _, ifso, ifnot, info) -> (match info with + | None -> begin Printf.printf "STOP Lcond None\n"; ([], [ifso; ifnot]) end + | Some false -> let ln, rem = traverse_fallthrough code ifnot in (ln, [ifso] @ rem) + | Some true -> failwith "Inconsistency detected: ifnot is not the preferred branch") + | Ljumptable(_, ln) -> begin Printf.printf "STOP Ljumptable\n"; ([], ln) end in ([node] @ ln, rem) end else ([], []) @@ -179,6 +180,7 @@ let forward_sequences code entry = in [fs] @ ((f code rem_from_node) @ (f code ln)) in (f code [entry]) +(** Unused code module PInt = struct type t = P.t let compare x y = compare (P.to_int x) (P.to_int y) @@ -219,7 +221,10 @@ let can_be_merged code s s' = | Lop _ | Lload _ | Lgetstack _ | Lsetstack _ | Lstore _ | Lcall _ | Lbuiltin _ | Ltailcall _ | Lreturn -> false | Lbranch n -> n == first_s' - | Lcond (_, _, ifso, ifnot, _) -> ifnot == first_s' + | Lcond (_, _, ifso, ifnot, info) -> (match info with + | None -> false + | Some false -> ifnot == first_s' + | Some true -> failwith "Inconsistency detected - ifnot is not the preferred branch") | Ljumptable (_, ln) -> match ln with | [] -> false @@ -256,6 +261,7 @@ let try_merge code (fs: (BinNums.positive list) list) = end done; !seqs +*) (** Code adapted from Duplicateaux.get_loop_headers * -- cgit From dd345b4fd33a9e59507029f30da9a09d5e450db6 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Fri, 13 Mar 2020 14:13:16 +0100 Subject: Added prediction info in the printers --- backend/PrintLTL.ml | 5 +++-- backend/PrintRTL.ml | 5 +++-- 2 files changed, 6 insertions(+), 4 deletions(-) diff --git a/backend/PrintLTL.ml b/backend/PrintLTL.ml index f173e374..d8f2ac12 100644 --- a/backend/PrintLTL.ml +++ b/backend/PrintLTL.ml @@ -83,10 +83,11 @@ let print_instruction pp succ = function (print_builtin_args loc) args | Lbranch s -> print_succ pp s succ - | Lcond(cond, args, s1, s2, _) -> - fprintf pp "if (%a) goto %d else goto %d" + | Lcond(cond, args, s1, s2, info) -> + fprintf pp "if (%a) goto %d else goto %d (prediction: %s)" (print_condition mreg) (cond, args) (P.to_int s1) (P.to_int s2) + (match info with None -> "none" | Some true -> "branch" | Some false -> "fallthrough") | Ljumptable(arg, tbl) -> let tbl = Array.of_list tbl in fprintf pp "jumptable (%a)" mreg arg; diff --git a/backend/PrintRTL.ml b/backend/PrintRTL.ml index 5eab9901..b2ef05ca 100644 --- a/backend/PrintRTL.ml +++ b/backend/PrintRTL.ml @@ -75,10 +75,11 @@ let print_instruction pp (pc, i) = (name_of_external ef) (print_builtin_args reg) args; print_succ pp s (pc - 1) - | Icond(cond, args, s1, s2, _) -> - fprintf pp "if (%a) goto %d else goto %d\n" + | Icond(cond, args, s1, s2, info) -> + fprintf pp "if (%a) goto %d else goto %d (prediction: %s)\n" (PrintOp.print_condition reg) (cond, args) (P.to_int s1) (P.to_int s2) + (match info with None -> "none" | Some true -> "branch" | Some false -> "fallthrough") | Ijumptable(arg, tbl) -> let tbl = Array.of_list tbl in fprintf pp "jumptable (%a)\n" reg arg; -- cgit From afa0407bc6474cbb1b519544ec0386ba7502ca62 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Fri, 13 Mar 2020 14:13:52 +0100 Subject: More helpful debug info in linearize --- backend/Linearizeaux.ml | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/backend/Linearizeaux.ml b/backend/Linearizeaux.ml index 5b3384f2..23ced4c2 100644 --- a/backend/Linearizeaux.ml +++ b/backend/Linearizeaux.ml @@ -168,7 +168,9 @@ let forward_sequences code entry = | Lcond (_, _, ifso, ifnot, info) -> (match info with | None -> begin Printf.printf "STOP Lcond None\n"; ([], [ifso; ifnot]) end | Some false -> let ln, rem = traverse_fallthrough code ifnot in (ln, [ifso] @ rem) - | Some true -> failwith "Inconsistency detected: ifnot is not the preferred branch") + | Some true -> + let errstr = Printf.sprintf ("Inconsistency detected in node %d: ifnot is not the preferred branch") (P.to_int node) in + failwith errstr) | Ljumptable(_, ln) -> begin Printf.printf "STOP Ljumptable\n"; ([], ln) end in ([node] @ ln, rem) end -- cgit From 5d1b51ae43e3b2784805aae60144870a6f14b6e0 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Fri, 13 Mar 2020 14:14:02 +0100 Subject: Fixing bug where conditions were not necessarily inverted --- backend/Duplicateaux.ml | 43 +++++++++++++++++++++---------------------- 1 file changed, 21 insertions(+), 22 deletions(-) diff --git a/backend/Duplicateaux.ml b/backend/Duplicateaux.ml index 37647714..c2839bc3 100644 --- a/backend/Duplicateaux.ml +++ b/backend/Duplicateaux.ml @@ -329,6 +329,18 @@ let best_predecessor_of node predecessors order is_visited = | Some lp -> try Some (List.find (fun n -> (List.mem n lp) && (not (ptree_get_some n is_visited))) order) with Not_found -> None +let print_trace t = print_intlist t + +let print_traces traces = + let rec f = function + | [] -> () + | t::lt -> Printf.printf "\n\t"; print_trace t; Printf.printf ",\n"; f lt + in begin + Printf.printf "Traces: {"; + f traces; + Printf.printf "}\n"; + end + (* Algorithm mostly inspired from Chang and Hwu 1988 * "Trace Selection for Compiling Large C Application Programs to Microcode" *) let select_traces code entrypoint = @@ -369,21 +381,10 @@ let select_traces code entrypoint = end done; (* Printf.printf "DFS: \t"; print_intlist order; Printf.printf "\n"; *) + Printf.printf "Traces: "; print_traces !traces; !traces end -let print_trace t = print_intlist t - -let print_traces traces = - let rec f = function - | [] -> () - | t::lt -> Printf.printf "\n\t"; print_trace t; Printf.printf ",\n"; f lt - in begin - Printf.printf "Traces: {"; - f traces; - Printf.printf "}\n"; - end - let rec make_identity_ptree_rec = function | [] -> PTree.empty | m::lm -> let (n, _) = m in PTree.set n n (make_identity_ptree_rec lm) @@ -487,18 +488,16 @@ let superblockify_traces code preds traces = let rec invert_iconds_trace code = function | [] -> code - | n::[] -> code - | n :: n' :: t -> + | n :: ln -> let code' = match ptree_get_some n code with - | Icond (c, lr, ifso, ifnot, i) -> - assert (n' == ifso || n' == ifnot); - if (n' == ifso) then ( - (* Printf.printf "Reversing ifso/ifnot for node %d\n" (P.to_int n); *) - let i' = match i with None -> None | Some b -> Some (not b) in - PTree.set n (Icond (Op.negate_condition c, lr, ifnot, ifso, i')) code ) - else code + | Icond (c, lr, ifso, ifnot, info) -> (match info with + | Some true -> begin + (* Printf.printf "Reversing ifso/ifnot for node %d\n" (P.to_int n); *) + PTree.set n (Icond (Op.negate_condition c, lr, ifnot, ifso, Some false)) code + end + | _ -> code) | _ -> code - in invert_iconds_trace code' (n'::t) + in invert_iconds_trace code' ln let rec invert_iconds code = function | [] -> code -- cgit From 192d5f379b3f1efa6f12b45af36f7cfea21d6d50 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sun, 15 Mar 2020 00:00:33 +0100 Subject: more inlining --- backend/Inliningaux.ml | 9 ++++++--- driver/Clflags.ml | 1 + driver/Driver.ml | 2 ++ 3 files changed, 9 insertions(+), 3 deletions(-) diff --git a/backend/Inliningaux.ml b/backend/Inliningaux.ml index d58704ca..cf308962 100644 --- a/backend/Inliningaux.ml +++ b/backend/Inliningaux.ml @@ -17,7 +17,8 @@ open Maps open Op open Ordered open! RTL - +open Camlcoq + module PSet = Make(OrderedPositive) type inlining_info = { @@ -83,13 +84,15 @@ let static_called_once id io = else false -(* To be considered: heuristics based on size of function? *) +(* D. Monniaux: attempt at heuristic based on size *) +let small_enough (f : coq_function) = + P.to_int (RTL.max_pc_function f) <= !Clflags.option_inline_auto_threshold let should_inline (io: inlining_info) (id: ident) (f: coq_function) = if !Clflags.option_finline then begin match C2C.atom_inline id with | C2C.Inline -> true | C2C.Noinline -> false - | C2C.No_specifier -> static_called_once id io + | C2C.No_specifier -> static_called_once id io || small_enough f end else false diff --git a/driver/Clflags.ml b/driver/Clflags.ml index 79c0bce0..ee5e9eeb 100644 --- a/driver/Clflags.ml +++ b/driver/Clflags.ml @@ -81,3 +81,4 @@ let option_faddx = ref false let option_fcoalesce_mem = ref true let option_fforward_moves = ref true let option_all_loads_nontrap = ref false +let option_inline_auto_threshold = ref 30 diff --git a/driver/Driver.ml b/driver/Driver.ml index 43aedf50..01451e07 100644 --- a/driver/Driver.ml +++ b/driver/Driver.ml @@ -190,6 +190,7 @@ Processing options: -Os Optimize for code size in preference to code speed -Obranchless Optimize to generate fewer conditional branches; try to produce branch-free instruction sequences as much as possible + -finline-auto-threshold n Inline functions under size n -ftailcalls Optimize function calls in tail position [on] -fconst-prop Perform global constant propagation [on] -ffloat-const-prop Control constant propagation of floats @@ -322,6 +323,7 @@ let cmdline_actions = _Regexp "-O[123]$", Unit (set_all optimization_options); Exact "-Os", Set option_Osize; Exact "-Obranchless", Set option_Obranchless; + Exact "-finline-auto-threshold", Integer (fun n -> option_inline_auto_threshold := n); Exact "-fsmall-data", Integer(fun n -> option_small_data := n); Exact "-fsmall-const", Integer(fun n -> option_small_const := n); Exact "-ffloat-const-prop", Integer(fun n -> option_ffloatconstprop := n); -- cgit From d0326db1105704e02e2b40facc2a85a267a2b9b5 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sun, 15 Mar 2020 09:18:02 +0100 Subject: by default do not inline much --- driver/Clflags.ml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/driver/Clflags.ml b/driver/Clflags.ml index ee5e9eeb..8054eb5b 100644 --- a/driver/Clflags.ml +++ b/driver/Clflags.ml @@ -81,4 +81,4 @@ let option_faddx = ref false let option_fcoalesce_mem = ref true let option_fforward_moves = ref true let option_all_loads_nontrap = ref false -let option_inline_auto_threshold = ref 30 +let option_inline_auto_threshold = ref 0 -- cgit From 1b111e3658b3f79a9814fd9799e2dbe0a921c768 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 17 Mar 2020 00:11:42 +0100 Subject: specify prefix with CCOMP_INSTALL_PREFIX --- config_simple.sh | 7 ++++++- 1 file changed, 6 insertions(+), 1 deletion(-) diff --git a/config_simple.sh b/config_simple.sh index f02680c4..e2d3844c 100755 --- a/config_simple.sh +++ b/config_simple.sh @@ -3,4 +3,9 @@ shift version=`git rev-parse --short HEAD` branch=`git rev-parse --abbrev-ref HEAD` date=`date -I` -./configure --prefix /opt/CompCert/${branch}/${date}_${version}/$arch "$@" $arch + +if test "x$CCOMP_INSTALL_PREFIX" = "x" ; +then CCOMP_INSTALL_PREFIX=/opt/CompCert ; +fi + +./configure --prefix ${CCOMP_INSTALL_PREFIX}/${branch}/${date}_${version}/$arch "$@" $arch -- cgit From fac63511aaf3fd3c77db69802c24474559365879 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 17 Mar 2020 07:44:45 +0100 Subject: fix for aarch64 DuplicateOpcodeHeuristic.ml --- aarch64/DuplicateOpcodeHeuristic.ml | 30 +++++++++++++++++++++++++++--- 1 file changed, 27 insertions(+), 3 deletions(-) diff --git a/aarch64/DuplicateOpcodeHeuristic.ml b/aarch64/DuplicateOpcodeHeuristic.ml index 85505245..5fc2156c 100644 --- a/aarch64/DuplicateOpcodeHeuristic.ml +++ b/aarch64/DuplicateOpcodeHeuristic.ml @@ -1,3 +1,27 @@ -exception HeuristicSucceeded - -let opcode_heuristic code cond ifso ifnot preferred = () +open Op +open Integers + +let opcode_heuristic code cond ifso ifnot is_loop_header = + match cond with + | Ccompimm (c, n) | Ccompuimm (c, n) -> if n == Integers.Int.zero then (match c with + | Clt | Cle -> Some false + | Cgt | Cge -> Some true + | _ -> None + ) else None + | Ccomplimm (c, n) | Ccompluimm (c, n) -> if n == Integers.Int64.zero then (match c with + | Clt | Cle -> Some false + | Cgt | Cge -> Some true + | _ -> None + ) else None + | Ccompf c | Ccompfs c -> (match c with + | Ceq -> Some false + | Cne -> Some true + | _ -> None + ) + | Cnotcompf c | Cnotcompfs c -> (match c with + | Ceq -> Some true + | Cne -> Some false + | _ -> None + ) + | _ -> None + -- cgit From 0ebdbc31c3e992e43d85699a039ebdd23e272df6 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 17 Mar 2020 07:49:46 +0100 Subject: DuplicateOpcodeHeuristic for ARM --- arm/DuplicateOpcodeHeuristic.ml | 25 ++++++++++++++++++++++--- 1 file changed, 22 insertions(+), 3 deletions(-) diff --git a/arm/DuplicateOpcodeHeuristic.ml b/arm/DuplicateOpcodeHeuristic.ml index 85505245..9b6a6409 100644 --- a/arm/DuplicateOpcodeHeuristic.ml +++ b/arm/DuplicateOpcodeHeuristic.ml @@ -1,3 +1,22 @@ -exception HeuristicSucceeded - -let opcode_heuristic code cond ifso ifnot preferred = () +open Op +open Integers + +let opcode_heuristic code cond ifso ifnot is_loop_header = + match cond with + | Ccompimm (c, n) | Ccompuimm (c, n) -> if n == Integers.Int.zero then (match c with + | Clt | Cle -> Some false + | Cgt | Cge -> Some true + | _ -> None + ) else None + | Ccompf c | Ccompfs c -> (match c with + | Ceq -> Some false + | Cne -> Some true + | _ -> None + ) + | Cnotcompf c | Cnotcompfs c -> (match c with + | Ceq -> Some true + | Cne -> Some false + | _ -> None + ) + | _ -> None + -- cgit From 2f19071e865181d9a0c2e61f5e57731fb86e4e5d Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 17 Mar 2020 07:57:56 +0100 Subject: riscV/DuplicateOpcodeHeuristic.ml --- config_rv64.sh | 2 +- riscV/DuplicateOpcodeHeuristic.ml | 30 +++++++++++++++++++++++++++--- 2 files changed, 28 insertions(+), 4 deletions(-) diff --git a/config_rv64.sh b/config_rv64.sh index e95f8a70..0698c2ff 100755 --- a/config_rv64.sh +++ b/config_rv64.sh @@ -1 +1 @@ -exec ./config_simple.sh rv64-linux --toolprefix riscv64-unknown-elf- "$@" +exec ./config_simple.sh rv64-linux --toolprefix riscv64-linux-gnu- "$@" diff --git a/riscV/DuplicateOpcodeHeuristic.ml b/riscV/DuplicateOpcodeHeuristic.ml index 85505245..2ec314c1 100644 --- a/riscV/DuplicateOpcodeHeuristic.ml +++ b/riscV/DuplicateOpcodeHeuristic.ml @@ -1,3 +1,27 @@ -exception HeuristicSucceeded - -let opcode_heuristic code cond ifso ifnot preferred = () +(* open Camlcoq *) +open Op +open Integers + +let opcode_heuristic code cond ifso ifnot is_loop_header = + match cond with + | Ccompimm (c, n) | Ccompuimm (c, n) -> if n == Integers.Int.zero then (match c with + | Clt | Cle -> Some false + | Cgt | Cge -> Some true + | _ -> None + ) else None + | Ccomplimm (c, n) | Ccompluimm (c, n) -> if n == Integers.Int64.zero then (match c with + | Clt | Cle -> Some false + | Cgt | Cge -> Some true + | _ -> None + ) else None + | Ccompf c | Ccompfs c -> (match c with + | Ceq -> Some false + | Cne -> Some true + | _ -> None + ) + | Cnotcompf c | Cnotcompfs c -> (match c with + | Ceq -> Some true + | Cne -> Some false + | _ -> None + ) + | _ -> None -- cgit From 0eb778a85b5b76ab6c7fd914ffaff1affcbde7bb Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 17 Mar 2020 08:05:17 +0100 Subject: DuplicateOpcodeHeuristic ppc --- powerpc/DuplicateOpcodeHeuristic.ml | 30 +++++++++++++++++++++++++++--- 1 file changed, 27 insertions(+), 3 deletions(-) diff --git a/powerpc/DuplicateOpcodeHeuristic.ml b/powerpc/DuplicateOpcodeHeuristic.ml index 85505245..33be79e8 100644 --- a/powerpc/DuplicateOpcodeHeuristic.ml +++ b/powerpc/DuplicateOpcodeHeuristic.ml @@ -1,3 +1,27 @@ -exception HeuristicSucceeded - -let opcode_heuristic code cond ifso ifnot preferred = () +(* open Camlcoq *) +open Op +open Integers + +let opcode_heuristic code cond ifso ifnot is_loop_header = + match cond with + | Ccompimm (c, n) | Ccompuimm (c, n) -> if n == Integers.Int.zero then (match c with + | Clt | Cle -> Some false + | Cgt | Cge -> Some true + | _ -> None + ) else None + | Ccomplimm (c, n) | Ccompluimm (c, n) -> if n == Integers.Int64.zero then (match c with + | Clt | Cle -> Some false + | Cgt | Cge -> Some true + | _ -> None + ) else None + | Ccompf c -> (match c with + | Ceq -> Some false + | Cne -> Some true + | _ -> None + ) + | Cnotcompf c -> (match c with + | Ceq -> Some true + | Cne -> Some false + | _ -> None + ) + | _ -> None -- cgit From aa2e3d776cb82ce01c4afdbacc52951e60ff2104 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 17 Mar 2020 08:10:26 +0100 Subject: DuplicateOpcodeHeuristic x86 --- x86/DuplicateOpcodeHeuristic.ml | 30 +++++++++++++++++++++++++++--- 1 file changed, 27 insertions(+), 3 deletions(-) diff --git a/x86/DuplicateOpcodeHeuristic.ml b/x86/DuplicateOpcodeHeuristic.ml index 85505245..2ec314c1 100644 --- a/x86/DuplicateOpcodeHeuristic.ml +++ b/x86/DuplicateOpcodeHeuristic.ml @@ -1,3 +1,27 @@ -exception HeuristicSucceeded - -let opcode_heuristic code cond ifso ifnot preferred = () +(* open Camlcoq *) +open Op +open Integers + +let opcode_heuristic code cond ifso ifnot is_loop_header = + match cond with + | Ccompimm (c, n) | Ccompuimm (c, n) -> if n == Integers.Int.zero then (match c with + | Clt | Cle -> Some false + | Cgt | Cge -> Some true + | _ -> None + ) else None + | Ccomplimm (c, n) | Ccompluimm (c, n) -> if n == Integers.Int64.zero then (match c with + | Clt | Cle -> Some false + | Cgt | Cge -> Some true + | _ -> None + ) else None + | Ccompf c | Ccompfs c -> (match c with + | Ceq -> Some false + | Cne -> Some true + | _ -> None + ) + | Cnotcompf c | Cnotcompfs c -> (match c with + | Ceq -> Some true + | Cne -> Some false + | _ -> None + ) + | _ -> None -- cgit From fb43d1078c0b0824132b30d7dd9bfe6b0ac47122 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Tue, 17 Mar 2020 15:12:06 +0100 Subject: Desactivating branch predictions by default --- backend/Duplicateaux.ml | 18 +++++++++++------- driver/Clflags.ml | 2 +- driver/Driver.ml | 6 +++++- 3 files changed, 17 insertions(+), 9 deletions(-) diff --git a/backend/Duplicateaux.ml b/backend/Duplicateaux.ml index 209527b9..a84f9754 100644 --- a/backend/Duplicateaux.ml +++ b/backend/Duplicateaux.ml @@ -502,11 +502,15 @@ let rec invert_iconds code = function let duplicate_aux f = let entrypoint = f.fn_entrypoint in let code = f.fn_code in - let traces = select_traces (to_ttl_code code entrypoint) entrypoint in - let icond_code = invert_iconds code traces in - let preds = get_predecessors_rtl icond_code in - if !Clflags.option_fduplicate >= 1 then - let (new_code, pTreeId) = ((* print_traces traces; *) superblockify_traces icond_code preds traces) in - ((new_code, f.fn_entrypoint), pTreeId) + if !Clflags.option_fduplicate < 0 then + ((code, entrypoint), make_identity_ptree code) else - ((icond_code, entrypoint), make_identity_ptree code) + let traces = select_traces (to_ttl_code code entrypoint) entrypoint in + let icond_code = invert_iconds code traces in + let preds = get_predecessors_rtl icond_code in + if !Clflags.option_fduplicate >= 1 then + let (new_code, pTreeId) = ((* print_traces traces; *) superblockify_traces icond_code preds traces) in + ((new_code, f.fn_entrypoint), pTreeId) + else + ((icond_code, entrypoint), make_identity_ptree code) + diff --git a/driver/Clflags.ml b/driver/Clflags.ml index 8054eb5b..6986fb96 100644 --- a/driver/Clflags.ml +++ b/driver/Clflags.ml @@ -28,7 +28,7 @@ let option_fconstprop = ref true let option_fcse = ref true let option_fcse2 = ref true let option_fredundancy = ref true -let option_fduplicate = ref 0 +let option_fduplicate = ref (-1) let option_finvertcond = ref true let option_ftracelinearize = ref false let option_fpostpass = ref true diff --git a/driver/Driver.ml b/driver/Driver.ml index 01451e07..388482a0 100644 --- a/driver/Driver.ml +++ b/driver/Driver.ml @@ -201,7 +201,11 @@ Processing options: -fpostpass Perform postpass scheduling (only for K1 architecture) [on] -fpostpass= Perform postpass scheduling with the specified optimization [list] (=list: list scheduling, =ilp: ILP, =greedy: just packing bundles) - -fduplicate Perform tail duplication to form superblocks on predicted traces + -fduplicate Perform tail duplication to form superblocks on predicted traces + nb_nodes control the heuristic deciding to duplicate or not + A value of -1 desactivates the entire pass (including branch prediction) + A value of 0 desactivates the duplication (but activates the branch prediction) + FIXME : this is desactivated by default for now -finvertcond Invert conditions based on predicted paths (to prefer fallthrough). Requires -fduplicate to be also activated [on] -ftracelinearize Linearizes based on the traces identified by duplicate phase -- cgit From 75be7f9206e495966fdc8c409007a2bcadedddbe Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 18 Mar 2020 15:28:48 +0100 Subject: Adding trace printf at start of some functions --- backend/Duplicateaux.ml | 20 +++++++++++++++----- 1 file changed, 15 insertions(+), 5 deletions(-) diff --git a/backend/Duplicateaux.ml b/backend/Duplicateaux.ml index 1548765e..99a2d9e3 100644 --- a/backend/Duplicateaux.ml +++ b/backend/Duplicateaux.ml @@ -6,7 +6,8 @@ let get_some = function | None -> failwith "Did not get some" | Some thing -> thing -let bfs code entrypoint = +let bfs code entrypoint = begin + Printf.printf "bfs\n"; flush stdout; let visited = ref (PTree.map (fun n i -> false) code) and bfs_list = ref [] and to_visit = Queue.create () @@ -32,6 +33,7 @@ let bfs code entrypoint = done; !bfs_list end +end let optbool o = match o with Some _ -> true | None -> false @@ -91,7 +93,8 @@ type vstate = Unvisited | Processed | Visited * * If we come accross an edge to a Processed node, it's a loop! *) -let get_loop_headers code entrypoint = +let get_loop_headers code entrypoint = begin + Printf.printf "get_loop_headers\n"; flush stdout; let visited = ref (PTree.map (fun n i -> Unvisited) code) and is_loop_header = ref (PTree.map (fun n i -> false) code) in let rec dfs_visit code = function @@ -121,6 +124,7 @@ let get_loop_headers code entrypoint = dfs_visit code [entrypoint]; !is_loop_header end +end let ptree_printbool pt = let elements = PTree.elements pt @@ -195,7 +199,8 @@ let do_loop_heuristic code cond ifso ifnot is_loop_header = else None end -let get_directions code entrypoint = +let get_directions code entrypoint = begin + Printf.printf "get_directions\n"; flush stdout; let bfs_order = bfs code entrypoint and is_loop_header = get_loop_headers code entrypoint and directions = ref (PTree.map (fun n i -> None) code) (* None <=> no predicted direction *) @@ -226,6 +231,7 @@ let get_directions code entrypoint = ) bfs_order; !directions end +end let update_direction direction = function | Icond (cond, lr, n, n', _) -> Icond (cond, lr, n, n', direction) @@ -238,7 +244,8 @@ let rec update_direction_rec directions = function in PTree.set n (update_direction direction i) (update_direction_rec directions lm) (* Uses branch prediction to write prediction annotations in Icond *) -let update_directions code entrypoint = +let update_directions code entrypoint = begin + Printf.printf "Update_directions\n"; flush stdout; let directions = get_directions code entrypoint in begin (* Printf.printf "Ifso directions: "; @@ -246,6 +253,7 @@ let update_directions code entrypoint = Printf.printf "\n"; *) update_direction_rec directions (PTree.elements code) end +end (** Trace selection *) @@ -343,7 +351,8 @@ let print_traces traces = (* Algorithm mostly inspired from Chang and Hwu 1988 * "Trace Selection for Compiling Large C Application Programs to Microcode" *) -let select_traces code entrypoint = +let select_traces code entrypoint = begin + Printf.printf "select_traces\n"; flush stdout; let order = dfs code entrypoint in let predecessors = get_predecessors_rtl code in let traces = ref [] in @@ -384,6 +393,7 @@ let select_traces code entrypoint = Printf.printf "Traces: "; print_traces !traces; !traces end +end let rec make_identity_ptree_rec = function | [] -> PTree.empty -- cgit From d4002956b3fbe9085e685e0c596f776ecfcdafd7 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 18 Mar 2020 15:29:09 +0100 Subject: Correcting inefficiency in Duplicateaux::bfs --- backend/Duplicateaux.ml | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/backend/Duplicateaux.ml b/backend/Duplicateaux.ml index 99a2d9e3..3ffe9aed 100644 --- a/backend/Duplicateaux.ml +++ b/backend/Duplicateaux.ml @@ -21,7 +21,7 @@ let bfs code entrypoint = begin match PTree.get !node code with | None -> failwith "No such node" | Some i -> - bfs_list := !bfs_list @ [!node]; + bfs_list := !node :: !bfs_list; match i with | Icall(_, _, _, _, n) -> Queue.add n to_visit | Ibuiltin(_, _, _, n) -> Queue.add n to_visit @@ -31,7 +31,7 @@ let bfs code entrypoint = begin | Inop n | Iop (_, _, _, n) | Iload (_, _, _, _, _, n) | Istore (_, _, _, _, n) -> Queue.add n to_visit end done; - !bfs_list + List.rev !bfs_list end end -- cgit From c6f8888aa89cfa86a9d61ecdc8d030cc8710ab6d Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 18 Mar 2020 15:46:19 +0100 Subject: Fixing inefficient implementation of Duplicateaux.dfs --- backend/Duplicateaux.ml | 42 ++++++++++++++++++++---------------------- 1 file changed, 20 insertions(+), 22 deletions(-) diff --git a/backend/Duplicateaux.ml b/backend/Duplicateaux.ml index 3ffe9aed..ae0c6252 100644 --- a/backend/Duplicateaux.ml +++ b/backend/Duplicateaux.ml @@ -39,7 +39,8 @@ let optbool o = match o with Some _ -> true | None -> false let ptree_get_some n ptree = get_some @@ PTree.get n ptree -let get_predecessors_rtl code = +let get_predecessors_rtl code = begin + Printf.printf "get_predecessors_rtl\n"; flush stdout; let preds = ref (PTree.map (fun n i -> []) code) in let process_inst (node, i) = let succ = match i with @@ -56,6 +57,7 @@ let get_predecessors_rtl code = List.iter process_inst (PTree.elements code); !preds end +end module PInt = struct type t = P.t @@ -264,32 +266,28 @@ let rec exists_false_rec = function let exists_false boolmap = exists_false_rec (PTree.elements boolmap) (* DFS using prediction info to guide the exploration *) -let dfs code entrypoint = +let dfs code entrypoint = begin + Printf.printf "dfs\n"; flush stdout; let visited = ref (PTree.map (fun n i -> false) code) in let rec dfs_list code = function | [] -> [] | node :: ln -> - let node_dfs = - if not (get_some @@ PTree.get node !visited) then begin - visited := PTree.set node true !visited; - match PTree.get node code with - | None -> failwith "No such node" - | Some i -> [node] @ match i with - | Icall(_, _, _, _, n) -> dfs_list code [n] - | Ibuiltin(_, _, _, n) -> dfs_list code [n] - | Iop (_, _, _, n) -> dfs_list code [n] - | Iload (_, _, _, _, _, n) -> dfs_list code [n] - | Istore (_, _, _, _, n) -> dfs_list code [n] - | Inop n -> dfs_list code [n] - | Ijumptable(_, ln) -> dfs_list code ln - | Itailcall _ | Ireturn _ -> [] - | Icond (_, _, n1, n2, info) -> match info with - | Some false -> dfs_list code [n2; n1] - | _ -> dfs_list code [n1; n2] - end - else [] - in node_dfs @ (dfs_list code ln) + if get_some @@ PTree.get node !visited then dfs_list code ln + else begin + visited := PTree.set node true !visited; + let next_nodes = (match get_some @@ PTree.get node code with + | Icall(_, _, _, _, n) | Ibuiltin (_, _, _, n) | Iop (_, _, _, n) + | Iload (_, _, _, _, _, n) | Istore (_, _, _, _, n) | Inop n -> [n] + | Ijumptable (_, ln) -> ln + | Itailcall _ | Ireturn _ -> [] + | Icond (_, _, n1, n2, info) -> (match info with + | Some false -> [n2; n1] + | _ -> [n1; n2] + ) + ) in node :: dfs_list code (next_nodes @ ln) + end in dfs_list code [entrypoint] +end (* let get_predecessors_ttl code = -- cgit From 66f96f7b3f84bf011be40b672e864c5c0f913f02 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Thu, 19 Mar 2020 12:02:52 +0100 Subject: New algo for Duplicateaux.select_traces in O(n) --- backend/Duplicateaux.ml | 75 ++++++++++++++++++++----------------------------- 1 file changed, 30 insertions(+), 45 deletions(-) diff --git a/backend/Duplicateaux.ml b/backend/Duplicateaux.ml index ae0c6252..91d313f7 100644 --- a/backend/Duplicateaux.ml +++ b/backend/Duplicateaux.ml @@ -52,9 +52,11 @@ let get_predecessors_rtl code = begin in List.iter (fun s -> let previous_preds = ptree_get_some s !preds in if optbool @@ List.find_opt (fun e -> e == node) previous_preds then () - else preds := PTree.set s (node::previous_preds) !preds) succ + else preds := PTree.set s (node::previous_preds) !preds) + succ in begin List.iter process_inst (PTree.elements code); + Printf.printf "get_predecessors_rtl done\n"; flush stdout; !preds end end @@ -329,10 +331,10 @@ let best_successor_of node code is_visited = | Some n -> if not (ptree_get_some n is_visited) then Some n else None (* FIXME - could be improved by selecting in priority the predicted paths *) -let best_predecessor_of node predecessors order is_visited = +let best_predecessor_of node predecessors order = match (PTree.get node predecessors) with | None -> failwith "No predecessor list found" - | Some lp -> try Some (List.find (fun n -> (List.mem n lp) && (not (ptree_get_some n is_visited))) order) + | Some lp -> try Some (List.find (fun n -> List.mem n lp) order) with Not_found -> None let print_trace t = print_intlist t @@ -347,51 +349,34 @@ let print_traces traces = Printf.printf "}\n"; end -(* Algorithm mostly inspired from Chang and Hwu 1988 - * "Trace Selection for Compiling Large C Application Programs to Microcode" *) -let select_traces code entrypoint = begin - Printf.printf "select_traces\n"; flush stdout; - let order = dfs code entrypoint in - let predecessors = get_predecessors_rtl code in - let traces = ref [] in - let is_visited = ref (PTree.map (fun n i -> false) code) in begin (* mark all nodes visited *) - while exists_false !is_visited do (* while (there are unvisited nodes) *) - let seed = select_unvisited_node !is_visited order in - let trace = ref [seed] in - let current = ref seed in begin - is_visited := PTree.set seed true !is_visited; (* mark seed visited *) - let quit_loop = ref false in begin - while not !quit_loop do - let s = best_successor_of !current code !is_visited in - match s with - | None -> quit_loop := true (* if (s==0) exit loop *) - | Some succ -> begin - trace := !trace @ [succ]; - is_visited := PTree.set succ true !is_visited; (* mark s visited *) - current := succ - end - done; - current := seed; - quit_loop := false; - while not !quit_loop do - let s = best_predecessor_of !current predecessors order !is_visited in - match s with - | None -> quit_loop := true (* if (s==0) exit loop *) - | Some pred -> begin - trace := pred :: !trace; - is_visited := PTree.set pred true !is_visited; (* mark s visited *) - current := pred - end - done; - traces := !trace :: !traces; - end - end - done; - (* Printf.printf "DFS: \t"; print_intlist order; Printf.printf "\n"; *) +let select_traces code entrypoint = + let is_visited = ref (PTree.map (fun n i -> false) code) in + let bfs_order = bfs code entrypoint in + let rec go_through node = begin + is_visited := PTree.set node true !is_visited; + let next_node = match (get_some @@ PTree.get node code) with + | Icall(_, _, _, _, n) | Ibuiltin (_, _, _, n) | Iop (_, _, _, n) + | Iload (_, _, _, _, _, n) | Istore (_, _, _, _, n) | Inop n -> Some n + | Ijumptable _ | Itailcall _ | Ireturn _ -> None + | Icond (_, _, n1, n2, info) -> (match info with + | Some false -> Some n2 + | Some true -> Some n1 + | None -> None + ) + in match next_node with + | None -> [node] + | Some n -> + if not (get_some @@ PTree.get n !is_visited) then node :: go_through n + else [node] + end + in let traces = ref [] in begin + List.iter (fun n -> + if not (get_some @@ PTree.get n !is_visited) then + traces := (go_through n) :: !traces + ) bfs_order; Printf.printf "Traces: "; print_traces !traces; !traces end -end let rec make_identity_ptree_rec = function | [] -> PTree.empty -- cgit From 88d3af62f4d0ed5f400c1e8690343a7a9ad15fe3 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 20 Mar 2020 05:59:01 +0100 Subject: progress in RA invariants --- mppa_k1c/Asmblockgenproof1.v | 47 ++++++++++++++++++++++---------------------- 1 file changed, 24 insertions(+), 23 deletions(-) diff --git a/mppa_k1c/Asmblockgenproof1.v b/mppa_k1c/Asmblockgenproof1.v index 5b44ddaa..6a3f2389 100644 --- a/mppa_k1c/Asmblockgenproof1.v +++ b/mppa_k1c/Asmblockgenproof1.v @@ -1517,21 +1517,21 @@ Opaque Int.eq. - (* Ocast8signed *) econstructor; split. eapply exec_straight_two. simpl;eauto. simpl;eauto. - split; intros; simpl; Simpl. + repeat split; intros; simpl; Simpl. assert (A: Int.ltu (Int.repr 24) Int.iwordsize = true) by auto. destruct (rs x0); auto; simpl. rewrite A; simpl. Simpl. unfold Val.shr. rewrite A. apply Val.lessdef_same. f_equal. apply Int.sign_ext_shr_shl. split; reflexivity. - (* Ocast16signed *) econstructor; split. eapply exec_straight_two. simpl;eauto. simpl;eauto. - split; intros; Simpl. + repeat split; intros; Simpl. assert (A: Int.ltu (Int.repr 16) Int.iwordsize = true) by auto. destruct (rs x0); auto; simpl. rewrite A; simpl. Simpl. unfold Val.shr. rewrite A. apply Val.lessdef_same. f_equal. apply Int.sign_ext_shr_shl. split; reflexivity. - (* Oshrximm *) econstructor; split. + apply exec_straight_one. simpl. eauto. - + split. + + repeat split. * rewrite Pregmap.gss. subst v. destruct (rs x0); simpl; trivial. @@ -1542,7 +1542,7 @@ Opaque Int.eq. - (* Oshrxlimm *) econstructor; split. + apply exec_straight_one. simpl. eauto. - + split. + + repeat split. * rewrite Pregmap.gss. subst v. destruct (rs x0); simpl; trivial. @@ -1553,7 +1553,7 @@ Opaque Int.eq. - (* Ocmp *) exploit transl_cond_op_correct; eauto. intros (rs' & A & B & C). - exists rs'; split. eexact A. eauto with asmgen. + exists rs'; repeat split; eauto with asmgen. - (* Osel *) unfold conditional_move in *. @@ -1572,24 +1572,25 @@ Opaque Int.eq. destruct c0; simpl in *. - all: - destruct c; simpl in *; inv EQ2; - econstructor; split; try (apply exec_straight_one; constructor); - split; try (simpl; intros; rewrite Pregmap.gso; trivial; assumption); - unfold Val.select; simpl; - unfold cmove, cmoveu; - rewrite Pregmap.gss; - destruct (rs x1); simpl; trivial; - try rewrite int_ltu_to_neq; - try rewrite int64_ltu_to_neq; - try change (Int64.eq Int64.zero Int64.zero) with true; - try destruct Archi.ptr64; - repeat rewrite if_neg; - simpl; - trivial; - try destruct (_ || _); - trivial; - try apply Val.lessdef_normalize. + all: destruct c. + all: simpl in *. + all: inv EQ2. + all: try (econstructor; split; [idtac | split ]). + all: try apply exec_straight_one. + all: intros; simpl; trivial. + all: unfold Val.select, cmove, cmoveu; simpl. + all: destruct (rs x1); simpl; trivial. + all: try rewrite int_ltu_to_neq. + all: try rewrite int64_ltu_to_neq. + all: try change (Int64.eq Int64.zero Int64.zero) with true. + all: try destruct Archi.ptr64. + all: try rewrite Pregmap.gss. + all: repeat rewrite if_neg. + all: simpl. + all: try destruct (_ || _). + all: try apply Val.lessdef_normalize. + all: trivial. (* no more lessdef *) + all: apply Pregmap.gso; congruence. - (* Oselimm *) unfold conditional_move_imm32 in *. -- cgit From 5e05d4acc53b4b098cb55006b8daa32149d7fba4 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 20 Mar 2020 06:16:12 +0100 Subject: more understandabe proofs --- mppa_k1c/Asmblockgenproof1.v | 76 ++++++++++++++++++++++---------------------- 1 file changed, 38 insertions(+), 38 deletions(-) diff --git a/mppa_k1c/Asmblockgenproof1.v b/mppa_k1c/Asmblockgenproof1.v index 6a3f2389..4c29867b 100644 --- a/mppa_k1c/Asmblockgenproof1.v +++ b/mppa_k1c/Asmblockgenproof1.v @@ -1596,49 +1596,49 @@ Opaque Int.eq. unfold conditional_move_imm32 in *. destruct c0; simpl in *. - all: - destruct c; simpl in *; inv EQ0; - econstructor; split; try (apply exec_straight_one; constructor); - split; try (simpl; intros; rewrite Pregmap.gso; trivial; assumption); - unfold Val.select; simpl; - unfold cmove, cmoveu; - rewrite Pregmap.gss; - destruct (rs x0); simpl; trivial; - try rewrite int_ltu_to_neq; - try rewrite int64_ltu_to_neq; - try change (Int64.eq Int64.zero Int64.zero) with true; - try destruct Archi.ptr64; - repeat rewrite if_neg; - simpl; - trivial; - try destruct (_ || _); - trivial; - try apply Val.lessdef_normalize. - + all: destruct c. + all: simpl in *. + all: inv EQ0. + all: try (econstructor; split; [idtac | split ]). + all: try apply exec_straight_one. + all: intros; simpl; trivial. + all: unfold Val.select, cmove, cmoveu; simpl. + all: destruct (rs x0); simpl; trivial. + all: try rewrite int_ltu_to_neq. + all: try rewrite int64_ltu_to_neq. + all: try change (Int64.eq Int64.zero Int64.zero) with true. + all: try destruct Archi.ptr64. + all: try rewrite Pregmap.gss. + all: repeat rewrite if_neg. + all: simpl. + all: try destruct (_ || _). + all: try apply Val.lessdef_normalize. + all: trivial. (* no more lessdef *) + all: apply Pregmap.gso; congruence. - (* Osellimm *) unfold conditional_move_imm64 in *. destruct c0; simpl in *. - all: - destruct c; simpl in *; inv EQ0; - econstructor; split; try (apply exec_straight_one; constructor); - split; try (simpl; intros; rewrite Pregmap.gso; trivial; assumption); - unfold Val.select; simpl; - unfold cmove, cmoveu; - rewrite Pregmap.gss; - destruct (rs x0); simpl; trivial; - try rewrite int_ltu_to_neq; - try rewrite int64_ltu_to_neq; - try change (Int64.eq Int64.zero Int64.zero) with true; - try destruct Archi.ptr64; - repeat rewrite if_neg; - simpl; - trivial; - try destruct (_ || _); - trivial; - try apply Val.lessdef_normalize. - + all: destruct c. + all: simpl in *. + all: inv EQ0. + all: try (econstructor; split; [idtac | split ]). + all: try apply exec_straight_one. + all: intros; simpl; trivial. + all: unfold Val.select, cmove, cmoveu; simpl. + all: destruct (rs x0); simpl; trivial. + all: try rewrite int_ltu_to_neq. + all: try rewrite int64_ltu_to_neq. + all: try change (Int64.eq Int64.zero Int64.zero) with true. + all: try destruct Archi.ptr64. + all: try rewrite Pregmap.gss. + all: repeat rewrite if_neg. + all: simpl. + all: try destruct (_ || _). + all: try apply Val.lessdef_normalize. + all: trivial. (* no more lessdef *) + all: apply Pregmap.gso; congruence. Qed. (** Memory accesses *) -- cgit From b1b2c6c6442a48c8eb2f7f378e440d8d4311048f Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 20 Mar 2020 06:35:43 +0100 Subject: proof clarification --- mppa_k1c/Asmblockgenproof1.v | 8 +++++--- 1 file changed, 5 insertions(+), 3 deletions(-) diff --git a/mppa_k1c/Asmblockgenproof1.v b/mppa_k1c/Asmblockgenproof1.v index 4c29867b..00df01e3 100644 --- a/mppa_k1c/Asmblockgenproof1.v +++ b/mppa_k1c/Asmblockgenproof1.v @@ -1483,6 +1483,8 @@ Proof. destruct (Z_lt_dec _ _); destruct (Z.eq_dec _ _); trivial; omega. Qed. +Ltac splitall := repeat match goal with |- _ /\ _ => split end. + Lemma transl_op_correct: forall op args res k (rs: regset) m v c, transl_op op args res k = OK c -> @@ -1575,7 +1577,7 @@ Opaque Int.eq. all: destruct c. all: simpl in *. all: inv EQ2. - all: try (econstructor; split; [idtac | split ]). + all: econstructor; splitall. all: try apply exec_straight_one. all: intros; simpl; trivial. all: unfold Val.select, cmove, cmoveu; simpl. @@ -1599,7 +1601,7 @@ Opaque Int.eq. all: destruct c. all: simpl in *. all: inv EQ0. - all: try (econstructor; split; [idtac | split ]). + all: econstructor; splitall. all: try apply exec_straight_one. all: intros; simpl; trivial. all: unfold Val.select, cmove, cmoveu; simpl. @@ -1623,7 +1625,7 @@ Opaque Int.eq. all: destruct c. all: simpl in *. all: inv EQ0. - all: try (econstructor; split; [idtac | split ]). + all: econstructor; splitall. all: try apply exec_straight_one. all: intros; simpl; trivial. all: unfold Val.select, cmove, cmoveu; simpl. -- cgit From 87b17fa1912da24ba114a181d1fbd1779d33e835 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Fri, 20 Mar 2020 15:56:00 +0100 Subject: Reintroducing the Chang algorithm - selecting algo based on size --- backend/Duplicateaux.ml | 63 ++++++++++++++++++++++++++++++++++++++++++++----- 1 file changed, 57 insertions(+), 6 deletions(-) diff --git a/backend/Duplicateaux.ml b/backend/Duplicateaux.ml index 91d313f7..b9f5cdf2 100644 --- a/backend/Duplicateaux.ml +++ b/backend/Duplicateaux.ml @@ -52,11 +52,9 @@ let get_predecessors_rtl code = begin in List.iter (fun s -> let previous_preds = ptree_get_some s !preds in if optbool @@ List.find_opt (fun e -> e == node) previous_preds then () - else preds := PTree.set s (node::previous_preds) !preds) - succ + else preds := PTree.set s (node::previous_preds) !preds) succ in begin List.iter process_inst (PTree.elements code); - Printf.printf "get_predecessors_rtl done\n"; flush stdout; !preds end end @@ -331,10 +329,10 @@ let best_successor_of node code is_visited = | Some n -> if not (ptree_get_some n is_visited) then Some n else None (* FIXME - could be improved by selecting in priority the predicted paths *) -let best_predecessor_of node predecessors order = +let best_predecessor_of node predecessors order is_visited = match (PTree.get node predecessors) with | None -> failwith "No predecessor list found" - | Some lp -> try Some (List.find (fun n -> List.mem n lp) order) + | Some lp -> try Some (List.find (fun n -> (List.mem n lp) && (not (ptree_get_some n is_visited))) order) with Not_found -> None let print_trace t = print_intlist t @@ -349,7 +347,8 @@ let print_traces traces = Printf.printf "}\n"; end -let select_traces code entrypoint = +(* Dumb (but linear) trace selection *) +let select_traces_linear code entrypoint = let is_visited = ref (PTree.map (fun n i -> false) code) in let bfs_order = bfs code entrypoint in let rec go_through node = begin @@ -374,9 +373,61 @@ let select_traces code entrypoint = if not (get_some @@ PTree.get n !is_visited) then traces := (go_through n) :: !traces ) bfs_order; + !traces + end + + +(* Algorithm mostly inspired from Chang and Hwu 1988 + * "Trace Selection for Compiling Large C Application Programs to Microcode" *) +let select_traces_chang code entrypoint = begin + Printf.printf "select_traces\n"; flush stdout; + let order = dfs code entrypoint in + let predecessors = get_predecessors_rtl code in + let traces = ref [] in + let is_visited = ref (PTree.map (fun n i -> false) code) in begin (* mark all nodes visited *) + Printf.printf "Length: %d\n" (List.length order); flush stdout; + while exists_false !is_visited do (* while (there are unvisited nodes) *) + let seed = select_unvisited_node !is_visited order in + let trace = ref [seed] in + let current = ref seed in begin + is_visited := PTree.set seed true !is_visited; (* mark seed visited *) + let quit_loop = ref false in begin + while not !quit_loop do + let s = best_successor_of !current code !is_visited in + match s with + | None -> quit_loop := true (* if (s==0) exit loop *) + | Some succ -> begin + trace := !trace @ [succ]; + is_visited := PTree.set succ true !is_visited; (* mark s visited *) + current := succ + end + done; + current := seed; + quit_loop := false; + while not !quit_loop do + let s = best_predecessor_of !current predecessors order !is_visited in + match s with + | None -> quit_loop := true (* if (s==0) exit loop *) + | Some pred -> begin + trace := pred :: !trace; + is_visited := PTree.set pred true !is_visited; (* mark s visited *) + current := pred + end + done; + traces := !trace :: !traces; + end + end + done; + (* Printf.printf "DFS: \t"; print_intlist order; Printf.printf "\n"; *) Printf.printf "Traces: "; print_traces !traces; !traces end +end + +let select_traces code entrypoint = + let length = List.length @@ PTree.elements code in + if (length < 5000) then select_traces_chang code entrypoint + else select_traces_linear code entrypoint let rec make_identity_ptree_rec = function | [] -> PTree.empty -- cgit From b675ed29c56a16cd35e5a8d7e49be7c582ebd4ab Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 20 Mar 2020 19:16:52 +0100 Subject: essai d'intégration continue MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- .gitlab-ci.yml | 11 +++++++++++ 1 file changed, 11 insertions(+) create mode 100644 .gitlab-ci.yml diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml new file mode 100644 index 00000000..8e840fd9 --- /dev/null +++ b/.gitlab-ci.yml @@ -0,0 +1,11 @@ +stages: + - build_x86_64 + +.build_x86_64: + stage: build_x86_64 + image: "coqorg/coq" + before_script: + - opam install -y menhir + script: + ./config_x86_64.sh + make -j "$NJOBS" \ No newline at end of file -- cgit From 404cf8865a61dc5c03d4255c42d47462666ace23 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 20 Mar 2020 19:20:00 +0100 Subject: fix YAML syntax --- .gitlab-ci.yml | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 8e840fd9..5e76b79f 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -1,11 +1,11 @@ stages: - build_x86_64 -.build_x86_64: +build_x86_64: stage: build_x86_64 image: "coqorg/coq" before_script: - opam install -y menhir script: ./config_x86_64.sh - make -j "$NJOBS" \ No newline at end of file + make -j "$NJOBS" -- cgit From aa926dbe4652b10ecc77347c99300bc6e00b5be4 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 20 Mar 2020 19:24:54 +0100 Subject: fix syntax --- .gitlab-ci.yml | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 5e76b79f..b6d78ef6 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -7,5 +7,5 @@ build_x86_64: before_script: - opam install -y menhir script: - ./config_x86_64.sh - make -j "$NJOBS" + - ./config_x86_64.sh + - make -j "$NJOBS" -- cgit From 64ce25f46ebe6558d252a878f2289e8ec3901f17 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 20 Mar 2020 19:43:17 +0100 Subject: use later ocaml compiler --- .gitlab-ci.yml | 1 + 1 file changed, 1 insertion(+) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index b6d78ef6..d7093a58 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -5,6 +5,7 @@ build_x86_64: stage: build_x86_64 image: "coqorg/coq" before_script: + - opam switch \$COMPILER_EDGE; eval \$(opam env) - opam install -y menhir script: - ./config_x86_64.sh -- cgit From d045132186970daa0b7b523e4fdbeb87a71fd3ff Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 20 Mar 2020 19:51:55 +0100 Subject: fix opam config env --- .gitlab-ci.yml | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index d7093a58..4f2b8413 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -5,7 +5,8 @@ build_x86_64: stage: build_x86_64 image: "coqorg/coq" before_script: - - opam switch \$COMPILER_EDGE; eval \$(opam env) + - opam switch \$COMPILER_EDGE + - eval `opam config env` - opam install -y menhir script: - ./config_x86_64.sh -- cgit From 43e091310b69f39e0591ee71922121a7508c3987 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 20 Mar 2020 19:55:22 +0100 Subject: fix COMPILER_EDGE --- .gitlab-ci.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 4f2b8413..2745b1eb 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -5,7 +5,7 @@ build_x86_64: stage: build_x86_64 image: "coqorg/coq" before_script: - - opam switch \$COMPILER_EDGE + - opam switch 4.07.1+flambda - eval `opam config env` - opam install -y menhir script: -- cgit From 599ef2ea4cf6f305097809ef123f948ac6c21429 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 20 Mar 2020 20:11:41 +0100 Subject: more architectures --- .gitlab-ci.yml | 77 ++++++++++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 77 insertions(+) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 2745b1eb..6eacd8bc 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -1,5 +1,11 @@ stages: + - build_ppc + - build_aarch64 + - build_arm + - build_rv64 + - build_rv32 - build_x86_64 + - build_x86 build_x86_64: stage: build_x86_64 @@ -11,3 +17,74 @@ build_x86_64: script: - ./config_x86_64.sh - make -j "$NJOBS" + +build_x86: + stage: build_x86 + image: "coqorg/coq" + before_script: + - opam switch 4.07.1+flambda + - eval `opam config env` + - opam install -y menhir + script: + - ./config_x86.sh + - make -j "$NJOBS" + +build_aarch64: + stage: build_aarch64 + image: "coqorg/coq" + before_script: + - sudo apt install gcc-aarch64-linux-gnu + - opam switch 4.07.1+flambda + - eval `opam config env` + - opam install -y menhir + script: + - ./config_aarch64.sh + - make -j "$NJOBS" + +build_arm: + stage: build_arm + image: "coqorg/coq" + before_script: + - sudo apt install gcc-arm-linux-gnueabihf + - opam switch 4.07.1+flambda + - eval `opam config env` + - opam install -y menhir + script: + - ./config_arm.sh + - make -j "$NJOBS" + +build_ppc: + stage: build_ppc + image: "coqorg/coq" + before_script: + - sudo apt install gcc-powerpc-linux + - opam switch 4.07.1+flambda + - eval `opam config env` + - opam install -y menhir + script: + - ./config_ppc.sh + - make -j "$NJOBS" + +build_rv64: + stage: build_rv64 + image: "coqorg/coq" + before_script: + - sudo apt install gcc-riscv64-linux-gnu + - opam switch 4.07.1+flambda + - eval `opam config env` + - opam install -y menhir + script: + - ./config_rv64.sh + - make -j "$NJOBS" + +build_rv32: + stage: build_rv32 + image: "coqorg/coq" + before_script: + - sudo apt install gcc-riscv64-linux-gnu + - opam switch 4.07.1+flambda + - eval `opam config env` + - opam install -y menhir + script: + - ./config_rv32.sh + - make -j "$NJOBS" -- cgit From 995df48bc3920b774a43da7612b1c63faa2ec7d4 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 20 Mar 2020 20:16:12 +0100 Subject: fix spelling --- .gitlab-ci.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 6eacd8bc..9e6d21b0 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -57,7 +57,7 @@ build_ppc: stage: build_ppc image: "coqorg/coq" before_script: - - sudo apt install gcc-powerpc-linux + - sudo apt install gcc-powerpc-linux-gnu - opam switch 4.07.1+flambda - eval `opam config env` - opam install -y menhir -- cgit From 4f3f8ce2a1da0362575087c642cc4e51fe0c9849 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 20 Mar 2020 20:19:18 +0100 Subject: with several jobs --- .gitlab-ci.yml | 22 ++++++++-------------- 1 file changed, 8 insertions(+), 14 deletions(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 9e6d21b0..b31d676f 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -1,14 +1,8 @@ stages: - - build_ppc - - build_aarch64 - - build_arm - - build_rv64 - - build_rv32 - - build_x86_64 - - build_x86 + - build build_x86_64: - stage: build_x86_64 + stage: build image: "coqorg/coq" before_script: - opam switch 4.07.1+flambda @@ -19,7 +13,7 @@ build_x86_64: - make -j "$NJOBS" build_x86: - stage: build_x86 + stage: build image: "coqorg/coq" before_script: - opam switch 4.07.1+flambda @@ -30,7 +24,7 @@ build_x86: - make -j "$NJOBS" build_aarch64: - stage: build_aarch64 + stage: build image: "coqorg/coq" before_script: - sudo apt install gcc-aarch64-linux-gnu @@ -42,7 +36,7 @@ build_aarch64: - make -j "$NJOBS" build_arm: - stage: build_arm + stage: build image: "coqorg/coq" before_script: - sudo apt install gcc-arm-linux-gnueabihf @@ -54,7 +48,7 @@ build_arm: - make -j "$NJOBS" build_ppc: - stage: build_ppc + stage: build image: "coqorg/coq" before_script: - sudo apt install gcc-powerpc-linux-gnu @@ -66,7 +60,7 @@ build_ppc: - make -j "$NJOBS" build_rv64: - stage: build_rv64 + stage: build image: "coqorg/coq" before_script: - sudo apt install gcc-riscv64-linux-gnu @@ -78,7 +72,7 @@ build_rv64: - make -j "$NJOBS" build_rv32: - stage: build_rv32 + stage: build image: "coqorg/coq" before_script: - sudo apt install gcc-riscv64-linux-gnu -- cgit From 64ed780f6ad9512bd4de2a6ab5cf79234e4d34e7 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 20 Mar 2020 20:31:48 +0100 Subject: ia32 --- .gitlab-ci.yml | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index b31d676f..a8120993 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -12,7 +12,7 @@ build_x86_64: - ./config_x86_64.sh - make -j "$NJOBS" -build_x86: +build_ia32: stage: build image: "coqorg/coq" before_script: @@ -20,7 +20,7 @@ build_x86: - eval `opam config env` - opam install -y menhir script: - - ./config_x86.sh + - ./config_ia32.sh - make -j "$NJOBS" build_aarch64: -- cgit From 1069688c9d49fc6ea0f667f640b6a60d7b1fdd84 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 20 Mar 2020 21:01:52 +0100 Subject: fixing aarch64? --- .gitlab-ci.yml | 24 +++++++++++++----------- 1 file changed, 13 insertions(+), 11 deletions(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index a8120993..144c920a 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -1,8 +1,9 @@ stages: - build + - build2 build_x86_64: - stage: build + stage: build2 image: "coqorg/coq" before_script: - opam switch 4.07.1+flambda @@ -13,7 +14,7 @@ build_x86_64: - make -j "$NJOBS" build_ia32: - stage: build + stage: build2 image: "coqorg/coq" before_script: - opam switch 4.07.1+flambda @@ -27,7 +28,8 @@ build_aarch64: stage: build image: "coqorg/coq" before_script: - - sudo apt install gcc-aarch64-linux-gnu + - sudo apt-get update + - sudo apt-get -y install gcc-aarch64-linux-gnu - opam switch 4.07.1+flambda - eval `opam config env` - opam install -y menhir @@ -36,10 +38,10 @@ build_aarch64: - make -j "$NJOBS" build_arm: - stage: build + stage: build2 image: "coqorg/coq" before_script: - - sudo apt install gcc-arm-linux-gnueabihf + - sudo apt-get -y install gcc-arm-linux-gnueabihf - opam switch 4.07.1+flambda - eval `opam config env` - opam install -y menhir @@ -48,10 +50,10 @@ build_arm: - make -j "$NJOBS" build_ppc: - stage: build + stage: build2 image: "coqorg/coq" before_script: - - sudo apt install gcc-powerpc-linux-gnu + - sudo apt-get -y install gcc-powerpc-linux-gnu - opam switch 4.07.1+flambda - eval `opam config env` - opam install -y menhir @@ -60,10 +62,10 @@ build_ppc: - make -j "$NJOBS" build_rv64: - stage: build + stage: build2 image: "coqorg/coq" before_script: - - sudo apt install gcc-riscv64-linux-gnu + - sudo apt-get -y install gcc-riscv64-linux-gnu - opam switch 4.07.1+flambda - eval `opam config env` - opam install -y menhir @@ -72,10 +74,10 @@ build_rv64: - make -j "$NJOBS" build_rv32: - stage: build + stage: build2 image: "coqorg/coq" before_script: - - sudo apt install gcc-riscv64-linux-gnu + - sudo apt-get -y install gcc-riscv64-linux-gnu - opam switch 4.07.1+flambda - eval `opam config env` - opam install -y menhir -- cgit From e8408d8d85b3961a1cf177b3e8a307f80319e4e7 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 20 Mar 2020 21:03:49 +0100 Subject: workaround for time issues --- .gitlab-ci.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 144c920a..abe1724a 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -28,7 +28,7 @@ build_aarch64: stage: build image: "coqorg/coq" before_script: - - sudo apt-get update + - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update - sudo apt-get -y install gcc-aarch64-linux-gnu - opam switch 4.07.1+flambda - eval `opam config env` -- cgit From 090dc09e4dbda97eea4ed28e67d84adf26f807f7 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 20 Mar 2020 21:18:28 +0100 Subject: apt update --- .gitlab-ci.yml | 17 ++++++++++------- 1 file changed, 10 insertions(+), 7 deletions(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index abe1724a..a285d8f2 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -1,9 +1,8 @@ stages: - build - - build2 build_x86_64: - stage: build2 + stage: build image: "coqorg/coq" before_script: - opam switch 4.07.1+flambda @@ -14,7 +13,7 @@ build_x86_64: - make -j "$NJOBS" build_ia32: - stage: build2 + stage: build image: "coqorg/coq" before_script: - opam switch 4.07.1+flambda @@ -38,9 +37,10 @@ build_aarch64: - make -j "$NJOBS" build_arm: - stage: build2 + stage: build image: "coqorg/coq" before_script: + - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update - sudo apt-get -y install gcc-arm-linux-gnueabihf - opam switch 4.07.1+flambda - eval `opam config env` @@ -50,9 +50,10 @@ build_arm: - make -j "$NJOBS" build_ppc: - stage: build2 + stage: build image: "coqorg/coq" before_script: + - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update - sudo apt-get -y install gcc-powerpc-linux-gnu - opam switch 4.07.1+flambda - eval `opam config env` @@ -62,9 +63,10 @@ build_ppc: - make -j "$NJOBS" build_rv64: - stage: build2 + stage: build image: "coqorg/coq" before_script: + - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update - sudo apt-get -y install gcc-riscv64-linux-gnu - opam switch 4.07.1+flambda - eval `opam config env` @@ -74,9 +76,10 @@ build_rv64: - make -j "$NJOBS" build_rv32: - stage: build2 + stage: build image: "coqorg/coq" before_script: + - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update - sudo apt-get -y install gcc-riscv64-linux-gnu - opam switch 4.07.1+flambda - eval `opam config env` -- cgit From 8f337f072a3731dbb778afdb4a5882fd744333a4 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 20 Mar 2020 21:33:13 +0100 Subject: +k1c target --- .gitlab-ci.yml | 11 +++++++++++ 1 file changed, 11 insertions(+) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index a285d8f2..9eafe558 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -87,3 +87,14 @@ build_rv32: script: - ./config_rv32.sh - make -j "$NJOBS" + +build_k1c: + stage: build + image: "coqorg/coq" + before_script: + - opam switch 4.07.1+flambda + - eval `opam config env` + - opam install -y menhir + script: + - ./config_k1c.sh -no-runtime-lib + - make -j "$NJOBS" -- cgit From 3ad897c65d5f16159694ede55958ac8c85339d55 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 20 Mar 2020 22:50:07 +0100 Subject: la lib standard ne passe pas en rv32, ne pas la tester en CI --- .gitlab-ci.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 9eafe558..b4571271 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -85,7 +85,7 @@ build_rv32: - eval `opam config env` - opam install -y menhir script: - - ./config_rv32.sh + - ./config_rv32.sh -no-runtime-lib - make -j "$NJOBS" build_k1c: -- cgit From d19e5a265cdee251b582792b63ed7e91c313579f Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sat, 21 Mar 2020 07:53:37 +0100 Subject: tentative pour n'avoir le gitlab-ci que sur mppa-work, mppa-k1c et master --- .gitlab-ci.yml | 31 +++++++++++++++++++++++++++++++ 1 file changed, 31 insertions(+) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index b4571271..79a00d3b 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -11,6 +11,10 @@ build_x86_64: script: - ./config_x86_64.sh - make -j "$NJOBS" + only: + - master + - mppa-k1c + - mppa-work build_ia32: stage: build @@ -22,6 +26,10 @@ build_ia32: script: - ./config_ia32.sh - make -j "$NJOBS" + only: + - master + - mppa-k1c + - mppa-work build_aarch64: stage: build @@ -35,6 +43,10 @@ build_aarch64: script: - ./config_aarch64.sh - make -j "$NJOBS" + only: + - master + - mppa-k1c + - mppa-work build_arm: stage: build @@ -48,6 +60,10 @@ build_arm: script: - ./config_arm.sh - make -j "$NJOBS" + only: + - master + - mppa-k1c + - mppa-work build_ppc: stage: build @@ -61,6 +77,10 @@ build_ppc: script: - ./config_ppc.sh - make -j "$NJOBS" + only: + - master + - mppa-k1c + - mppa-work build_rv64: stage: build @@ -74,6 +94,10 @@ build_rv64: script: - ./config_rv64.sh - make -j "$NJOBS" + only: + - master + - mppa-k1c + - mppa-work build_rv32: stage: build @@ -87,6 +111,10 @@ build_rv32: script: - ./config_rv32.sh -no-runtime-lib - make -j "$NJOBS" + only: + - master + - mppa-k1c + - mppa-work build_k1c: stage: build @@ -98,3 +126,6 @@ build_k1c: script: - ./config_k1c.sh -no-runtime-lib - make -j "$NJOBS" + only: + - mppa-k1c + - mppa-work -- cgit From 9ce733bd06a6f36a144769b05b9405ea7ebbbfb9 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sat, 21 Mar 2020 10:50:41 +0100 Subject: essai d'avoir le pipeline en manuel OU sur certaines branches --- .gitlab-ci.yml | 95 +++++++++++++++++++++++++++++++++++++++------------------- 1 file changed, 64 insertions(+), 31 deletions(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 79a00d3b..ad684229 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -11,10 +11,14 @@ build_x86_64: script: - ./config_x86_64.sh - make -j "$NJOBS" - only: - - master - - mppa-k1c - - mppa-work + rules: + - if: '$CI_COMMIT_BRANCH == "mppa-work"' + when: always + - if: '$CI_COMMIT_BRANCH == "mppa-k1c"' + when: always + - if: '$CI_COMMIT_BRANCH == "master"' + when: always + - when: manual build_ia32: stage: build @@ -26,10 +30,14 @@ build_ia32: script: - ./config_ia32.sh - make -j "$NJOBS" - only: - - master - - mppa-k1c - - mppa-work + rules: + - if: '$CI_COMMIT_BRANCH == "mppa-work"' + when: always + - if: '$CI_COMMIT_BRANCH == "mppa-k1c"' + when: always + - if: '$CI_COMMIT_BRANCH == "master"' + when: always + - when: manual build_aarch64: stage: build @@ -43,10 +51,14 @@ build_aarch64: script: - ./config_aarch64.sh - make -j "$NJOBS" - only: - - master - - mppa-k1c - - mppa-work + rules: + - if: '$CI_COMMIT_BRANCH == "mppa-work"' + when: always + - if: '$CI_COMMIT_BRANCH == "mppa-k1c"' + when: always + - if: '$CI_COMMIT_BRANCH == "master"' + when: always + - when: manual build_arm: stage: build @@ -60,10 +72,14 @@ build_arm: script: - ./config_arm.sh - make -j "$NJOBS" - only: - - master - - mppa-k1c - - mppa-work + rules: + - if: '$CI_COMMIT_BRANCH == "mppa-work"' + when: always + - if: '$CI_COMMIT_BRANCH == "mppa-k1c"' + when: always + - if: '$CI_COMMIT_BRANCH == "master"' + when: always + - when: manual build_ppc: stage: build @@ -77,10 +93,14 @@ build_ppc: script: - ./config_ppc.sh - make -j "$NJOBS" - only: - - master - - mppa-k1c - - mppa-work + rules: + - if: '$CI_COMMIT_BRANCH == "mppa-work"' + when: always + - if: '$CI_COMMIT_BRANCH == "mppa-k1c"' + when: always + - if: '$CI_COMMIT_BRANCH == "master"' + when: always + - when: manual build_rv64: stage: build @@ -94,10 +114,14 @@ build_rv64: script: - ./config_rv64.sh - make -j "$NJOBS" - only: - - master - - mppa-k1c - - mppa-work + rules: + - if: '$CI_COMMIT_BRANCH == "mppa-work"' + when: always + - if: '$CI_COMMIT_BRANCH == "mppa-k1c"' + when: always + - if: '$CI_COMMIT_BRANCH == "master"' + when: always + - when: manual build_rv32: stage: build @@ -111,10 +135,14 @@ build_rv32: script: - ./config_rv32.sh -no-runtime-lib - make -j "$NJOBS" - only: - - master - - mppa-k1c - - mppa-work + rules: + - if: '$CI_COMMIT_BRANCH == "mppa-work"' + when: always + - if: '$CI_COMMIT_BRANCH == "mppa-k1c"' + when: always + - if: '$CI_COMMIT_BRANCH == "master"' + when: always + - when: manual build_k1c: stage: build @@ -126,6 +154,11 @@ build_k1c: script: - ./config_k1c.sh -no-runtime-lib - make -j "$NJOBS" - only: - - mppa-k1c - - mppa-work + rules: + - if: '$CI_COMMIT_BRANCH == "mppa-work"' + when: always + - if: '$CI_COMMIT_BRANCH == "mppa-k1c"' + when: always + - if: '$CI_COMMIT_BRANCH == "master"' + when: always + - when: manual -- cgit From 1033c2a0ffefc336c343888e1abda02d7a1db228 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Mon, 23 Mar 2020 16:32:23 +0100 Subject: Removing store heuristic and more fine tuning loop heuristic --- backend/Duplicateaux.ml | 12 ++++++++---- backend/Linearizeaux.ml | 2 +- 2 files changed, 9 insertions(+), 5 deletions(-) diff --git a/backend/Duplicateaux.ml b/backend/Duplicateaux.ml index b9f5cdf2..9ee082ea 100644 --- a/backend/Duplicateaux.ml +++ b/backend/Duplicateaux.ml @@ -195,12 +195,16 @@ let do_store_heuristic code cond ifso ifnot is_loop_header = let do_loop_heuristic code cond ifso ifnot is_loop_header = begin Printf.printf "\tLoop heuristic..\n"; - let predicate n = get_some @@ PTree.get n is_loop_header - in if (look_ahead code ifso is_loop_header predicate) then Some true - else if (look_ahead code ifnot is_loop_header predicate) then Some false + let predicate n = get_some @@ PTree.get n is_loop_header in + let ifso_loop = look_ahead code ifso is_loop_header predicate in + let ifnot_loop = look_ahead code ifnot is_loop_header predicate in + if ifso_loop && ifnot_loop then None (* TODO - take the innermost loop ? *) + else if ifso_loop then Some true + else if ifnot_loop then Some false else None end + (* Remark - compared to the original paper, we don't use the store heuristic *) let get_directions code entrypoint = begin Printf.printf "get_directions\n"; flush stdout; let bfs_order = bfs code entrypoint @@ -214,7 +218,7 @@ let get_directions code entrypoint = begin | Icond (cond, lr, ifso, ifnot, _) -> (* Printf.printf "Analyzing %d.." (P.to_int n); *) let heuristics = [ do_call_heuristic; do_opcode_heuristic; - do_return_heuristic; do_store_heuristic; do_loop_heuristic ] in + do_return_heuristic; do_loop_heuristic; (* do_store_heuristic *) ] in let preferred = ref None in begin Printf.printf "Deciding condition for RTL node %d\n" (P.to_int n); diff --git a/backend/Linearizeaux.ml b/backend/Linearizeaux.ml index 23ced4c2..e68a9b9a 100644 --- a/backend/Linearizeaux.ml +++ b/backend/Linearizeaux.ml @@ -154,7 +154,7 @@ let forward_sequences code entry = let visited = ref (PTree.map (fun n i -> false) code) in (* returns the list of traversed nodes, and a list of nodes to start traversing next *) let rec traverse_fallthrough code node = - Printf.printf "Traversing %d..\n" (P.to_int node); + (* Printf.printf "Traversing %d..\n" (P.to_int node); *) if not (get_some @@ PTree.get node !visited) then begin visited := PTree.set node true !visited; match PTree.get node code with -- cgit From 815fd43dc43ea85d35f1275bd701f5b370ced2a5 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Tue, 24 Mar 2020 11:46:13 +0100 Subject: Linearizeaux: Refining block selection in case of tie --- backend/Linearizeaux.ml | 56 ++++++++++++++++++++++++++++++++----------------- 1 file changed, 37 insertions(+), 19 deletions(-) diff --git a/backend/Linearizeaux.ml b/backend/Linearizeaux.ml index e68a9b9a..3f207d9e 100644 --- a/backend/Linearizeaux.ml +++ b/backend/Linearizeaux.ml @@ -163,15 +163,15 @@ let forward_sequences code entry = let ln, rem = match (last_element bb) with | Lop _ | Lload _ | Lgetstack _ | Lsetstack _ | Lstore _ | Lcall _ | Lbuiltin _ -> assert false - | Ltailcall _ | Lreturn -> begin Printf.printf "STOP tailcall/return\n"; ([], []) end + | Ltailcall _ | Lreturn -> begin (* Printf.printf "STOP tailcall/return\n"; *) ([], []) end | Lbranch n -> let ln, rem = traverse_fallthrough code n in (ln, rem) | Lcond (_, _, ifso, ifnot, info) -> (match info with - | None -> begin Printf.printf "STOP Lcond None\n"; ([], [ifso; ifnot]) end + | None -> begin (* Printf.printf "STOP Lcond None\n"; *) ([], [ifso; ifnot]) end | Some false -> let ln, rem = traverse_fallthrough code ifnot in (ln, [ifso] @ rem) | Some true -> let errstr = Printf.sprintf ("Inconsistency detected in node %d: ifnot is not the preferred branch") (P.to_int node) in failwith errstr) - | Ljumptable(_, ln) -> begin Printf.printf "STOP Ljumptable\n"; ([], ln) end + | Ljumptable(_, ln) -> begin (* Printf.printf "STOP Ljumptable\n"; *) ([], ln) end in ([node] @ ln, rem) end else ([], []) @@ -423,31 +423,49 @@ let order_sequences code entry fs = depmap.(i) <- ISet.remove s_id deps ) depmap end + in let choose_best_of candidates = + let current_best_id = ref None in + let current_best_score = ref None in + begin + List.iter (fun id -> + match !current_best_id with + | None -> begin + current_best_id := Some id; + match fs_a.(id) with + | [] -> current_best_score := None + | n::l -> current_best_score := Some (P.to_int n) + end + | Some b -> begin + match fs_a.(id) with + | [] -> () + | n::l -> let nscore = P.to_int n in + match !current_best_score with + | None -> (current_best_id := Some id; current_best_score := Some nscore) + | Some bs -> if nscore > bs then (current_best_id := Some id; current_best_score := Some nscore) + end + ) candidates; + !current_best_id + end in let select_next () = - let selected_id = ref None in + let candidates = ref [] in begin Array.iteri (fun i deps -> begin (* Printf.printf "Deps: "; print_iset deps; Printf.printf "\n"; *) - match !selected_id with - | None -> if (deps == ISet.empty && not fs_evaluated.(i)) then selected_id := Some i - | Some id -> () + if (deps == ISet.empty && not fs_evaluated.(i)) then candidates := i :: !candidates end ) depmap; - match !selected_id with - | Some id -> id - | None -> begin - Array.iteri (fun i deps -> - match !selected_id with - | None -> if not fs_evaluated.(i) then selected_id := Some i - | Some id -> () - ) depmap; - get_some !selected_id - end + if not (List.length !candidates > 0) then begin + Array.iteri (fun i deps -> + if (not fs_evaluated.(i)) then candidates := i :: !candidates + ) depmap; + end; + get_some (choose_best_of !candidates) end in begin - (* Printf.printf "depmap: "; print_depmap depmap; *) - (* Printf.printf "forward sequences identified: "; print_ssequence fs; *) + Printf.printf "-------------------------------\n"; + Printf.printf "depmap: "; print_depmap depmap; + Printf.printf "forward sequences identified: "; print_ssequence fs; while List.length !ordered_fs != List.length fs do let next_id = select_next () in evaluate next_id -- cgit From c1aa5f9678c2a90453c57f9918b349753fdf50be Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Tue, 24 Mar 2020 17:08:13 +0100 Subject: Duplicate: added another loop heuristic which should detect loop branches better --- backend/Duplicateaux.ml | 62 +++++++++++++++++++++++++++++++++++++++++++++---- 1 file changed, 57 insertions(+), 5 deletions(-) diff --git a/backend/Duplicateaux.ml b/backend/Duplicateaux.ml index 9ee082ea..fedb63fe 100644 --- a/backend/Duplicateaux.ml +++ b/backend/Duplicateaux.ml @@ -204,13 +204,65 @@ let do_loop_heuristic code cond ifso ifnot is_loop_header = else None end +let do_loop2_heuristic loop_info n code cond ifso ifnot is_loop_header = + begin + Printf.printf "\tLoop2 heuristic..\n"; + match get_some @@ PTree.get n loop_info with + | None -> None + | Some b -> Some b + end + +(* Returns a PTree of either None or Some b where b determines the node following the loop, for a cb instruction *) +(* It uses the fact that loops in CompCert are done by a branch (backedge) instruction followed by a cb *) +let get_loop_info is_loop_header bfs_order code = + let loop_info = ref (PTree.map (fun n i -> None) code) in + let mark_path s n = + let visited = ref (PTree.map (fun n i -> false) code) in + let rec explore src dest = + if (get_some @@ PTree.get src !visited) then false + else if src == dest then true + else begin + visited := PTree.set src true !visited; + match get_some @@ PTree.get src code with + | Inop s | Iop (_, _, _, s) | Iload (_,_,_,_,_,s) | Istore (_,_,_,_,s) | Icall (_,_,_,_,s) + | Ibuiltin (_,_,_,s) -> explore s dest + | Icond (_,_,s1,s2,_) -> (explore s1 dest) || (explore s2 dest) + | Ijumptable _ | Itailcall _ | Ireturn _ -> false + end + in match get_some @@ PTree.get s !loop_info with + | None -> begin + match get_some @@ PTree.get s code with + | Icond (_, _, n1, n2, _) -> + let b1 = explore n1 n in + let b2 = explore n2 n in + if (b1 && b2) then () + else if b1 then loop_info := PTree.set s (Some true) !loop_info + else if b2 then loop_info := PTree.set s (Some false) !loop_info + else () + | _ -> () + end + | Some _ -> () + in begin + List.iter (fun n -> + match get_some @@ PTree.get n code with + | Inop s | Iop (_,_,_,s) | Iload (_,_,_,_,_,s) | Istore (_,_,_,_,s) | Icall (_,_,_,_,s) + | Ibuiltin (_, _, _, s) -> + if get_some @@ PTree.get s is_loop_header then mark_path s n + | Icond _ -> () (* loop backedges are never Icond in CompCert *) + | Ijumptable _ -> () + | Itailcall _ | Ireturn _ -> () + ) bfs_order; + !loop_info + end + (* Remark - compared to the original paper, we don't use the store heuristic *) let get_directions code entrypoint = begin Printf.printf "get_directions\n"; flush stdout; - let bfs_order = bfs code entrypoint - and is_loop_header = get_loop_headers code entrypoint - and directions = ref (PTree.map (fun n i -> None) code) (* None <=> no predicted direction *) - in begin + let bfs_order = bfs code entrypoint in + let is_loop_header = get_loop_headers code entrypoint in + let loop_info = get_loop_info is_loop_header bfs_order code in + let directions = ref (PTree.map (fun n i -> None) code) in (* None <=> no predicted direction *) + begin (* ptree_printbool is_loop_header; *) (* Printf.printf "\n"; *) List.iter (fun n -> @@ -218,7 +270,7 @@ let get_directions code entrypoint = begin | Icond (cond, lr, ifso, ifnot, _) -> (* Printf.printf "Analyzing %d.." (P.to_int n); *) let heuristics = [ do_call_heuristic; do_opcode_heuristic; - do_return_heuristic; do_loop_heuristic; (* do_store_heuristic *) ] in + do_return_heuristic; do_loop2_heuristic loop_info n; do_loop_heuristic; (* do_store_heuristic *) ] in let preferred = ref None in begin Printf.printf "Deciding condition for RTL node %d\n" (P.to_int n); -- cgit From ee7cd36732efd3af91f8d6cb9be18a58e0ff43a3 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 24 Mar 2020 20:03:00 +0100 Subject: exec_straight_steps exec_straight_steps_goto exec_straight_opt_steps_goto --- aarch64/Asmgenproof.v | 12 +++++++----- 1 file changed, 7 insertions(+), 5 deletions(-) diff --git a/aarch64/Asmgenproof.v b/aarch64/Asmgenproof.v index eeff1956..5353d1ab 100644 --- a/aarch64/Asmgenproof.v +++ b/aarch64/Asmgenproof.v @@ -472,7 +472,8 @@ Inductive match_states: Mach.state -> Asm.state -> Prop := (MEXT: Mem.extends m m') (AT: transl_code_at_pc ge (rs PC) fb f c ep tf tc) (AG: agree ms sp rs) - (DXP: ep = true -> rs#X29 = parent_sp s), + (DXP: ep = true -> rs#X29 = parent_sp s) + (LEAF: is_leaf_function f = true -> rs#X30 = parent_ra s), match_states (Mach.State s fb sp c ms m) (Asm.State rs m') | match_states_call: @@ -503,16 +504,17 @@ Lemma exec_straight_steps: exists rs2, exec_straight tge tf c rs1 m1' k rs2 m2' /\ agree ms2 sp rs2 - /\ (it1_is_parent ep i = true -> rs2#X29 = parent_sp s)) -> + /\ (it1_is_parent ep i = true -> rs2#X29 = parent_sp s) + /\ (is_leaf_function f = true -> rs2#X30 = parent_ra s)) -> exists st', plus step tge (State rs1 m1') E0 st' /\ match_states (Mach.State s fb sp c ms2 m2) st'. Proof. intros. inversion H2. subst. monadInv H7. - exploit H3; eauto. intros [rs2 [A [B C]]]. + exploit H3; eauto. intros [rs2 [A [B [C D]]]]. exists (State rs2 m2'); split. - eapply exec_straight_exec; eauto. - econstructor; eauto. eapply exec_straight_at; eauto. + - eapply exec_straight_exec; eauto. + - econstructor; eauto. eapply exec_straight_at; eauto. Qed. Lemma exec_straight_steps_goto: -- cgit From b59b1c908b1f412591accba7d2ecb5818062c3f9 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 24 Mar 2020 20:47:36 +0100 Subject: progress in proofs about RA --- aarch64/Asmgenproof.v | 55 +++++++++++++++++++++++-------------- aarch64/Asmgenproof1.v | 73 +++++++++++++++++++++++++++++++++++++++----------- 2 files changed, 92 insertions(+), 36 deletions(-) diff --git a/aarch64/Asmgenproof.v b/aarch64/Asmgenproof.v index 5353d1ab..0dc37f36 100644 --- a/aarch64/Asmgenproof.v +++ b/aarch64/Asmgenproof.v @@ -473,7 +473,7 @@ Inductive match_states: Mach.state -> Asm.state -> Prop := (AT: transl_code_at_pc ge (rs PC) fb f c ep tf tc) (AG: agree ms sp rs) (DXP: ep = true -> rs#X29 = parent_sp s) - (LEAF: is_leaf_function f = true -> rs#X30 = parent_ra s), + (LEAF: is_leaf_function f = true -> rs#RA = parent_ra s), match_states (Mach.State s fb sp c ms m) (Asm.State rs m') | match_states_call: @@ -505,7 +505,7 @@ Lemma exec_straight_steps: exec_straight tge tf c rs1 m1' k rs2 m2' /\ agree ms2 sp rs2 /\ (it1_is_parent ep i = true -> rs2#X29 = parent_sp s) - /\ (is_leaf_function f = true -> rs2#X30 = parent_ra s)) -> + /\ (is_leaf_function f = true -> rs2#RA = parent_ra s)) -> exists st', plus step tge (State rs1 m1') E0 st' /\ match_states (Mach.State s fb sp c ms2 m2) st'. @@ -529,13 +529,14 @@ Lemma exec_straight_steps_goto: exists jmp, exists k', exists rs2, exec_straight tge tf c rs1 m1' (jmp :: k') rs2 m2' /\ agree ms2 sp rs2 - /\ exec_instr tge tf jmp rs2 m2' = goto_label tf lbl rs2 m2') -> + /\ exec_instr tge tf jmp rs2 m2' = goto_label tf lbl rs2 m2' + /\ (is_leaf_function f = true -> rs2#RA = parent_ra s)) -> exists st', plus step tge (State rs1 m1') E0 st' /\ match_states (Mach.State s fb sp c' ms2 m2) st'. Proof. intros. inversion H3. subst. monadInv H9. - exploit H5; eauto. intros [jmp [k' [rs2 [A [B C]]]]]. + exploit H5; eauto. intros [jmp [k' [rs2 [A [B [C D]]]]]]. generalize (functions_transl _ _ _ H7 H8); intro FN. generalize (transf_function_no_overflow _ _ H8); intro NOOV. exploit exec_straight_steps_2; eauto. @@ -552,6 +553,7 @@ Proof. econstructor; eauto. apply agree_exten with rs2; auto with asmgen. congruence. + rewrite OTH by congruence; auto. Qed. Lemma exec_straight_opt_steps_goto: @@ -566,13 +568,14 @@ Lemma exec_straight_opt_steps_goto: exists jmp, exists k', exists rs2, exec_straight_opt tge tf c rs1 m1' (jmp :: k') rs2 m2' /\ agree ms2 sp rs2 - /\ exec_instr tge tf jmp rs2 m2' = goto_label tf lbl rs2 m2') -> + /\ exec_instr tge tf jmp rs2 m2' = goto_label tf lbl rs2 m2' + /\ (is_leaf_function f = true -> rs2#RA = parent_ra s)) -> exists st', plus step tge (State rs1 m1') E0 st' /\ match_states (Mach.State s fb sp c' ms2 m2) st'. Proof. intros. inversion H3. subst. monadInv H9. - exploit H5; eauto. intros [jmp [k' [rs2 [A [B C]]]]]. + exploit H5; eauto. intros [jmp [k' [rs2 [A [B [C D]]]]]]. generalize (functions_transl _ _ _ H7 H8); intro FN. generalize (transf_function_no_overflow _ _ H8); intro NOOV. inv A. @@ -585,6 +588,7 @@ Proof. econstructor; eauto. apply agree_exten with rs2; auto with asmgen. congruence. + rewrite OTH by congruence; auto. - exploit exec_straight_steps_2; eauto. intros [ofs' [PC2 CT2]]. exploit find_label_goto_label; eauto. @@ -599,6 +603,7 @@ Proof. econstructor; eauto. apply agree_exten with rs2; auto with asmgen. congruence. + rewrite OTH by congruence; auto. Qed. (** We need to show that, in the simulation diagram, we cannot @@ -640,17 +645,20 @@ Proof. - (* Mlabel *) left; eapply exec_straight_steps; eauto; intros. monadInv TR. econstructor; split. apply exec_straight_one. simpl; eauto. auto. - split. apply agree_nextinstr; auto. simpl; congruence. + split. { apply agree_nextinstr; auto. } + split. { simpl; congruence. } + rewrite nextinstr_inv by congruence; assumption. - (* Mgetstack *) unfold load_stack in H. exploit Mem.loadv_extends; eauto. intros [v' [A B]]. rewrite (sp_val _ _ _ AG) in A. left; eapply exec_straight_steps; eauto. intros. simpl in TR. - exploit loadind_correct; eauto with asmgen. intros [rs' [P [Q R]]]. + exploit loadind_correct; eauto with asmgen. intros [rs' [P [Q [R S]]]]. exists rs'; split. eauto. - split. eapply agree_set_mreg; eauto with asmgen. congruence. - simpl; congruence. + split. { eapply agree_set_mreg; eauto with asmgen. congruence. } + split. { simpl; congruence. } + rewrite S. assumption. - (* Msetstack *) unfold store_stack in H. @@ -658,10 +666,12 @@ Proof. exploit Mem.storev_extends; eauto. intros [m2' [A B]]. left; eapply exec_straight_steps; eauto. rewrite (sp_val _ _ _ AG) in A. intros. simpl in TR. - exploit storeind_correct; eauto with asmgen. intros [rs' [P Q]]. + exploit storeind_correct; eauto with asmgen. intros [rs' [P [Q R]]]. exists rs'; split. eauto. split. eapply agree_undef_regs; eauto with asmgen. - simpl; intros. rewrite Q; auto with asmgen. + simpl; intros. + split. rewrite Q; auto with asmgen. + rewrite R. assumption. - (* Mgetparam *) assert (f0 = f) by congruence; subst f0. @@ -677,24 +687,29 @@ Opaque loadind. (* X30 contains parent *) exploit loadind_correct. eexact EQ. instantiate (2 := rs0). simpl; rewrite DXP; eauto. simpl; congruence. - intros [rs1 [P [Q R]]]. + intros [rs1 [P [Q [R S]]]]. exists rs1; split. eauto. split. eapply agree_set_mreg. eapply agree_set_mreg; eauto. congruence. auto with asmgen. - simpl; intros. rewrite R; auto with asmgen. - apply preg_of_not_X29; auto. + simpl; split; intros. + { rewrite R; auto with asmgen. + apply preg_of_not_X29; auto. + } + { rewrite S; auto. } + (* X30 does not contain parent *) exploit loadptr_correct. eexact A. simpl; congruence. intros [rs1 [P [Q R]]]. exploit loadind_correct. eexact EQ. instantiate (2 := rs1). simpl; rewrite Q. eauto. simpl; congruence. - intros [rs2 [S [T U]]]. + intros [rs2 [S [T [U V]]]]. exists rs2; split. eapply exec_straight_trans; eauto. split. eapply agree_set_mreg. eapply agree_set_mreg. eauto. eauto. instantiate (1 := rs1#X29 <- (rs2#X29)). intros. rewrite Pregmap.gso; auto with asmgen. congruence. intros. unfold Pregmap.set. destruct (PregEq.eq r' X29). congruence. auto with asmgen. - simpl; intros. rewrite U; auto with asmgen. + split; simpl; intros. rewrite U; auto with asmgen. apply preg_of_not_X29; auto. - + rewrite V. rewrite R by congruence. auto. + - (* Mop *) assert (eval_operation tge sp op (map rs args) m = Some v). { rewrite <- H. apply eval_operation_preserved. exact symbols_preserved. } @@ -705,11 +720,11 @@ Opaque loadind. exists rs2; split. eauto. split. apply agree_set_undef_mreg with rs0; auto. apply Val.lessdef_trans with v'; auto. - simpl; intros. InvBooleans. + split; simpl; intros. InvBooleans. rewrite R; auto. apply preg_of_not_X29; auto. Local Transparent destroyed_by_op. destruct op; try exact I; simpl; congruence. - + - (* Mload *) assert (Op.eval_addressing tge sp addr (map rs args) = Some a). { rewrite <- H. apply eval_addressing_preserved. exact symbols_preserved. } diff --git a/aarch64/Asmgenproof1.v b/aarch64/Asmgenproof1.v index 6d44bcc8..c85543f3 100644 --- a/aarch64/Asmgenproof1.v +++ b/aarch64/Asmgenproof1.v @@ -22,6 +22,37 @@ Local Transparent Archi.ptr64. (** Properties of registers *) +Lemma preg_of_not_RA: + forall r, (preg_of r) <> RA. +Proof. + destruct r; discriminate. +Qed. + +Lemma RA_not_written: + forall (rs : regset) dst v, + rs # (preg_of dst) <- v RA = rs RA. +Proof. + intros. + apply Pregmap.gso. + intro. + symmetry in H. + exact (preg_of_not_RA dst H). +Qed. + +Hint Resolve RA_not_written : asmgen. + +Lemma RA_not_written2: + forall (rs : regset) dst v i, + preg_of dst = i -> + rs # i <- v RA = rs RA. +Proof. + intros. + subst i. + apply RA_not_written. +Qed. + +Hint Resolve RA_not_written2 : asmgen. + Lemma preg_of_iregsp_not_PC: forall r, preg_of_iregsp r <> PC. Proof. destruct r; simpl; congruence. @@ -1347,13 +1378,15 @@ Ltac TranslOpSimpl := [ apply exec_straight_one; [simpl; eauto | reflexivity] | split; [ rewrite ? transl_eval_shift, ? transl_eval_shiftl; apply Val.lessdef_same; Simpl; fail - | intros; Simpl; fail ] ]. + | split; [ intros; Simpl; fail + | intros; Simpl; eapply RA_not_written2; eauto] ]]. Ltac TranslOpBase := econstructor; split; [ apply exec_straight_one; [simpl; eauto | reflexivity] | split; [ rewrite ? transl_eval_shift, ? transl_eval_shiftl; Simpl - | intros; Simpl; fail ] ]. + | split; [ intros; Simpl; fail + | intros; Simpl; eapply RA_not_written2; eauto] ]]. Lemma transl_op_correct: forall op args res k (rs: regset) m v c, @@ -1362,15 +1395,15 @@ Lemma transl_op_correct: 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. + /\ (forall r, data_preg r = true -> r <> preg_of res -> preg_notin r (destroyed_by_op op) -> rs' r = rs r) + /\ rs' RA = rs RA. Proof. Local Opaque Int.eq Int64.eq Val.add Val.addl Int.zwordsize Int64.zwordsize. 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) eqn:RR; try discriminate; destruct (preg_of m0) eqn:R1; inv TR. -+ TranslOpSimpl. -+ TranslOpSimpl. + all: TranslOpSimpl. - (* intconst *) exploit exec_loadimm32. intros (rs' & A & B & C). exists rs'; split. eexact A. split. rewrite B; auto. intros; auto with asmgen. @@ -1712,7 +1745,7 @@ Lemma loadptr_correct: forall (base: iregsp) ofs dst k m v (rs: regset), exists rs', exec_straight ge fn (loadptr base ofs dst k) rs m k rs' m /\ rs'#dst = v - /\ forall r, r <> PC -> r <> X16 -> r <> dst -> rs' r = rs r. + /\ (forall r, r <> PC -> r <> X16 -> r <> dst -> rs' r = rs r). Proof. intros. destruct (Val.offset_ptr rs#base ofs) eqn:V; try discriminate. @@ -1720,7 +1753,8 @@ Proof. econstructor; split. eapply exec_straight_opt_right. eexact A. apply exec_straight_one. simpl. unfold exec_load. rewrite B, H. eauto. auto. - split. Simpl. intros; Simpl. + split. Simpl. + intros; Simpl. Qed. Lemma storeptr_correct: forall (base: iregsp) ofs (src: ireg) k m m' (rs: regset), @@ -1729,7 +1763,8 @@ Lemma storeptr_correct: forall (base: iregsp) ofs (src: ireg) k m m' (rs: regset src <> X16 -> exists rs', exec_straight ge fn (storeptr src base ofs k) rs m k rs' m' - /\ forall r, r <> PC -> r <> X16 -> rs' r = rs r. + /\ (forall r, r <> PC -> r <> X16 -> rs' r = rs r) + /\ rs' RA = rs RA. Proof. intros. destruct (Val.offset_ptr rs#base ofs) eqn:V; try discriminate. @@ -1737,7 +1772,7 @@ Proof. econstructor; split. eapply exec_straight_opt_right. eexact A. apply exec_straight_one. simpl. unfold exec_store. rewrite B, C, H by eauto with asmgen. eauto. auto. - intros; Simpl. + split; intros; Simpl. Qed. Lemma loadind_correct: forall (base: iregsp) ofs ty dst k c (rs: regset) m v, @@ -1747,7 +1782,8 @@ Lemma loadind_correct: forall (base: iregsp) ofs ty dst k c (rs: regset) m v, exists rs', exec_straight ge fn c rs m k rs' m /\ rs'#(preg_of dst) = v - /\ forall r, data_preg r = true -> r <> preg_of dst -> rs' r = rs r. + /\ (forall r, data_preg r = true -> r <> preg_of dst -> rs' r = rs r) + /\ rs' RA = rs RA. Proof. intros. destruct (Val.offset_ptr rs#base ofs) eqn:V; try discriminate. @@ -1763,7 +1799,10 @@ Proof. econstructor; split. eapply exec_straight_opt_right. eexact A. apply exec_straight_one. rewrite SEM. unfold exec_load. rewrite B, H0. eauto. Simpl. - split. Simpl. intros; Simpl. + split. Simpl. + split. intros; Simpl. + Simpl. rewrite RA_not_written. + apply C; congruence. Qed. Lemma storeind_correct: forall (base: iregsp) ofs ty src k c (rs: regset) m m', @@ -1772,7 +1811,8 @@ Lemma storeind_correct: forall (base: iregsp) ofs ty src k c (rs: regset) m m', preg_of_iregsp base <> IR X16 -> exists rs', exec_straight ge fn c rs m k rs' m' - /\ forall r, data_preg r = true -> rs' r = rs r. + /\ (forall r, data_preg r = true -> rs' r = rs r) + /\ rs' RA = rs RA. Proof. intros. destruct (Val.offset_ptr rs#base ofs) eqn:V; try discriminate. @@ -1790,7 +1830,8 @@ Proof. apply exec_straight_one. rewrite SEM. unfold exec_store. rewrite B, C, H0 by eauto with asmgen. eauto. Simpl. - intros; Simpl. + split. intros; Simpl. + Simpl. Qed. Lemma make_epilogue_correct: @@ -1807,7 +1848,7 @@ Lemma make_epilogue_correct: /\ Mem.extends m' tm' /\ rs'#RA = parent_ra cs /\ rs'#SP = parent_sp cs - /\ (forall r, r <> PC -> r <> SP -> r <> X30 -> r <> X16 -> rs'#r = rs#r). + /\ (forall r, r <> PC -> r <> SP -> r <> RA -> r <> X16 -> 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'). @@ -1815,7 +1856,7 @@ Proof. 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. + unfold make_epilogue. exploit (loadptr_correct XSP (fn_retaddr_ofs f)). instantiate (2 := rs). simpl. rewrite <- (sp_val _ _ _ AG). simpl. eexact LRA'. simpl; congruence. intros (rs1 & A1 & B1 & C1). @@ -1833,4 +1874,4 @@ Proof. intros. Simpl. Qed. -End CONSTRUCTORS. \ No newline at end of file +End CONSTRUCTORS. -- cgit From 4f29e1e0a8cea00a52fa31f42f050fcd90eb739c Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 24 Mar 2020 23:36:56 +0100 Subject: transl_cond --- aarch64/Asmgenproof1.v | 475 ++++++++++++++++++++++++++++++++++--------------- 1 file changed, 332 insertions(+), 143 deletions(-) diff --git a/aarch64/Asmgenproof1.v b/aarch64/Asmgenproof1.v index c85543f3..96561da7 100644 --- a/aarch64/Asmgenproof1.v +++ b/aarch64/Asmgenproof1.v @@ -53,6 +53,20 @@ Qed. Hint Resolve RA_not_written2 : asmgen. +Lemma RA_not_written3: + forall (rs : regset) dst v i, + ireg_of dst = OK i -> + rs # i <- v RA = rs RA. +Proof. + intros. + unfold ireg_of in H. + destruct preg_of eqn:PREG; try discriminate. + replace i0 with i in * by congruence. + eapply RA_not_written2; eassumption. +Qed. + +Hint Resolve RA_not_written3 : asmgen. + Lemma preg_of_iregsp_not_PC: forall r, preg_of_iregsp r <> PC. Proof. destruct r; simpl; congruence. @@ -70,6 +84,26 @@ Proof. red; intros; subst x. elim (preg_of_not_X16 r); auto. Qed. +Lemma ireg_of_not_RA: forall r x, ireg_of r = OK x -> x <> RA. +Proof. + unfold ireg_of; intros. destruct (preg_of r) eqn:E; inv H. + red; intros; subst x. elim (preg_of_not_RA r); auto. +Qed. + +Lemma ireg_of_not_RA': forall r x, ireg_of r = OK x -> RA <> x. +Proof. + intros. intro. + apply (ireg_of_not_RA r x); auto. +Qed. + +Lemma ireg_of_not_RA'': forall r x, ireg_of r = OK x -> IR RA <> IR x. +Proof. + intros. intro. + apply (ireg_of_not_RA' r x); auto. congruence. +Qed. + +Hint Resolve ireg_of_not_RA ireg_of_not_RA' ireg_of_not_RA'' : asmgen. + Lemma ireg_of_not_X16': forall r x, ireg_of r = OK x -> IR x <> IR X16. Proof. intros. apply ireg_of_not_X16 in H. congruence. @@ -236,42 +270,49 @@ Qed. Lemma exec_loadimm_k_w: forall (rd: ireg) k m l, wf_decomposition l -> + rd <> RA -> forall (rs: regset) accu, rs#rd = Vint (Int.repr accu) -> exists rs', exec_straight_opt ge fn (loadimm_k W rd l k) rs m k rs' m /\ rs'#rd = Vint (Int.repr (recompose_int accu l)) - /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r. + /\ (forall r, r <> PC -> r <> rd -> rs'#r = rs#r) + /\ rs' # RA = rs # RA. Proof. - induction 1; intros rs accu ACCU; simpl. + induction 1; intros RD_NOT_RA rs accu ACCU; simpl. - exists rs; split. apply exec_straight_opt_refl. auto. -- destruct (IHwf_decomposition +- destruct (IHwf_decomposition RD_NOT_RA (nextinstr (rs#rd <- (insert_in_int rs#rd n p 16))) (Zinsert accu n p 16)) - as (rs' & P & Q & R). + as (rs' & P & Q & R & S). Simpl. rewrite ACCU. simpl. f_equal. apply Int.eqm_samerepr. apply Zinsert_eqmod. auto. omega. apply Int.eqm_sym; apply Int.eqm_unsigned_repr. exists rs'; split. eapply exec_straight_opt_step_opt. simpl; eauto. auto. exact P. - split. exact Q. intros; Simpl. rewrite R by auto. Simpl. + split. exact Q. + split. + { intros; Simpl. + rewrite R by auto. Simpl. } + { rewrite S. Simpl. } Qed. Lemma exec_loadimm_z_w: forall rd l k rs m, wf_decomposition l -> + rd <> RA -> exists rs', exec_straight ge fn (loadimm_z W rd l k) rs m k rs' m /\ rs'#rd = Vint (Int.repr (recompose_int 0 l)) /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r. Proof. - unfold loadimm_z; destruct 1. + unfold loadimm_z; destruct 1; intro RD_NOT_RA. - econstructor; split. apply exec_straight_one. simpl; eauto. auto. split. Simpl. intros; Simpl. - set (accu0 := Zinsert 0 n p 16). set (rs1 := nextinstr (rs#rd <- (Vint (Int.repr accu0)))). - destruct (exec_loadimm_k_w rd k m l H1 rs1 accu0) as (rs2 & P & Q & R); auto. + destruct (exec_loadimm_k_w rd k m l H1 RD_NOT_RA rs1 accu0) as (rs2 & P & Q & R & S); auto. unfold rs1; Simpl. exists rs2; split. eapply exec_straight_opt_step; eauto. @@ -284,12 +325,13 @@ Qed. Lemma exec_loadimm_n_w: forall rd l k rs m, wf_decomposition l -> + rd <> RA -> exists rs', exec_straight ge fn (loadimm_n W rd l k) rs m k rs' m /\ rs'#rd = Vint (Int.repr (Z.lnot (recompose_int 0 l))) /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r. Proof. - unfold loadimm_n; destruct 1. + unfold loadimm_n; destruct 1; intro RD_NOT_RA. - econstructor; split. apply exec_straight_one. simpl; eauto. auto. split. Simpl. @@ -298,7 +340,8 @@ Proof. set (rs1 := nextinstr (rs#rd <- (Vint (Int.repr accu0)))). destruct (exec_loadimm_k_w rd k m (negate_decomposition l) (negate_decomposition_wf l H1) - rs1 accu0) as (rs2 & P & Q & R). + RD_NOT_RA rs1 accu0) + as (rs2 & P & Q & R & S). unfold rs1; Simpl. exists rs2; split. eapply exec_straight_opt_step; eauto. @@ -310,7 +353,8 @@ Proof. Qed. Lemma exec_loadimm32: - forall rd n k rs m, + forall rd n k rs m + (RD_NOT_RA : rd <> RA), exists rs', exec_straight ge fn (loadimm32 rd n k) rs m k rs' m /\ rs'#rd = Vint n @@ -333,13 +377,14 @@ Proof. apply Int.eqm_samerepr. apply decompose_notint_eqmod. apply Int.repr_unsigned. } destruct Nat.leb. -+ rewrite <- A. apply exec_loadimm_z_w. apply decompose_int_wf; omega. -+ rewrite <- B. apply exec_loadimm_n_w. apply decompose_int_wf; omega. ++ rewrite <- A. apply exec_loadimm_z_w. apply decompose_int_wf; omega. trivial. ++ rewrite <- B. apply exec_loadimm_n_w. apply decompose_int_wf; omega. trivial. Qed. Lemma exec_loadimm_k_x: forall (rd: ireg) k m l, - wf_decomposition l -> + wf_decomposition l -> + rd <> RA -> forall (rs: regset) accu, rs#rd = Vlong (Int64.repr accu) -> exists rs', @@ -347,9 +392,9 @@ Lemma exec_loadimm_k_x: /\ rs'#rd = Vlong (Int64.repr (recompose_int accu l)) /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r. Proof. - induction 1; intros rs accu ACCU; simpl. + induction 1; intros RD_NOT_RA rs accu ACCU; simpl. - exists rs; split. apply exec_straight_opt_refl. auto. -- destruct (IHwf_decomposition +- destruct (IHwf_decomposition RD_NOT_RA (nextinstr (rs#rd <- (insert_in_long rs#rd n p 16))) (Zinsert accu n p 16)) as (rs' & P & Q & R). @@ -363,19 +408,20 @@ Qed. Lemma exec_loadimm_z_x: forall rd l k rs m, wf_decomposition l -> + rd <> RA -> exists rs', exec_straight ge fn (loadimm_z X rd l k) rs m k rs' m /\ rs'#rd = Vlong (Int64.repr (recompose_int 0 l)) /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r. Proof. - unfold loadimm_z; destruct 1. + unfold loadimm_z; destruct 1; intro RD_NOT_RA. - econstructor; split. apply exec_straight_one. simpl; eauto. auto. split. Simpl. intros; Simpl. - set (accu0 := Zinsert 0 n p 16). set (rs1 := nextinstr (rs#rd <- (Vlong (Int64.repr accu0)))). - destruct (exec_loadimm_k_x rd k m l H1 rs1 accu0) as (rs2 & P & Q & R); auto. + destruct (exec_loadimm_k_x rd k m l H1 RD_NOT_RA rs1 accu0) as (rs2 & P & Q & R); auto. unfold rs1; Simpl. exists rs2; split. eapply exec_straight_opt_step; eauto. @@ -388,12 +434,13 @@ Qed. Lemma exec_loadimm_n_x: forall rd l k rs m, wf_decomposition l -> + rd <> RA -> exists rs', exec_straight ge fn (loadimm_n X rd l k) rs m k rs' m /\ rs'#rd = Vlong (Int64.repr (Z.lnot (recompose_int 0 l))) /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r. Proof. - unfold loadimm_n; destruct 1. + unfold loadimm_n; destruct 1; intro RD_NOT_RA. - econstructor; split. apply exec_straight_one. simpl; eauto. auto. split. Simpl. @@ -402,7 +449,7 @@ Proof. set (rs1 := nextinstr (rs#rd <- (Vlong (Int64.repr accu0)))). destruct (exec_loadimm_k_x rd k m (negate_decomposition l) (negate_decomposition_wf l H1) - rs1 accu0) as (rs2 & P & Q & R). + RD_NOT_RA rs1 accu0) as (rs2 & P & Q & R). unfold rs1; Simpl. exists rs2; split. eapply exec_straight_opt_step; eauto. @@ -415,12 +462,13 @@ Qed. Lemma exec_loadimm64: forall rd n k rs m, + rd <> RA -> 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 -> rs'#r = rs#r. Proof. - unfold loadimm64, loadimm; intros. + unfold loadimm64, loadimm; intros until m; intro RD_NOT_RA. destruct (is_logical_imm64 n). - econstructor; split. apply exec_straight_one. simpl; eauto. auto. @@ -437,8 +485,8 @@ Proof. apply Int64.eqm_samerepr. apply decompose_notint_eqmod. apply Int64.repr_unsigned. } destruct Nat.leb. -+ rewrite <- A. apply exec_loadimm_z_x. apply decompose_int_wf; omega. -+ rewrite <- B. apply exec_loadimm_n_x. apply decompose_int_wf; omega. ++ rewrite <- A. apply exec_loadimm_z_x. apply decompose_int_wf; omega. trivial. ++ rewrite <- B. apply exec_loadimm_n_x. apply decompose_int_wf; omega. trivial. Qed. (** Add immediate *) @@ -450,55 +498,59 @@ Lemma exec_addimm_aux_32: Next (nextinstr (rs#rd <- (sem rs#r1 (Vint (Int.repr n))))) m) -> (forall v n1 n2, sem (sem v (Vint n1)) (Vint n2) = sem v (Vint (Int.add n1 n2))) -> forall rd r1 n k rs m, + (IR RA) <> (preg_of_iregsp (RR1 rd)) -> exists rs', exec_straight ge fn (addimm_aux insn rd r1 (Int.unsigned n) k) rs m k rs' m /\ rs'#rd = sem rs#r1 (Vint n) - /\ forall r, data_preg r = true -> r <> rd -> rs'#r = rs#r. + /\ (forall r, data_preg r = true -> r <> rd -> rs'#r = rs#r) + /\ rs' # RA = rs # RA. Proof. - intros insn sem SEM ASSOC; intros. unfold addimm_aux. + intros insn sem SEM ASSOC; intros until m; intro RD_NOT_RA. unfold addimm_aux. set (nlo := Zzero_ext 12 (Int.unsigned n)). set (nhi := Int.unsigned n - nlo). assert (E: Int.unsigned n = nhi + nlo) by (unfold nhi; omega). rewrite <- (Int.repr_unsigned n). destruct (Z.eqb_spec nhi 0); [|destruct (Z.eqb_spec nlo 0)]. - econstructor; split. apply exec_straight_one. apply SEM. Simpl. split. Simpl. do 3 f_equal; omega. - intros; Simpl. + split; intros; Simpl. - econstructor; split. apply exec_straight_one. apply SEM. Simpl. split. Simpl. do 3 f_equal; omega. - intros; Simpl. + split; intros; Simpl. - econstructor; split. eapply exec_straight_two. apply SEM. apply SEM. Simpl. Simpl. split. Simpl. rewrite ASSOC. do 2 f_equal. apply Int.eqm_samerepr. rewrite E. auto with ints. - intros; Simpl. + split; intros; Simpl. Qed. Lemma exec_addimm32: forall rd r1 n k rs m, r1 <> X16 -> + (IR RA) <> (preg_of_iregsp (RR1 rd)) -> exists rs', exec_straight ge fn (addimm32 rd r1 n k) rs m k rs' m /\ rs'#rd = Val.add rs#r1 (Vint n) - /\ forall r, data_preg r = true -> r <> rd -> rs'#r = rs#r. + /\ (forall r, data_preg r = true -> r <> rd -> rs'#r = rs#r) + /\ rs' # RA = rs # RA. Proof. intros. unfold addimm32. set (nn := Int.neg n). destruct (Int.eq n (Int.zero_ext 24 n)); [| destruct (Int.eq nn (Int.zero_ext 24 nn))]. -- apply exec_addimm_aux_32 with (sem := Val.add). auto. intros; apply Val.add_assoc. +- apply exec_addimm_aux_32 with (sem := Val.add); auto. intros; apply Val.add_assoc. - rewrite <- Val.sub_opp_add. - apply exec_addimm_aux_32 with (sem := Val.sub). auto. + apply exec_addimm_aux_32 with (sem := Val.sub); auto. intros. rewrite ! Val.sub_add_opp, Val.add_assoc. rewrite Int.neg_add_distr. auto. - destruct (Int.lt n Int.zero). + rewrite <- Val.sub_opp_add; fold nn. - edestruct (exec_loadimm32 X16 nn) as (rs1 & A & B & C). + edestruct (exec_loadimm32 X16 nn) as (rs1 & A & B & C). congruence. econstructor; split. eapply exec_straight_trans. eexact A. eapply exec_straight_one. simpl; eauto. auto. split. Simpl. rewrite B, C; eauto with asmgen. - intros; Simpl. -+ edestruct (exec_loadimm32 X16 n) as (rs1 & A & B & C). + split; intros; Simpl. ++ edestruct (exec_loadimm32 X16 n) as (rs1 & A & B & C). congruence. econstructor; split. eapply exec_straight_trans. eexact A. eapply exec_straight_one. simpl; eauto. auto. split. Simpl. rewrite B, C; eauto with asmgen. - intros; Simpl. + split; intros; Simpl. Qed. Lemma exec_addimm_aux_64: @@ -508,10 +560,12 @@ Lemma exec_addimm_aux_64: Next (nextinstr (rs#rd <- (sem rs#r1 (Vlong (Int64.repr n))))) m) -> (forall v n1 n2, sem (sem v (Vlong n1)) (Vlong n2) = sem v (Vlong (Int64.add n1 n2))) -> forall rd r1 n k rs m, + (IR RA) <> (preg_of_iregsp (RR1 rd)) -> exists rs', exec_straight ge fn (addimm_aux insn rd r1 (Int64.unsigned n) k) rs m k rs' m /\ rs'#rd = sem rs#r1 (Vlong n) - /\ forall r, data_preg r = true -> r <> rd -> rs'#r = rs#r. + /\ (forall r, data_preg r = true -> r <> rd -> rs'#r = rs#r) + /\ rs' # RA = rs # RA. Proof. intros insn sem SEM ASSOC; intros. unfold addimm_aux. set (nlo := Zzero_ext 12 (Int64.unsigned n)). set (nhi := Int64.unsigned n - nlo). @@ -520,44 +574,46 @@ Proof. destruct (Z.eqb_spec nhi 0); [|destruct (Z.eqb_spec nlo 0)]. - econstructor; split. apply exec_straight_one. apply SEM. Simpl. split. Simpl. do 3 f_equal; omega. - intros; Simpl. + split; intros; Simpl. - econstructor; split. apply exec_straight_one. apply SEM. Simpl. split. Simpl. do 3 f_equal; omega. - intros; Simpl. + split; intros; Simpl. - econstructor; split. eapply exec_straight_two. apply SEM. apply SEM. Simpl. Simpl. split. Simpl. rewrite ASSOC. do 2 f_equal. apply Int64.eqm_samerepr. rewrite E. auto with ints. - intros; Simpl. + split; intros; Simpl. Qed. Lemma exec_addimm64: forall rd r1 n k rs m, preg_of_iregsp r1 <> X16 -> + (IR RA) <> (preg_of_iregsp (RR1 rd)) -> exists rs', exec_straight ge fn (addimm64 rd r1 n k) rs m k rs' m /\ rs'#rd = Val.addl rs#r1 (Vlong n) - /\ forall r, data_preg r = true -> r <> rd -> rs'#r = rs#r. + /\ (forall r, data_preg r = true -> r <> rd -> rs'#r = rs#r) + /\ rs' # RA = rs # RA. Proof. intros. unfold addimm64. set (nn := Int64.neg n). destruct (Int64.eq n (Int64.zero_ext 24 n)); [| destruct (Int64.eq nn (Int64.zero_ext 24 nn))]. -- apply exec_addimm_aux_64 with (sem := Val.addl). auto. intros; apply Val.addl_assoc. +- apply exec_addimm_aux_64 with (sem := Val.addl); auto. intros; apply Val.addl_assoc. - rewrite <- Val.subl_opp_addl. - apply exec_addimm_aux_64 with (sem := Val.subl). auto. + apply exec_addimm_aux_64 with (sem := Val.subl); auto. intros. rewrite ! Val.subl_addl_opp, Val.addl_assoc. rewrite Int64.neg_add_distr. auto. - destruct (Int64.lt n Int64.zero). + rewrite <- Val.subl_opp_addl; fold nn. - edestruct (exec_loadimm64 X16 nn) as (rs1 & A & B & C). + edestruct (exec_loadimm64 X16 nn) as (rs1 & A & B & C). congruence. econstructor; split. eapply exec_straight_trans. eexact A. eapply exec_straight_one. simpl; eauto. Simpl. split. Simpl. rewrite B, C; eauto with asmgen. simpl. rewrite Int64.shl'_zero. auto. - intros; Simpl. -+ edestruct (exec_loadimm64 X16 n) as (rs1 & A & B & C). + split; intros; Simpl. ++ edestruct (exec_loadimm64 X16 n) as (rs1 & A & B & C). congruence. econstructor; split. eapply exec_straight_trans. eexact A. eapply exec_straight_one. simpl; eauto. Simpl. split. Simpl. rewrite B, C; eauto with asmgen. simpl. rewrite Int64.shl'_zero. auto. - intros; Simpl. + split; intros; Simpl. Qed. (** Logical immediate *) @@ -574,22 +630,25 @@ Lemma exec_logicalimm32: Next (nextinstr (rs#rd <- (sem rs##r1 (eval_shift_op_int rs#r2 s)))) m) -> forall rd r1 n k rs m, r1 <> X16 -> + (IR RA) <> (preg_of_iregsp (RR1 rd)) -> exists rs', exec_straight ge fn (logicalimm32 insn1 insn2 rd r1 n k) rs m k rs' m /\ rs'#rd = sem rs#r1 (Vint n) - /\ forall r, data_preg r = true -> r <> rd -> rs'#r = rs#r. + /\ (forall r, data_preg r = true -> r <> rd -> rs'#r = rs#r) + /\ rs' # RA = rs # RA. Proof. intros until sem; intros SEM1 SEM2; intros. unfold logicalimm32. destruct (is_logical_imm32 n). - econstructor; split. apply exec_straight_one. apply SEM1. reflexivity. - split. Simpl. rewrite Int.repr_unsigned; auto. intros; Simpl. -- edestruct (exec_loadimm32 X16 n) as (rs1 & A & B & C). + split. Simpl. rewrite Int.repr_unsigned; auto. + split; intros; Simpl. +- edestruct (exec_loadimm32 X16 n) as (rs1 & A & B & C). congruence. econstructor; split. eapply exec_straight_trans. eexact A. apply exec_straight_one. apply SEM2. reflexivity. split. Simpl. f_equal; auto. apply C; auto with asmgen. - intros; Simpl. + split; intros; Simpl. Qed. Lemma exec_logicalimm64: @@ -604,50 +663,58 @@ Lemma exec_logicalimm64: Next (nextinstr (rs#rd <- (sem rs###r1 (eval_shift_op_long rs#r2 s)))) m) -> forall rd r1 n k rs m, r1 <> X16 -> + (IR RA) <> (preg_of_iregsp (RR1 rd)) -> exists rs', exec_straight ge fn (logicalimm64 insn1 insn2 rd r1 n k) rs m k rs' m /\ rs'#rd = sem rs#r1 (Vlong n) - /\ forall r, data_preg r = true -> r <> rd -> rs'#r = rs#r. + /\ (forall r, data_preg r = true -> r <> rd -> rs'#r = rs#r) + /\ rs' # RA = rs # RA. Proof. intros until sem; intros SEM1 SEM2; intros. unfold logicalimm64. destruct (is_logical_imm64 n). - econstructor; split. apply exec_straight_one. apply SEM1. reflexivity. - split. Simpl. rewrite Int64.repr_unsigned. auto. intros; Simpl. -- edestruct (exec_loadimm64 X16 n) as (rs1 & A & B & C). + split. Simpl. rewrite Int64.repr_unsigned. auto. + split; intros; Simpl. +- edestruct (exec_loadimm64 X16 n) as (rs1 & A & B & C). congruence. econstructor; split. eapply exec_straight_trans. eexact A. apply exec_straight_one. apply SEM2. reflexivity. split. Simpl. f_equal; auto. apply C; auto with asmgen. - intros; Simpl. + split; intros; Simpl. Qed. (** Load address of symbol *) Lemma exec_loadsymbol: forall rd s ofs k rs m, - rd <> X16 \/ Archi.pic_code tt = false -> + rd <> X16 \/ Archi.pic_code tt = false -> + (IR RA) <> (preg_of_iregsp (RR1 rd)) -> exists rs', exec_straight ge fn (loadsymbol rd s ofs k) rs m k rs' m /\ rs'#rd = Genv.symbol_address ge s ofs - /\ forall r, data_preg r = true -> r <> rd -> rs'#r = rs#r. + /\ (forall r, data_preg r = true -> r <> rd -> rs'#r = rs#r) + /\ rs'#RA = rs#RA. Proof. unfold loadsymbol; intros. destruct (Archi.pic_code tt). - predSpec Ptrofs.eq Ptrofs.eq_spec ofs Ptrofs.zero. + subst ofs. econstructor; split. apply exec_straight_one; [simpl; eauto | reflexivity]. - split. Simpl. intros; Simpl. + split. Simpl. split; intros; Simpl. + + exploit exec_addimm64. instantiate (1 := rd). simpl. destruct H; congruence. - intros (rs1 & A & B & C). + instantiate (1 := rd). assumption. + intros (rs1 & A & B & C & D). econstructor; split. econstructor. simpl; eauto. auto. eexact A. split. simpl in B; rewrite B. Simpl. rewrite <- Genv.shift_symbol_address_64 by auto. rewrite Ptrofs.add_zero_l, Ptrofs.of_int64_to_int64 by auto. auto. - intros. rewrite C by auto. Simpl. + split; intros. rewrite C by auto; Simpl. + rewrite D. Simpl. - econstructor; split. eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto. split. Simpl. rewrite symbol_high_low; auto. - intros; Simpl. + split; intros; Simpl. Qed. (** Shifted operands *) @@ -756,23 +823,25 @@ Lemma exec_arith_extended: Next (nextinstr (rs#rd <- (sem rs###r1 (eval_shift_op_long rs#r2 s)))) m) -> forall (rd r1 r2: ireg) (ex: extension) (a: amount64) (k: code) rs m, r1 <> X16 -> + (IR RA) <> (preg_of_iregsp (RR1 rd)) -> exists rs', exec_straight ge fn (arith_extended insnX insnS rd r1 r2 ex a k) rs m k rs' m /\ rs'#rd = sem rs#r1 (Op.eval_extend ex rs#r2 a) - /\ forall r, data_preg r = true -> r <> rd -> rs'#r = rs#r. + /\ (forall r, data_preg r = true -> r <> rd -> rs'#r = rs#r) + /\ rs' # RA = rs # RA. Proof. intros sem insnX insnS EX ES; intros. unfold arith_extended. destruct (Int.ltu a (Int.repr 5)). - econstructor; split. apply exec_straight_one. rewrite EX; eauto. auto. split. Simpl. f_equal. destruct ex; auto. - intros; Simpl. + split; intros; Simpl. - exploit (exec_move_extended_base X16 r2 ex). intros (rs' & A & B & C). econstructor; split. eapply exec_straight_trans. eexact A. apply exec_straight_one. rewrite ES. eauto. auto. split. Simpl. unfold ir0x. rewrite C by eauto with asmgen. f_equal. rewrite B. destruct ex; auto. - intros; Simpl. + split; intros; Simpl. Qed. (** Extended right shift *) @@ -780,41 +849,49 @@ Qed. Lemma exec_shrx32: forall (rd r1: ireg) (n: int) k v (rs: regset) m, Val.shrx rs#r1 (Vint n) = Some v -> r1 <> X16 -> + (IR RA) <> (preg_of_iregsp (RR1 rd)) -> exists rs', exec_straight ge fn (shrx32 rd r1 n k) rs m k rs' m /\ rs'#rd = v - /\ forall r, data_preg r = true -> r <> rd -> rs'#r = rs#r. + /\ (forall r, data_preg r = true -> r <> rd -> rs'#r = rs#r) + /\ rs' # RA = rs # RA. Proof. unfold shrx32; intros. apply Val.shrx_shr_2 in H. destruct (Int.eq n Int.zero) eqn:E. - econstructor; split. apply exec_straight_one; [simpl;eauto|auto]. - split. Simpl. subst v; auto. intros; Simpl. + split. Simpl. subst v; auto. + split; intros; Simpl. - econstructor; split. eapply exec_straight_three. unfold exec_instr. rewrite or_zero_eval_shift_op_int by congruence. eauto. simpl; eauto. unfold exec_instr. rewrite or_zero_eval_shift_op_int by congruence. eauto. auto. auto. auto. - split. subst v; Simpl. intros; Simpl. + split. subst v; Simpl. + split; intros; Simpl. Qed. Lemma exec_shrx64: forall (rd r1: ireg) (n: int) k v (rs: regset) m, Val.shrxl rs#r1 (Vint n) = Some v -> r1 <> X16 -> + (IR RA) <> (preg_of_iregsp (RR1 rd)) -> exists rs', exec_straight ge fn (shrx64 rd r1 n k) rs m k rs' m /\ rs'#rd = v - /\ forall r, data_preg r = true -> r <> rd -> rs'#r = rs#r. + /\ (forall r, data_preg r = true -> r <> rd -> rs'#r = rs#r) + /\ rs' # RA = rs # RA. Proof. unfold shrx64; intros. apply Val.shrxl_shrl_2 in H. destruct (Int.eq n Int.zero) eqn:E. - econstructor; split. apply exec_straight_one; [simpl;eauto|auto]. - split. Simpl. subst v; auto. intros; Simpl. + split. Simpl. subst v; auto. + split; intros; Simpl. - econstructor; split. eapply exec_straight_three. unfold exec_instr. rewrite or_zero_eval_shift_op_long by congruence. eauto. simpl; eauto. unfold exec_instr. rewrite or_zero_eval_shift_op_long by congruence. eauto. auto. auto. auto. - split. subst v; Simpl. intros; Simpl. + split. subst v; Simpl. + split; intros; Simpl. Qed. (** Condition bits *) @@ -1070,6 +1147,56 @@ Ltac ArgsInv := | [ H: freg_of _ = OK _ |- _ ] => simpl in *; rewrite (freg_of_eq _ _ H) in * end). +Lemma compare_int_RA: + forall rs a b m, + compare_int rs a b m X30 = rs X30. +Proof. + unfold compare_int. + intros. + repeat rewrite Pregmap.gso by congruence. + trivial. +Qed. + +Hint Resolve compare_int_RA : asmgen. + +Lemma compare_long_RA: + forall rs a b m, + compare_long rs a b m X30 = rs X30. +Proof. + unfold compare_long. + intros. + repeat rewrite Pregmap.gso by congruence. + trivial. +Qed. + +Hint Resolve compare_long_RA : asmgen. + +Lemma compare_float_RA: + forall rs a b, + compare_float rs a b X30 = rs X30. +Proof. + unfold compare_float. + intros. + destruct a; destruct b. + all: repeat rewrite Pregmap.gso by congruence; trivial. +Qed. + +Hint Resolve compare_float_RA : asmgen. + + +Lemma compare_single_RA: + forall rs a b, + compare_single rs a b X30 = rs X30. +Proof. + unfold compare_single. + intros. + destruct a; destruct b. + all: repeat rewrite Pregmap.gso by congruence; trivial. +Qed. + +Hint Resolve compare_single_RA : asmgen. + + Lemma transl_cond_correct: forall cond args k c rs m, transl_cond cond args k = OK c -> @@ -1078,185 +1205,218 @@ Lemma transl_cond_correct: /\ (forall b, eval_condition cond (map rs (map preg_of args)) m = Some b -> eval_testcond (cond_for_cond cond) rs' = Some b) - /\ forall r, data_preg r = true -> rs'#r = rs#r. + /\ (forall r, data_preg r = true -> rs'#r = rs#r) + /\ rs' # RA = rs # RA. Proof. intros until m; intros TR. destruct cond; simpl in TR; ArgsInv. - (* Ccomp *) econstructor; split. apply exec_straight_one. simpl; eauto. auto. - split; intros. apply eval_testcond_compare_sint; auto. + repeat split; intros. apply eval_testcond_compare_sint; auto. destruct r; reflexivity || discriminate. - (* Ccompu *) econstructor; split. apply exec_straight_one. simpl; eauto. auto. - split; intros. apply eval_testcond_compare_uint; auto. + repeat split; intros. apply eval_testcond_compare_uint; auto. destruct r; reflexivity || discriminate. - (* Ccompimm *) destruct (is_arith_imm32 n); [|destruct (is_arith_imm32 (Int.neg n))]. + econstructor; split. apply exec_straight_one. simpl; eauto. auto. - split; intros. rewrite Int.repr_unsigned. apply eval_testcond_compare_sint; auto. + repeat split; intros. rewrite Int.repr_unsigned. apply eval_testcond_compare_sint; auto. destruct r; reflexivity || discriminate. + econstructor; split. apply exec_straight_one. simpl. rewrite Int.repr_unsigned, Int.neg_involutive. eauto. auto. - split; intros. apply eval_testcond_compare_sint; auto. + repeat split; intros. apply eval_testcond_compare_sint; auto. destruct r; reflexivity || discriminate. -+ exploit (exec_loadimm32 X16 n). intros (rs' & A & B & C). ++ exploit (exec_loadimm32 X16 n). congruence. intros (rs' & A & B & C). econstructor; split. eapply exec_straight_trans. eexact A. apply exec_straight_one. simpl. rewrite B, C by eauto with asmgen. eauto. auto. - split; intros. apply eval_testcond_compare_sint; auto. - transitivity (rs' r). destruct r; reflexivity || discriminate. auto with asmgen. + repeat split; intros. apply eval_testcond_compare_sint; auto. + transitivity (rs' r). destruct r; reflexivity || discriminate. + auto with asmgen. + Simpl. rewrite compare_int_RA. + apply C; congruence. - (* Ccompuimm *) destruct (is_arith_imm32 n); [|destruct (is_arith_imm32 (Int.neg n))]. + econstructor; split. apply exec_straight_one. simpl; eauto. auto. - split; intros. rewrite Int.repr_unsigned. apply eval_testcond_compare_uint; auto. + repeat split; intros. rewrite Int.repr_unsigned. apply eval_testcond_compare_uint; auto. destruct r; reflexivity || discriminate. + econstructor; split. apply exec_straight_one. simpl. rewrite Int.repr_unsigned, Int.neg_involutive. eauto. auto. - split; intros. apply eval_testcond_compare_uint; auto. + repeat split; intros. apply eval_testcond_compare_uint; auto. destruct r; reflexivity || discriminate. -+ exploit (exec_loadimm32 X16 n). intros (rs' & A & B & C). ++ exploit (exec_loadimm32 X16 n). congruence. intros (rs' & A & B & C). econstructor; split. eapply exec_straight_trans. eexact A. apply exec_straight_one. simpl. rewrite B, C by eauto with asmgen. eauto. auto. - split; intros. apply eval_testcond_compare_uint; auto. + repeat split; intros. apply eval_testcond_compare_uint; auto. transitivity (rs' r). destruct r; reflexivity || discriminate. auto with asmgen. + Simpl. rewrite compare_int_RA. + apply C; congruence. - (* Ccompshift *) econstructor; split. apply exec_straight_one. simpl; eauto. auto. - split; intros. rewrite transl_eval_shift. apply eval_testcond_compare_sint; auto. + repeat split; intros. rewrite transl_eval_shift. apply eval_testcond_compare_sint; auto. destruct r; reflexivity || discriminate. - (* Ccompushift *) econstructor; split. apply exec_straight_one. simpl; eauto. auto. - split; intros. rewrite transl_eval_shift. apply eval_testcond_compare_uint; auto. + repeat split; intros. rewrite transl_eval_shift. apply eval_testcond_compare_uint; auto. destruct r; reflexivity || discriminate. - (* Cmaskzero *) destruct (is_logical_imm32 n). + econstructor; split. apply exec_straight_one. simpl; eauto. auto. - split; intros. rewrite Int.repr_unsigned. apply (eval_testcond_compare_sint Ceq); auto. + repeat split; intros. rewrite Int.repr_unsigned. apply (eval_testcond_compare_sint Ceq); auto. destruct r; reflexivity || discriminate. -+ exploit (exec_loadimm32 X16 n). intros (rs' & A & B & C). ++ exploit (exec_loadimm32 X16 n). congruence. intros (rs' & A & B & C). econstructor; split. eapply exec_straight_trans. eexact A. apply exec_straight_one. simpl. rewrite B, C by eauto with asmgen. eauto. auto. - split; intros. apply (eval_testcond_compare_sint Ceq); auto. + repeat split; intros. apply (eval_testcond_compare_sint Ceq); auto. transitivity (rs' r). destruct r; reflexivity || discriminate. auto with asmgen. + Simpl. rewrite compare_int_RA. + apply C; congruence. + - (* Cmasknotzero *) destruct (is_logical_imm32 n). + econstructor; split. apply exec_straight_one. simpl; eauto. auto. - split; intros. rewrite Int.repr_unsigned. apply (eval_testcond_compare_sint Cne); auto. + repeat split; intros. rewrite Int.repr_unsigned. apply (eval_testcond_compare_sint Cne); auto. destruct r; reflexivity || discriminate. -+ exploit (exec_loadimm32 X16 n). intros (rs' & A & B & C). + ++ exploit (exec_loadimm32 X16 n). congruence. intros (rs' & A & B & C). econstructor; split. eapply exec_straight_trans. eexact A. apply exec_straight_one. simpl. rewrite B, C by eauto with asmgen. eauto. auto. - split; intros. apply (eval_testcond_compare_sint Cne); auto. + repeat split; intros. apply (eval_testcond_compare_sint Cne); auto. transitivity (rs' r). destruct r; reflexivity || discriminate. auto with asmgen. + Simpl. rewrite compare_int_RA. + apply C; congruence. + - (* Ccompl *) econstructor; split. apply exec_straight_one. simpl; eauto. auto. - split; intros. apply eval_testcond_compare_slong; auto. + repeat split; intros. apply eval_testcond_compare_slong; auto. destruct r; reflexivity || discriminate. - (* Ccomplu *) econstructor; split. apply exec_straight_one. simpl; eauto. auto. - split; intros. apply eval_testcond_compare_ulong; auto. + repeat split; intros. apply eval_testcond_compare_ulong; auto. destruct r; reflexivity || discriminate. - (* Ccomplimm *) destruct (is_arith_imm64 n); [|destruct (is_arith_imm64 (Int64.neg n))]. + econstructor; split. apply exec_straight_one. simpl; eauto. auto. - split; intros. rewrite Int64.repr_unsigned. apply eval_testcond_compare_slong; auto. + repeat split; intros. rewrite Int64.repr_unsigned. apply eval_testcond_compare_slong; auto. destruct r; reflexivity || discriminate. + econstructor; split. apply exec_straight_one. simpl. rewrite Int64.repr_unsigned, Int64.neg_involutive. eauto. auto. - split; intros. apply eval_testcond_compare_slong; auto. + repeat split; intros. apply eval_testcond_compare_slong; auto. destruct r; reflexivity || discriminate. -+ exploit (exec_loadimm64 X16 n). intros (rs' & A & B & C). ++ exploit (exec_loadimm64 X16 n). congruence. intros (rs' & A & B & C). econstructor; split. eapply exec_straight_trans. eexact A. apply exec_straight_one. simpl. rewrite B, C by eauto with asmgen. eauto. auto. - split; intros. apply eval_testcond_compare_slong; auto. + repeat split; intros. apply eval_testcond_compare_slong; auto. transitivity (rs' r). destruct r; reflexivity || discriminate. auto with asmgen. + Simpl. rewrite compare_long_RA. + apply C; congruence. + - (* Ccompluimm *) destruct (is_arith_imm64 n); [|destruct (is_arith_imm64 (Int64.neg n))]. + econstructor; split. apply exec_straight_one. simpl; eauto. auto. - split; intros. rewrite Int64.repr_unsigned. apply eval_testcond_compare_ulong; auto. + repeat split; intros. rewrite Int64.repr_unsigned. apply eval_testcond_compare_ulong; auto. destruct r; reflexivity || discriminate. + econstructor; split. apply exec_straight_one. simpl. rewrite Int64.repr_unsigned, Int64.neg_involutive. eauto. auto. - split; intros. apply eval_testcond_compare_ulong; auto. + repeat split; intros. apply eval_testcond_compare_ulong; auto. destruct r; reflexivity || discriminate. -+ exploit (exec_loadimm64 X16 n). intros (rs' & A & B & C). ++ exploit (exec_loadimm64 X16 n). congruence. intros (rs' & A & B & C). econstructor; split. eapply exec_straight_trans. eexact A. apply exec_straight_one. simpl. rewrite B, C by eauto with asmgen. eauto. auto. - split; intros. apply eval_testcond_compare_ulong; auto. + repeat split; intros. apply eval_testcond_compare_ulong; auto. transitivity (rs' r). destruct r; reflexivity || discriminate. auto with asmgen. + Simpl. rewrite compare_long_RA. + apply C; congruence. + - (* Ccomplshift *) econstructor; split. apply exec_straight_one. simpl; eauto. auto. - split; intros. rewrite transl_eval_shiftl. apply eval_testcond_compare_slong; auto. + repeat split; intros. rewrite transl_eval_shiftl. apply eval_testcond_compare_slong; auto. destruct r; reflexivity || discriminate. - (* Ccomplushift *) econstructor; split. apply exec_straight_one. simpl; eauto. auto. - split; intros. rewrite transl_eval_shiftl. apply eval_testcond_compare_ulong; auto. + repeat split; intros. rewrite transl_eval_shiftl. apply eval_testcond_compare_ulong; auto. destruct r; reflexivity || discriminate. - (* Cmasklzero *) destruct (is_logical_imm64 n). + econstructor; split. apply exec_straight_one. simpl; eauto. auto. - split; intros. rewrite Int64.repr_unsigned. apply (eval_testcond_compare_slong Ceq); auto. + repeat split; intros. rewrite Int64.repr_unsigned. apply (eval_testcond_compare_slong Ceq); auto. destruct r; reflexivity || discriminate. -+ exploit (exec_loadimm64 X16 n). intros (rs' & A & B & C). ++ exploit (exec_loadimm64 X16 n). congruence. intros (rs' & A & B & C). econstructor; split. eapply exec_straight_trans. eexact A. apply exec_straight_one. simpl. rewrite B, C by eauto with asmgen. eauto. auto. - split; intros. apply (eval_testcond_compare_slong Ceq); auto. + repeat split; intros. apply (eval_testcond_compare_slong Ceq); auto. transitivity (rs' r). destruct r; reflexivity || discriminate. auto with asmgen. + Simpl. rewrite compare_long_RA. + apply C; congruence. + - (* Cmasknotzero *) destruct (is_logical_imm64 n). + econstructor; split. apply exec_straight_one. simpl; eauto. auto. - split; intros. rewrite Int64.repr_unsigned. apply (eval_testcond_compare_slong Cne); auto. + repeat split; intros. rewrite Int64.repr_unsigned. apply (eval_testcond_compare_slong Cne); auto. destruct r; reflexivity || discriminate. -+ exploit (exec_loadimm64 X16 n). intros (rs' & A & B & C). ++ exploit (exec_loadimm64 X16 n). congruence. intros (rs' & A & B & C). econstructor; split. eapply exec_straight_trans. eexact A. apply exec_straight_one. simpl. rewrite B, C by eauto with asmgen. eauto. auto. - split; intros. apply (eval_testcond_compare_slong Cne); auto. + repeat split; intros. apply (eval_testcond_compare_slong Cne); auto. transitivity (rs' r). destruct r; reflexivity || discriminate. auto with asmgen. + Simpl. rewrite compare_long_RA. + apply C; congruence. + - (* Ccompf *) econstructor; split. apply exec_straight_one. simpl; eauto. rewrite compare_float_inv; auto. - split; intros. apply eval_testcond_compare_float; auto. + repeat split; intros. apply eval_testcond_compare_float; auto. destruct r; discriminate || rewrite compare_float_inv; auto. + Simpl. - (* Cnotcompf *) econstructor; split. apply exec_straight_one. simpl; eauto. rewrite compare_float_inv; auto. - split; intros. apply eval_testcond_compare_not_float; auto. + repeat split; intros. apply eval_testcond_compare_not_float; auto. destruct r; discriminate || rewrite compare_float_inv; auto. + Simpl. - (* Ccompfzero *) econstructor; split. apply exec_straight_one. simpl; eauto. rewrite compare_float_inv; auto. - split; intros. apply eval_testcond_compare_float; auto. + repeat split; intros. apply eval_testcond_compare_float; auto. destruct r; discriminate || rewrite compare_float_inv; auto. + Simpl. - (* Cnotcompfzero *) econstructor; split. apply exec_straight_one. simpl; eauto. rewrite compare_float_inv; auto. - split; intros. apply eval_testcond_compare_not_float; auto. + repeat split; intros. apply eval_testcond_compare_not_float; auto. destruct r; discriminate || rewrite compare_float_inv; auto. + Simpl. - (* Ccompfs *) econstructor; split. apply exec_straight_one. simpl; eauto. rewrite compare_single_inv; auto. - split; intros. apply eval_testcond_compare_single; auto. + repeat split; intros. apply eval_testcond_compare_single; auto. destruct r; discriminate || rewrite compare_single_inv; auto. + Simpl. - (* Cnotcompfs *) econstructor; split. apply exec_straight_one. simpl; eauto. rewrite compare_single_inv; auto. - split; intros. apply eval_testcond_compare_not_single; auto. + repeat split; intros. apply eval_testcond_compare_not_single; auto. destruct r; discriminate || rewrite compare_single_inv; auto. + Simpl. - (* Ccompfszero *) econstructor; split. apply exec_straight_one. simpl; eauto. rewrite compare_single_inv; auto. - split; intros. apply eval_testcond_compare_single; auto. + repeat split; intros. apply eval_testcond_compare_single; auto. destruct r; discriminate || rewrite compare_single_inv; auto. + Simpl. - (* Cnotcompfszero *) econstructor; split. apply exec_straight_one. simpl; eauto. rewrite compare_single_inv; auto. - split; intros. apply eval_testcond_compare_not_single; auto. + repeat split; intros. apply eval_testcond_compare_not_single; auto. destruct r; discriminate || rewrite compare_single_inv; auto. + Simpl. Qed. (** Translation of conditional branches *) @@ -1379,7 +1539,7 @@ Ltac TranslOpSimpl := | split; [ rewrite ? transl_eval_shift, ? transl_eval_shiftl; apply Val.lessdef_same; Simpl; fail | split; [ intros; Simpl; fail - | intros; Simpl; eapply RA_not_written2; eauto] ]]. + | intros; Simpl; eauto with asmgen; fail] ]]. Ltac TranslOpBase := econstructor; split; @@ -1405,11 +1565,19 @@ Local Opaque Int.eq Int64.eq Val.add Val.addl Int.zwordsize Int64.zwordsize. destruct (preg_of res) eqn:RR; try discriminate; destruct (preg_of m0) eqn:R1; inv TR. all: TranslOpSimpl. - (* intconst *) - exploit exec_loadimm32. intros (rs' & A & B & C). - exists rs'; split. eexact A. split. rewrite B; auto. intros; auto with asmgen. + exploit exec_loadimm32. apply (ireg_of_not_RA res); eassumption. + intros (rs' & A & B & C). + exists rs'; split. eexact A. split. rewrite B; auto. + split. intros; auto with asmgen. + apply C. congruence. + eapply ireg_of_not_RA''; eauto. - (* longconst *) - exploit exec_loadimm64. intros (rs' & A & B & C). - exists rs'; split. eexact A. split. rewrite B; auto. intros; auto with asmgen. + exploit exec_loadimm64. apply (ireg_of_not_RA res); eassumption. + intros (rs' & A & B & C). + exists rs'; split. eexact A. split. rewrite B; auto. + split. intros; auto with asmgen. + apply C. congruence. + eapply ireg_of_not_RA''; eauto. - (* floatconst *) destruct (Float.eq_dec n Float.zero). + subst n. TranslOpSimpl. @@ -1419,11 +1587,15 @@ Local Opaque Int.eq Int64.eq Val.add Val.addl Int.zwordsize Int64.zwordsize. + subst n. TranslOpSimpl. + TranslOpSimpl. - (* loadsymbol *) - exploit (exec_loadsymbol x id ofs). eauto with asmgen. intros (rs' & A & B & C). - exists rs'; split. eexact A. split. rewrite B; auto. auto. + exploit (exec_loadsymbol x id ofs). eauto with asmgen. + apply (ireg_of_not_RA'' res); eassumption. + intros (rs' & A & B & C & D). + exists rs'; split. eexact A. split. rewrite B; auto. + split; auto. - (* addrstack *) exploit (exec_addimm64 x XSP (Ptrofs.to_int64 ofs)). simpl; eauto with asmgen. - intros (rs' & A & B & C). + apply (ireg_of_not_RA'' res); eassumption. + intros (rs' & A & B & C & D). exists rs'; split. eexact A. split. simpl in B; rewrite B. Local Transparent Val.addl. destruct (rs SP); simpl; auto. rewrite Ptrofs.of_int64_to_int64 by auto. auto. @@ -1431,7 +1603,8 @@ Local Transparent Val.addl. - (* shift *) rewrite <- transl_eval_shift'. TranslOpSimpl. - (* addimm *) - exploit (exec_addimm32 x x0 n). eauto with asmgen. intros (rs' & A & B & C). + exploit (exec_addimm32 x x0 n). eauto with asmgen. eapply ireg_of_not_RA''; eassumption. + intros (rs' & A & B & C & D). exists rs'; split. eexact A. split. rewrite B; auto. auto. - (* mul *) TranslOpBase. @@ -1439,18 +1612,20 @@ Local Transparent Val.add. destruct (rs x0); auto; destruct (rs x1); auto. simpl. rewrite Int.add_zero_l; auto. - (* andimm *) exploit (exec_logicalimm32 (Pandimm W) (Pand W)). - intros; reflexivity. intros; reflexivity. instantiate (1 := x0). eauto with asmgen. - intros (rs' & A & B & C). - exists rs'; split. eexact A. split. rewrite B; auto. auto. + intros; reflexivity. intros; reflexivity. instantiate (1 := x0). eauto with asmgen. apply (ireg_of_not_RA'' res); eassumption. + intros (rs' & A & B & C & D). + exists rs'; split. eexact A. split. rewrite B; auto. + split; auto. - (* orimm *) exploit (exec_logicalimm32 (Porrimm W) (Porr W)). - intros; reflexivity. intros; reflexivity. instantiate (1 := x0). eauto with asmgen. - intros (rs' & A & B & C). - exists rs'; split. eexact A. split. rewrite B; auto. auto. + intros; reflexivity. intros; reflexivity. instantiate (1 := x0). eauto with asmgen. apply (ireg_of_not_RA'' res); eassumption. + intros (rs' & A & B & C & D). + exists rs'; split. eexact A. split. rewrite B; auto. + split; auto. - (* xorimm *) exploit (exec_logicalimm32 (Peorimm W) (Peor W)). - intros; reflexivity. intros; reflexivity. instantiate (1 := x0). eauto with asmgen. - intros (rs' & A & B & C). + intros; reflexivity. intros; reflexivity. instantiate (1 := x0). eauto with asmgen. apply (ireg_of_not_RA'' res); eassumption. + intros (rs' & A & B & C & D). exists rs'; split. eexact A. split. rewrite B; auto. auto. - (* not *) TranslOpBase. @@ -1459,8 +1634,10 @@ Local Transparent Val.add. TranslOpBase. destruct (eval_shift s (rs x0) a); auto. simpl. rewrite Int.or_zero_l; auto. - (* shrx *) - exploit (exec_shrx32 x x0 n); eauto with asmgen. intros (rs' & A & B & C). - econstructor; split. eexact A. split. rewrite B; auto. auto. + exploit (exec_shrx32 x x0 n); eauto with asmgen. apply (ireg_of_not_RA'' res); eassumption. + intros (rs' & A & B & C & D). + econstructor; split. eexact A. split. rewrite B; auto. + split; auto. - (* zero-ext *) TranslOpBase. destruct (rs x0); auto; simpl. rewrite Int.shl_zero. auto. @@ -1484,36 +1661,47 @@ Local Transparent Val.add. - (* extend *) exploit (exec_move_extended x0 x1 x a k). intros (rs' & A & B & C). econstructor; split. eexact A. - split. rewrite B; auto. eauto with asmgen. + split. rewrite B; auto. + split; eauto with asmgen. - (* addext *) exploit (exec_arith_extended Val.addl Paddext (Padd X)). - auto. auto. instantiate (1 := x1). eauto with asmgen. intros (rs' & A & B & C). - econstructor; split. eexact A. split. rewrite B; auto. auto. + auto. auto. instantiate (1 := x1). eauto with asmgen. + apply (ireg_of_not_RA'' res); eassumption. + intros (rs' & A & B & C & D). + econstructor; split. eexact A. split. rewrite B; auto. + split; auto. - (* addlimm *) exploit (exec_addimm64 x x0 n). simpl. generalize (ireg_of_not_X16 _ _ EQ1). congruence. - intros (rs' & A & B & C). + apply (ireg_of_not_RA'' res); eassumption. + intros (rs' & A & B & C & D). exists rs'; split. eexact A. split. simpl in B; rewrite B; auto. auto. - (* subext *) exploit (exec_arith_extended Val.subl Psubext (Psub X)). - auto. auto. instantiate (1 := x1). eauto with asmgen. intros (rs' & A & B & C). - econstructor; split. eexact A. split. rewrite B; auto. auto. + auto. auto. instantiate (1 := x1). eauto with asmgen. + apply (ireg_of_not_RA'' res); eassumption. + intros (rs' & A & B & C & D). + econstructor; split. eexact A. split. rewrite B; auto. + split; auto. - (* mull *) TranslOpBase. destruct (rs x0); auto; destruct (rs x1); auto. simpl. rewrite Int64.add_zero_l; auto. - (* andlimm *) exploit (exec_logicalimm64 (Pandimm X) (Pand X)). intros; reflexivity. intros; reflexivity. instantiate (1 := x0). eauto with asmgen. - intros (rs' & A & B & C). + apply (ireg_of_not_RA'' res); eassumption. + intros (rs' & A & B & C & D). exists rs'; split. eexact A. split. rewrite B; auto. auto. - (* orlimm *) exploit (exec_logicalimm64 (Porrimm X) (Porr X)). intros; reflexivity. intros; reflexivity. instantiate (1 := x0). eauto with asmgen. - intros (rs' & A & B & C). + apply (ireg_of_not_RA'' res); eassumption. + intros (rs' & A & B & C & D). exists rs'; split. eexact A. split. rewrite B; auto. auto. - (* xorlimm *) exploit (exec_logicalimm64 (Peorimm X) (Peor X)). intros; reflexivity. intros; reflexivity. instantiate (1 := x0). eauto with asmgen. - intros (rs' & A & B & C). + apply (ireg_of_not_RA'' res); eassumption. + intros (rs' & A & B & C & D). exists rs'; split. eexact A. split. rewrite B; auto. auto. - (* notl *) TranslOpBase. @@ -1522,7 +1710,8 @@ Local Transparent Val.add. TranslOpBase. destruct (eval_shiftl s (rs x0) a); auto. simpl. rewrite Int64.or_zero_l; auto. - (* shrx *) - exploit (exec_shrx64 x x0 n); eauto with asmgen. intros (rs' & A & B & C). + exploit (exec_shrx64 x x0 n); eauto with asmgen. + apply (ireg_of_not_RA'' res); eassumption. intros (rs' & A & B & C & D ). econstructor; split. eexact A. split. rewrite B; auto. auto. - (* zero-ext-l *) TranslOpBase. -- cgit From 6d5718dae13e8db5fc85feb86395dd3dc785dfda Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 24 Mar 2020 23:46:19 +0100 Subject: transl_op_correct --- aarch64/Asmgenproof1.v | 24 ++++++++++++++---------- 1 file changed, 14 insertions(+), 10 deletions(-) diff --git a/aarch64/Asmgenproof1.v b/aarch64/Asmgenproof1.v index 96561da7..1471ee4f 100644 --- a/aarch64/Asmgenproof1.v +++ b/aarch64/Asmgenproof1.v @@ -1429,7 +1429,8 @@ Lemma transl_cond_branch_correct: exec_straight_opt ge fn c rs m (insn :: k) rs' m /\ exec_instr ge fn insn rs' m = (if b then goto_label fn lbl rs' m else Next (nextinstr rs') m) - /\ forall r, data_preg r = true -> rs'#r = rs#r. + /\ (forall r, data_preg r = true -> rs'#r = rs#r) + /\ rs' # RA = rs # RA. Proof. intros until b; intros TR EV. assert (DFL: @@ -1438,13 +1439,14 @@ Proof. exec_straight_opt ge fn c rs m (insn :: k) rs' m /\ exec_instr ge fn insn rs' m = (if b then goto_label fn lbl rs' m else Next (nextinstr rs') m) - /\ forall r, data_preg r = true -> rs'#r = rs#r). + /\ (forall r, data_preg r = true -> rs'#r = rs#r) + /\ rs' # RA = rs # RA ). { unfold transl_cond_branch_default; intros. - exploit transl_cond_correct; eauto. intros (rs' & A & B & C). + exploit transl_cond_correct; eauto. intros (rs' & A & B & C & D). exists rs', (Pbc (cond_for_cond cond) lbl); split. apply exec_straight_opt_intro. eexact A. - split; auto. simpl. rewrite (B b) by auto. auto. + repeat split; auto. simpl. rewrite (B b) by auto. auto. } Local Opaque transl_cond transl_cond_branch_default. destruct args as [ | a1 args]; simpl in TR; auto. @@ -1732,35 +1734,37 @@ Local Transparent Val.add. TranslOpBase. destruct (rs x0); simpl; auto. rewrite ! a64_range; simpl. rewrite <- Int64.sign_ext_shr'_min; auto using a64_range. - (* condition *) - exploit (transl_cond_correct cond args); eauto. intros (rs' & A & B & C). + exploit (transl_cond_correct cond args); eauto. intros (rs' & A & B & C & D). econstructor; split. eapply exec_straight_trans. eexact A. apply exec_straight_one. simpl; eauto. auto. split. Simpl. destruct (eval_condition cond (map rs (map preg_of args)) m) as [b|]; simpl in *. rewrite (B b) by auto. auto. auto. - intros; Simpl. + split; intros; Simpl. - (* select *) destruct (preg_of res) eqn:RES; monadInv TR. + (* integer *) generalize (ireg_of_eq _ _ EQ) (ireg_of_eq _ _ EQ1); intros E1 E2; rewrite E1, E2. - exploit (transl_cond_correct cond args); eauto. intros (rs' & A & B & C). + exploit (transl_cond_correct cond args); eauto. intros (rs' & A & B & C & D). econstructor; split. eapply exec_straight_trans. eexact A. apply exec_straight_one. simpl; eauto. auto. split. Simpl. destruct (eval_condition cond (map rs (map preg_of args)) m) as [b|]; simpl in *. rewrite (B b) by auto. rewrite !C. apply Val.lessdef_normalize. rewrite <- E2; auto with asmgen. rewrite <- E1; auto with asmgen. auto. - intros; Simpl. + split; intros; Simpl. + rewrite <- D. + eapply RA_not_written2; eassumption. + (* FP *) generalize (freg_of_eq _ _ EQ) (freg_of_eq _ _ EQ1); intros E1 E2; rewrite E1, E2. - exploit (transl_cond_correct cond args); eauto. intros (rs' & A & B & C). + exploit (transl_cond_correct cond args); eauto. intros (rs' & A & B & C & D). econstructor; split. eapply exec_straight_trans. eexact A. apply exec_straight_one. simpl; eauto. auto. split. Simpl. destruct (eval_condition cond (map rs (map preg_of args)) m) as [b|]; simpl in *. rewrite (B b) by auto. rewrite !C. apply Val.lessdef_normalize. rewrite <- E2; auto with asmgen. rewrite <- E1; auto with asmgen. auto. - intros; Simpl. + split; intros; Simpl. Qed. (** Translation of addressing modes, loads, stores *) -- cgit From 2729bdeb4aced04a1301d5696574ff8610072395 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 24 Mar 2020 23:52:22 +0100 Subject: transl_addressing_correct --- aarch64/Asmgenproof1.v | 13 +++++++++---- 1 file changed, 9 insertions(+), 4 deletions(-) diff --git a/aarch64/Asmgenproof1.v b/aarch64/Asmgenproof1.v index 1471ee4f..4da1b52b 100644 --- a/aarch64/Asmgenproof1.v +++ b/aarch64/Asmgenproof1.v @@ -1784,7 +1784,7 @@ Proof. destruct (offset_representable sz ofs); inv EQ0. + econstructor; econstructor; split. apply exec_straight_opt_refl. auto. -+ exploit (exec_loadimm64 X16 ofs). intros (rs' & A & B & C). ++ exploit (exec_loadimm64 X16 ofs). congruence. intros (rs' & A & B & C). econstructor; exists rs'; split. apply exec_straight_opt_intro; eexact A. split. simpl. rewrite B, C by eauto with asmgen. auto. eauto with asmgen. @@ -1810,7 +1810,8 @@ Proof. split; auto. destruct x; auto. + exploit (exec_arith_extended Val.addl Paddext (Padd X)); auto. instantiate (1 := x0). eauto with asmgen. - intros (rs' & A & B & C). + instantiate (1 := X16). simpl. congruence. + intros (rs' & A & B & C & D). econstructor; exists rs'; split. apply exec_straight_opt_intro. eexact A. split. simpl. rewrite B. rewrite Val.addl_assoc. f_equal. @@ -1823,7 +1824,9 @@ Proof. apply exec_straight_opt_intro. apply exec_straight_one. simpl; eauto. auto. split. simpl. Simpl. rewrite symbol_high_low. simpl in EV. congruence. intros; Simpl. -+ exploit (exec_loadsymbol X16 id ofs). auto. intros (rs' & A & B & C). ++ exploit (exec_loadsymbol X16 id ofs). auto. + simpl. congruence. + intros (rs' & A & B & C & D). econstructor; exists rs'; split. apply exec_straight_opt_intro. eexact A. split. simpl. @@ -1837,7 +1840,9 @@ Proof. destruct (offset_representable sz (Ptrofs.to_int64 ofs)); inv TR. + econstructor; econstructor; split. apply exec_straight_opt_refl. auto. -+ exploit (exec_loadimm64 X16 (Ptrofs.to_int64 ofs)). intros (rs' & A & B & C). ++ exploit (exec_loadimm64 X16 (Ptrofs.to_int64 ofs)). + simpl. congruence. + intros (rs' & A & B & C). econstructor; exists rs'; split. apply exec_straight_opt_intro. eexact A. split. simpl. rewrite B, C by eauto with asmgen. auto. -- cgit From 6a05fbf4b55dfdf2884c3e0b2cb57d707ad1598d Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 24 Mar 2020 23:53:42 +0100 Subject: Asmgenproof1 --- aarch64/Asmgenproof1.v | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/aarch64/Asmgenproof1.v b/aarch64/Asmgenproof1.v index 4da1b52b..bd1474b6 100644 --- a/aarch64/Asmgenproof1.v +++ b/aarch64/Asmgenproof1.v @@ -1932,7 +1932,9 @@ Proof. { destruct (rs base); try discriminate. simpl in *. rewrite Ptrofs.of_int64_to_int64 by auto. auto. } destruct offset_representable. - econstructor; econstructor; split. apply exec_straight_opt_refl. auto. -- exploit (exec_loadimm64 X16); eauto. intros (rs' & A & B & C). +- exploit (exec_loadimm64 X16); eauto. + simpl. congruence. + intros (rs' & A & B & C). econstructor; econstructor; split. apply exec_straight_opt_intro; eexact A. split. simpl. rewrite B, C by eauto with asmgen. auto. auto. Qed. -- cgit From ef9c9d4eb1e6e4ed1db4b57d647d60b0491e63ca Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 25 Mar 2020 00:13:38 +0100 Subject: proof forward --- aarch64/Asmgenproof.v | 30 ++++++++++++++++++++++++------ 1 file changed, 24 insertions(+), 6 deletions(-) diff --git a/aarch64/Asmgenproof.v b/aarch64/Asmgenproof.v index 0dc37f36..0b863162 100644 --- a/aarch64/Asmgenproof.v +++ b/aarch64/Asmgenproof.v @@ -716,7 +716,7 @@ Opaque loadind. exploit eval_operation_lessdef. eapply preg_vals; eauto. eauto. eexact H0. intros [v' [A B]]. rewrite (sp_val _ _ _ AG) in A. left; eapply exec_straight_steps; eauto; intros. simpl in TR. - exploit transl_op_correct; eauto. intros [rs2 [P [Q R]]]. + exploit transl_op_correct; eauto. intros [rs2 [P [Q [R S]]]]. exists rs2; split. eauto. split. apply agree_set_undef_mreg with rs0; auto. apply Val.lessdef_trans with v'; auto. @@ -724,7 +724,8 @@ Opaque loadind. rewrite R; auto. apply preg_of_not_X29; auto. Local Transparent destroyed_by_op. destruct op; try exact I; simpl; congruence. - + rewrite S. + auto. - (* Mload *) assert (Op.eval_addressing tge sp addr (map rs args) = Some a). { rewrite <- H. apply eval_addressing_preserved. exact symbols_preserved. } @@ -732,10 +733,11 @@ Local Transparent destroyed_by_op. intros [a' [A B]]. rewrite (sp_val _ _ _ AG) in A. exploit Mem.loadv_extends; eauto. intros [v' [C D]]. left; eapply exec_straight_steps; eauto; intros. simpl in TR. - exploit transl_load_correct; eauto. intros [rs2 [P [Q R]]]. + exploit transl_load_correct; eauto. intros [rs2 [P [Q [R S]]]]. exists rs2; split. eauto. split. eapply agree_set_undef_mreg; eauto. congruence. - simpl; congruence. + split. simpl; congruence. + rewrite S. assumption. - (* Mstore *) assert (Op.eval_addressing tge sp addr (map rs args) = Some a). @@ -745,10 +747,11 @@ Local Transparent destroyed_by_op. assert (Val.lessdef (rs src) (rs0 (preg_of src))) by (eapply preg_val; eauto). exploit Mem.storev_extends; eauto. intros [m2' [C D]]. left; eapply exec_straight_steps; eauto. - intros. simpl in TR. exploit transl_store_correct; eauto. intros [rs2 [P Q]]. + intros. simpl in TR. exploit transl_store_correct; eauto. intros [rs2 [P [Q R]]]. exists rs2; split. eauto. split. eapply agree_undef_regs; eauto with asmgen. - simpl; congruence. + split. simpl; congruence. + rewrite R. assumption. - (* Mcall *) assert (f0 = f) by congruence. subst f0. @@ -857,6 +860,18 @@ Local Transparent destroyed_by_op. eapply agree_undef_regs; eauto. intros. rewrite undef_regs_other_2; auto. congruence. + Simpl. + rewrite set_res_other by trivial. + rewrite undef_regs_other. + assumption. + intro. + rewrite in_map_iff. + intros (x0 & PREG & IN). + subst r'. + intro. + apply (preg_of_not_RA x0). + congruence. + - (* Mgoto *) assert (f0 = f) by congruence. subst f0. inv AT. monadInv H4. @@ -870,6 +885,9 @@ Local Transparent destroyed_by_op. eapply agree_exten; eauto with asmgen. congruence. + rewrite INV by congruence. + assumption. + - (* Mcond true *) assert (f0 = f) by congruence. subst f0. exploit eval_condition_lessdef. eapply preg_vals; eauto. eauto. eauto. intros EC. -- cgit From 0e42b14d8e3c1a87a0242468bb5ace8ec8f9ef9a Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 25 Mar 2020 00:25:05 +0100 Subject: proof forward --- aarch64/Asmgenproof.v | 20 +++++++++++++++----- 1 file changed, 15 insertions(+), 5 deletions(-) diff --git a/aarch64/Asmgenproof.v b/aarch64/Asmgenproof.v index 0b863162..e7b13cc9 100644 --- a/aarch64/Asmgenproof.v +++ b/aarch64/Asmgenproof.v @@ -893,20 +893,25 @@ Local Transparent destroyed_by_op. exploit eval_condition_lessdef. eapply preg_vals; eauto. eauto. eauto. intros EC. left; eapply exec_straight_opt_steps_goto; eauto. intros. simpl in TR. - exploit transl_cond_branch_correct; eauto. intros (rs' & jmp & A & B & C). + exploit transl_cond_branch_correct; eauto. intros (rs' & jmp & A & B & C & D). exists jmp; exists k; exists rs'. split. eexact A. split. apply agree_exten with rs0; auto with asmgen. - exact B. + split. + exact B. + rewrite D. exact LEAF. - (* Mcond false *) exploit eval_condition_lessdef. eapply preg_vals; eauto. eauto. eauto. intros EC. left; eapply exec_straight_steps; eauto. intros. simpl in TR. - exploit transl_cond_branch_correct; eauto. intros (rs' & jmp & A & B & C). + exploit transl_cond_branch_correct; eauto. intros (rs' & jmp & A & B & C & D). econstructor; split. eapply exec_straight_opt_right. eexact A. apply exec_straight_one. eexact B. auto. split. apply agree_exten with rs0; auto. intros. Simpl. + split. simpl; congruence. + Simpl. rewrite D. + exact LEAF. - (* Mjumptable *) assert (f0 = f) by congruence. subst f0. @@ -928,6 +933,10 @@ Local Transparent destroyed_by_op. simpl. intros. rewrite C; auto with asmgen. Simpl. congruence. + rewrite C by congruence. + repeat rewrite Pregmap.gso by congruence. + assumption. + - (* Mreturn *) assert (f0 = f) by congruence. subst f0. inversion AT; subst. simpl in H6; monadInv H6. @@ -970,7 +979,7 @@ Local Transparent destroyed_by_op. simpl preg_of_iregsp. change (rs2 X30) with (rs0 X30). rewrite ATLR. change (rs2 X2) with sp. eexact P. simpl; congruence. congruence. - intros (rs3 & U & V). + intros (rs3 & U & V & W). assert (EXEC_PROLOGUE: exec_straight tge tf tf.(fn_code) rs0 m' @@ -996,7 +1005,8 @@ Local Transparent destroyed_at_function_entry. simpl. simpl; intros; Simpl. unfold sp; congruence. intros. rewrite V by auto with asmgen. reflexivity. - + intro IS_LEAF. + - (* external function *) exploit functions_translated; eauto. intros [tf [A B]]. simpl in B. inv B. -- cgit From eb6e959c60a799c368faf3a59b565217d52376f1 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 25 Mar 2020 00:33:10 +0100 Subject: proof forward --- aarch64/Asmgenproof.v | 5 ++++- 1 file changed, 4 insertions(+), 1 deletion(-) diff --git a/aarch64/Asmgenproof.v b/aarch64/Asmgenproof.v index e7b13cc9..1e443486 100644 --- a/aarch64/Asmgenproof.v +++ b/aarch64/Asmgenproof.v @@ -1005,7 +1005,10 @@ Local Transparent destroyed_at_function_entry. simpl. simpl; intros; Simpl. unfold sp; congruence. intros. rewrite V by auto with asmgen. reflexivity. - intro IS_LEAF. + + rewrite W. + unfold rs2. + Simpl. - (* external function *) exploit functions_translated; eauto. -- cgit From b096dac760b7d306c85e2b6b9b56779018596916 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 25 Mar 2020 00:40:01 +0100 Subject: RA is preserved --- aarch64/Asmgenproof.v | 22 ++++++++++++++++------ aarch64/Asmgenproof1.v | 32 ++++++++++++++++++++------------ 2 files changed, 36 insertions(+), 18 deletions(-) diff --git a/aarch64/Asmgenproof.v b/aarch64/Asmgenproof.v index 1e443486..5f88b99b 100644 --- a/aarch64/Asmgenproof.v +++ b/aarch64/Asmgenproof.v @@ -636,7 +636,7 @@ Qed. Theorem step_simulation: forall S1 t S2, Mach.step return_address_offset ge S1 t S2 -> - forall S1' (MS: match_states S1 S1'), + forall S1' (MS: match_states S1 S1') (WF: wf_state ge S1), (exists S2', plus step tge S1' t S2' /\ match_states S2 S2') \/ (measure S2 < measure S1 /\ t = E0 /\ match_states S2 S1')%nat. Proof. @@ -1029,6 +1029,10 @@ Local Transparent destroyed_at_function_entry. simpl. right. split. omega. split. auto. rewrite <- ATPC in H5. econstructor; eauto. congruence. + inv WF. + inv STACK. + inv H1. + congruence. Qed. Lemma transf_initial_states: @@ -1064,11 +1068,17 @@ Qed. Theorem transf_program_correct: forward_simulation (Mach.semantics return_address_offset prog) (Asm.semantics tprog). Proof. - eapply forward_simulation_star with (measure := measure). - apply senv_preserved. - eexact transf_initial_states. - eexact transf_final_states. - exact step_simulation. + eapply forward_simulation_star with (measure := measure) + (match_states := fun S1 S2 => match_states S1 S2 /\ wf_state ge S1). + - apply senv_preserved. + - simpl; intros. exploit transf_initial_states; eauto. + intros (s2 & A & B). + exists s2; intuition auto. apply wf_initial; auto. + - simpl; intros. destruct H as [MS WF]. eapply transf_final_states; eauto. + - simpl; intros. destruct H0 as [MS WF]. + exploit step_simulation; eauto. intros [ (s2' & A & B) | (A & B & C) ]. + + left; exists s2'; intuition auto. eapply wf_step; eauto. + + right; intuition auto. eapply wf_step; eauto. Qed. End PRESERVATION. diff --git a/aarch64/Asmgenproof1.v b/aarch64/Asmgenproof1.v index bd1474b6..b851966d 100644 --- a/aarch64/Asmgenproof1.v +++ b/aarch64/Asmgenproof1.v @@ -1776,7 +1776,8 @@ Lemma transl_addressing_correct: exists ad rs', exec_straight_opt ge fn c rs m (insn ad :: k) rs' m /\ Asm.eval_addressing ge ad rs' = Vptr b o - /\ forall r, data_preg r = true -> rs' r = rs r. + /\ (forall r, data_preg r = true -> rs' r = rs r) + /\ rs' # RA = rs # RA. Proof. intros until o; intros TR EV. unfold transl_addressing in TR; destruct addr; ArgsInv; SimplEval EV. @@ -1787,7 +1788,7 @@ Proof. + exploit (exec_loadimm64 X16 ofs). congruence. intros (rs' & A & B & C). econstructor; exists rs'; split. apply exec_straight_opt_intro; eexact A. split. simpl. rewrite B, C by eauto with asmgen. auto. - eauto with asmgen. + split; eauto with asmgen. - (* Aindexed2 *) econstructor; econstructor; split. apply exec_straight_opt_refl. auto. @@ -1803,7 +1804,7 @@ Proof. + econstructor; econstructor; split. apply exec_straight_opt_intro. apply exec_straight_one. simpl; eauto. auto. split. simpl. Simpl. rewrite H0. simpl. rewrite Ptrofs.add_zero. auto. - intros; Simpl. + split; intros; Simpl. - (* Aindexed2ext *) destruct (Int.eq a Int.zero || Int.eq (Int.shl Int.one a) (Int.repr sz)); inv EQ2. + econstructor; econstructor; split. apply exec_straight_opt_refl. @@ -1817,13 +1818,15 @@ Proof. split. simpl. rewrite B. rewrite Val.addl_assoc. f_equal. unfold Op.eval_extend; destruct x, (rs x1); simpl; auto; rewrite ! a64_range; simpl; rewrite Int64.add_zero; auto. - intros. apply C; eauto with asmgen. + split; intros. + apply C; eauto with asmgen. + trivial. - (* Aglobal *) destruct (Ptrofs.eq (Ptrofs.modu ofs (Ptrofs.repr sz)) Ptrofs.zero && symbol_is_aligned id sz); inv TR. + econstructor; econstructor; split. apply exec_straight_opt_intro. apply exec_straight_one. simpl; eauto. auto. split. simpl. Simpl. rewrite symbol_high_low. simpl in EV. congruence. - intros; Simpl. + split; intros; Simpl. + exploit (exec_loadsymbol X16 id ofs). auto. simpl. congruence. intros (rs' & A & B & C & D). @@ -1832,7 +1835,7 @@ Proof. split. simpl. rewrite B. rewrite <- Genv.shift_symbol_address_64, Ptrofs.add_zero by auto. simpl in EV. congruence. - auto with asmgen. + split; auto with asmgen. - (* Ainstrack *) assert (E: Val.addl (rs SP) (Vlong (Ptrofs.to_int64 ofs)) = Vptr b o). { simpl in EV. inv EV. destruct (rs SP); simpl in H1; inv H1. simpl. @@ -1857,7 +1860,8 @@ Lemma transl_load_correct: exists rs', exec_straight ge fn c rs m k rs' m /\ rs'#(preg_of dst) = v - /\ forall r, data_preg r = true -> r <> preg_of dst -> rs' r = rs r. + /\ (forall r, data_preg r = true -> r <> preg_of dst -> rs' r = rs r) + /\ rs' # RA = rs # RA. Proof. intros. destruct vaddr; try discriminate. assert (A: exists sz insn, @@ -1870,14 +1874,17 @@ Proof. do 2 econstructor; (split; [eassumption|auto]). } destruct A as (sz & insn & B & C). - exploit transl_addressing_correct. eexact B. eexact H0. intros (ad & rs' & P & Q & R). + exploit transl_addressing_correct. eexact B. eexact H0. intros (ad & rs' & P & Q & R & S). assert (X: exec_load ge chunk (fun v => v) ad (preg_of dst) rs' m = Next (nextinstr (rs'#(preg_of dst) <- v)) m). { unfold exec_load. rewrite Q, H1. auto. } econstructor; split. eapply exec_straight_opt_right. eexact P. apply exec_straight_one. rewrite C, X; eauto. Simpl. - split. Simpl. intros; Simpl. + split. Simpl. + split; intros; Simpl. + rewrite <- S. + apply RA_not_written. Qed. Lemma transl_store_correct: @@ -1887,7 +1894,8 @@ Lemma transl_store_correct: Mem.storev chunk m vaddr rs#(preg_of src) = Some m' -> exists rs', exec_straight ge fn c rs m k rs' m' - /\ forall r, data_preg r = true -> rs' r = rs r. + /\ (forall r, data_preg r = true -> rs' r = rs r) + /\ rs' # RA = rs # RA. Proof. intros. destruct vaddr; try discriminate. set (chunk' := match chunk with Mint8signed => Mint8unsigned @@ -1903,7 +1911,7 @@ Proof. do 2 econstructor; (split; [eassumption|auto]). } destruct A as (sz & insn & B & C). - exploit transl_addressing_correct. eexact B. eexact H0. intros (ad & rs' & P & Q & R). + exploit transl_addressing_correct. eexact B. eexact H0. intros (ad & rs' & P & Q & R & S). assert (X: Mem.storev chunk' m (Vptr b i) rs#(preg_of src) = Some m'). { rewrite <- H1. unfold chunk'. destruct chunk; auto; simpl; symmetry. apply Mem.store_signed_unsigned_8. @@ -1914,7 +1922,7 @@ Proof. econstructor; split. eapply exec_straight_opt_right. eexact P. apply exec_straight_one. rewrite C, Y; eauto. Simpl. - intros; Simpl. + split; intros; Simpl. Qed. (** Translation of indexed memory accesses *) -- cgit From 469a282add58074e2bc1a2822125d18e4dc6a80d Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 25 Mar 2020 00:57:57 +0100 Subject: removed RA restoration --- aarch64/Asmgen.v | 6 ++++-- aarch64/Asmgenproof.v | 7 ++++++- aarch64/Asmgenproof1.v | 18 +++++++++++++++++- 3 files changed, 27 insertions(+), 4 deletions(-) diff --git a/aarch64/Asmgen.v b/aarch64/Asmgen.v index 875f3fd1..fc083223 100644 --- a/aarch64/Asmgen.v +++ b/aarch64/Asmgen.v @@ -1050,8 +1050,10 @@ Definition storeptr (src: ireg) (base: iregsp) (ofs: ptrofs) (k: code) := (** Function epilogue *) Definition make_epilogue (f: Mach.function) (k: code) := - loadptr XSP f.(fn_retaddr_ofs) RA - (Pfreeframe f.(fn_stacksize) f.(fn_link_ofs) :: k). + if is_leaf_function f + then Pfreeframe f.(fn_stacksize) f.(fn_link_ofs) :: k + else loadptr XSP f.(fn_retaddr_ofs) RA + (Pfreeframe f.(fn_stacksize) f.(fn_link_ofs) :: k). (** Translation of a Mach instruction. *) diff --git a/aarch64/Asmgenproof.v b/aarch64/Asmgenproof.v index 5f88b99b..90a06144 100644 --- a/aarch64/Asmgenproof.v +++ b/aarch64/Asmgenproof.v @@ -337,7 +337,12 @@ Qed. Remark make_epilogue_label: forall f k, tail_nolabel k (make_epilogue f k). Proof. - unfold make_epilogue; intros. eapply tail_nolabel_trans. apply loadptr_label. TailNoLabel. + unfold make_epilogue; intros. + destruct is_leaf_function. + { TailNoLabel. } + eapply tail_nolabel_trans. + apply loadptr_label. + TailNoLabel. Qed. Lemma transl_instr_label: diff --git a/aarch64/Asmgenproof1.v b/aarch64/Asmgenproof1.v index b851966d..91c5c306 100644 --- a/aarch64/Asmgenproof1.v +++ b/aarch64/Asmgenproof1.v @@ -2044,6 +2044,7 @@ Qed. Lemma make_epilogue_correct: forall ge0 f m stk soff cs m' ms rs k tm, + (is_leaf_function f = true -> rs # (IR RA) = parent_ra cs) -> 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' -> @@ -2058,7 +2059,7 @@ Lemma make_epilogue_correct: /\ rs'#SP = parent_sp cs /\ (forall r, r <> PC -> r <> SP -> r <> RA -> r <> X16 -> rs'#r = rs#r). Proof. - intros until tm; intros LP LRA FREE AG MEXT MCS. + intros until tm; intros LEAF_RA 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'. @@ -2068,6 +2069,21 @@ Proof. exploit (loadptr_correct XSP (fn_retaddr_ofs f)). instantiate (2 := rs). simpl. rewrite <- (sp_val _ _ _ AG). simpl. eexact LRA'. simpl; congruence. intros (rs1 & A1 & B1 & C1). + destruct (is_leaf_function f). + { + econstructor; econstructor; split. + apply exec_straight_one. simpl. + rewrite <- (sp_val _ _ _ AG). simpl; rewrite LP'. + rewrite FREE'. eauto. auto. + split. apply agree_nextinstr. apply agree_set_other; auto. + apply agree_change_sp with (Vptr stk soff). + apply agree_exten with rs; auto. + eapply parent_sp_def; eauto. + split. auto. + split. Simpl. + split. Simpl. + intros. Simpl. + } econstructor; econstructor; split. eapply exec_straight_trans. eexact A1. apply exec_straight_one. simpl. simpl; rewrite (C1 SP) by auto with asmgen. rewrite <- (sp_val _ _ _ AG). simpl; rewrite LP'. -- cgit From e991d17839bf1c4736bdb5d8cbbd956be6fb2a1e Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 25 Mar 2020 07:38:07 +0100 Subject: better epilogue proof --- aarch64/Asmgenproof1.v | 26 ++++++++++++++++++-------- 1 file changed, 18 insertions(+), 8 deletions(-) diff --git a/aarch64/Asmgenproof1.v b/aarch64/Asmgenproof1.v index 91c5c306..48c7f4e6 100644 --- a/aarch64/Asmgenproof1.v +++ b/aarch64/Asmgenproof1.v @@ -2046,7 +2046,7 @@ Lemma make_epilogue_correct: forall ge0 f m stk soff cs m' ms rs k tm, (is_leaf_function f = true -> rs # (IR RA) = parent_ra cs) -> 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) -> + (is_leaf_function f = false -> 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 -> @@ -2060,17 +2060,14 @@ Lemma make_epilogue_correct: /\ (forall r, r <> PC -> r <> SP -> r <> RA -> r <> X16 -> rs'#r = rs#r). Proof. intros until tm; intros LEAF_RA LP LRA FREE AG MEXT MCS. + destruct (is_leaf_function f) eqn:IS_LEAF. + { 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. - exploit (loadptr_correct XSP (fn_retaddr_ofs f)). - instantiate (2 := rs). simpl. rewrite <- (sp_val _ _ _ AG). simpl. eexact LRA'. simpl; congruence. - intros (rs1 & A1 & B1 & C1). - destruct (is_leaf_function f). - { + rewrite IS_LEAF. + econstructor; econstructor; split. apply exec_straight_one. simpl. rewrite <- (sp_val _ _ _ AG). simpl; rewrite LP'. @@ -2084,6 +2081,19 @@ Proof. split. Simpl. intros. Simpl. } + lapply LRA. 2: reflexivity. + clear LRA. intro LRA. + 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 IS_LEAF. + exploit (loadptr_correct XSP (fn_retaddr_ofs f)). + instantiate (2 := rs). simpl. rewrite <- (sp_val _ _ _ AG). simpl. eexact LRA'. simpl; congruence. + intros (rs1 & A1 & B1 & C1). + econstructor; econstructor; split. eapply exec_straight_trans. eexact A1. apply exec_straight_one. simpl. simpl; rewrite (C1 SP) by auto with asmgen. rewrite <- (sp_val _ _ _ AG). simpl; rewrite LP'. -- cgit From ed399d8dcb3b41dfacf8257c22c608061503fd3d Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 25 Mar 2020 16:49:29 +0100 Subject: Loop heuristic > Call heuristic --- backend/Duplicateaux.ml | 5 +++-- 1 file changed, 3 insertions(+), 2 deletions(-) diff --git a/backend/Duplicateaux.ml b/backend/Duplicateaux.ml index fedb63fe..54d60d24 100644 --- a/backend/Duplicateaux.ml +++ b/backend/Duplicateaux.ml @@ -269,8 +269,9 @@ let get_directions code entrypoint = begin match (get_some @@ PTree.get n code) with | Icond (cond, lr, ifso, ifnot, _) -> (* Printf.printf "Analyzing %d.." (P.to_int n); *) - let heuristics = [ do_call_heuristic; do_opcode_heuristic; - do_return_heuristic; do_loop2_heuristic loop_info n; do_loop_heuristic; (* do_store_heuristic *) ] in + let heuristics = [ do_opcode_heuristic; + do_return_heuristic; do_loop2_heuristic loop_info n; do_loop_heuristic; do_call_heuristic; + (* do_store_heuristic *) ] in let preferred = ref None in begin Printf.printf "Deciding condition for RTL node %d\n" (P.to_int n); -- cgit From 36589dd043392d4d8672a82f24975371c102c286 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 25 Mar 2020 16:49:46 +0100 Subject: Linearize: Scheduling based on maxpc instead of dependencies --- backend/Linearizeaux.ml | 12 +++++++++--- 1 file changed, 9 insertions(+), 3 deletions(-) diff --git a/backend/Linearizeaux.ml b/backend/Linearizeaux.ml index 3f207d9e..605a5db5 100644 --- a/backend/Linearizeaux.ml +++ b/backend/Linearizeaux.ml @@ -419,9 +419,13 @@ let order_sequences code entry fs = assert (not fs_evaluated.(s_id)); ordered_fs := fs_a.(s_id) :: !ordered_fs; fs_evaluated.(s_id) <- true; + (* Printf.printf "++++++\n"; + Printf.printf "Scheduling %d\n" s_id; + Printf.printf "Initial depmap: "; print_depmap depmap; *) Array.iteri (fun i deps -> depmap.(i) <- ISet.remove s_id deps - ) depmap + ) depmap; + (* Printf.printf "Final depmap: "; print_depmap depmap; *) end in let choose_best_of candidates = let current_best_id = ref None in @@ -451,8 +455,10 @@ let order_sequences code entry fs = begin Array.iteri (fun i deps -> begin - (* Printf.printf "Deps: "; print_iset deps; Printf.printf "\n"; *) - if (deps == ISet.empty && not fs_evaluated.(i)) then candidates := i :: !candidates + (* Printf.printf "Deps of %d: " i; print_iset deps; Printf.printf "\n"; *) + (* FIXME - if we keep it that way (no dependency check), remove all the unneeded stuff *) + if ((* deps == ISet.empty && *) not fs_evaluated.(i)) then + candidates := i :: !candidates end ) depmap; if not (List.length !candidates > 0) then begin -- cgit From b46cdc3ade397a57a7b748946fb58e16e95bf42b Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 27 Mar 2020 16:53:59 +0100 Subject: fix broken test Makefile fix math.h so that it does special things only on K1C --- runtime/include/math.h | 7 +++++++ test/Makefile | 4 +++- 2 files changed, 10 insertions(+), 1 deletion(-) diff --git a/runtime/include/math.h b/runtime/include/math.h index d6475df1..01b8d8d8 100644 --- a/runtime/include/math.h +++ b/runtime/include/math.h @@ -1,6 +1,8 @@ #ifndef _COMPCERT_MATH_H #define _COMPCERT_MATH_H +#ifdef __K1C__ + #define isfinite(__y) (fpclassify((__y)) >= FP_ZERO) #include_next @@ -16,4 +18,9 @@ #define fmaf(x, y, z) __builtin_fmaf((x),(y),(z)) #endif +#else + +#include_next + +#endif #endif diff --git a/test/Makefile b/test/Makefile index 7efcd8f1..e9c5d6a1 100644 --- a/test/Makefile +++ b/test/Makefile @@ -4,7 +4,9 @@ include ../Makefile.config # Kalray note - removing compression, raytracer and spass that cannot be executed by the simulator in reasonable time ifeq ($(ARCH),mppa_k1c) - DIRS:=c regression + DIRS=c regression +else + DIRS=c compression raytracer spass regression endif ifeq ($(CLIGHTGEN),true) -- cgit From a12c5d99df634bec3f95f2e10664b429173e49aa Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 27 Mar 2020 17:13:08 +0100 Subject: fix issues in Mandelbrot due to modifications in the source code --- test/c/mandelbrot.c | 2 -- 1 file changed, 2 deletions(-) diff --git a/test/c/mandelbrot.c b/test/c/mandelbrot.c index fb8b929c..548c3ffa 100644 --- a/test/c/mandelbrot.c +++ b/test/c/mandelbrot.c @@ -59,7 +59,6 @@ int main (int argc, char **argv) if(bit_num == 8) { - printf("%c", byte_acc); putc(byte_acc,stdout); #ifdef __K1C__ // stdout isn't flushed enough when --syscall=libstd_scalls.so is passed to the simulator k1-cluster fflush(stdout); @@ -70,7 +69,6 @@ int main (int argc, char **argv) else if(x == w-1) { byte_acc <<= (8-w%8); - printf("%c", byte_acc); putc(byte_acc,stdout); #ifdef __K1C__ // stdout isn't flushed enough when --syscall=libstd_scalls.so is passed to the simulator k1-cluster fflush(stdout); -- cgit From e0a48d116d222425bf40d1fc5f79b68c7ce9a37f Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 27 Mar 2020 17:24:02 +0100 Subject: disable some tests --- test/regression/Makefile | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/test/regression/Makefile b/test/regression/Makefile index 3447d6a5..ad3ffd99 100644 --- a/test/regression/Makefile +++ b/test/regression/Makefile @@ -26,18 +26,18 @@ TESTS_COMP?=attribs1 bitfields1 bitfields2 bitfields3 bitfields4 \ varargs1 varargs2 varargs3 sections alias aligned\ packedstruct1 packedstruct2 -ifeq ($(ARCH),mppa_k1c) +# FIXME ifeq ($(ARCH),mppa_k1c) TESTS_COMP:=$(filter-out packedstruct1,$(TESTS_COMP)) TESTS_COMP:=$(filter-out packedstruct2,$(TESTS_COMP)) -endif +# endif # Can run, both in compiled mode and in interpreter mode, # but produce processor-dependent results, so no reference output in Results TESTS_DIFF=NaNs -ifeq ($(ARCH),mppa_k1c) +# FIXME ifeq ($(ARCH),mppa_k1c) TESTS_DIFF:=$(filter-out NaNs,$(TESTS_DIFF)) -endif +# endif endif # Other tests: should compile to .s without errors (but expect warnings) -- cgit From 210541294da0f718be67a427426f2ee4c353e858 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 27 Mar 2020 17:29:41 +0100 Subject: run tests on aarch64 --- .gitlab-ci.yml | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index ad684229..0d07db41 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -44,13 +44,14 @@ build_aarch64: image: "coqorg/coq" before_script: - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update - - sudo apt-get -y install gcc-aarch64-linux-gnu + - sudo apt-get -y install gcc-aarch64-linux-gnu qemu-user-binfmt - opam switch 4.07.1+flambda - eval `opam config env` - opam install -y menhir script: - ./config_aarch64.sh - make -j "$NJOBS" + - cd test && make CCOMPOPTS=-static all test rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always -- cgit From 3a81fb827a35577f13bcc6594a503449a1d180f8 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 27 Mar 2020 17:47:15 +0100 Subject: do not use binfmt --- .gitlab-ci.yml | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 0d07db41..2bcd99ca 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -44,14 +44,14 @@ build_aarch64: image: "coqorg/coq" before_script: - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update - - sudo apt-get -y install gcc-aarch64-linux-gnu qemu-user-binfmt + - sudo apt-get -y install gcc-aarch64-linux-gnu qemu-user - opam switch 4.07.1+flambda - eval `opam config env` - opam install -y menhir script: - ./config_aarch64.sh - make -j "$NJOBS" - - cd test && make CCOMPOPTS=-static all test + - cd test && make CCOMPOPTS=-static SIMU='/opt/qemu/4.2.0/bin/qemu-aarch64' EXECUTE='/opt/qemu/4.2.0/bin/qemu-aarch64' all test rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always -- cgit From d77f715cf71e7afd71f374d41a1b158362af4f62 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 27 Mar 2020 18:00:38 +0100 Subject: call standard qemu not mine! --- .gitlab-ci.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 2bcd99ca..c57a5a84 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -51,7 +51,7 @@ build_aarch64: script: - ./config_aarch64.sh - make -j "$NJOBS" - - cd test && make CCOMPOPTS=-static SIMU='/opt/qemu/4.2.0/bin/qemu-aarch64' EXECUTE='/opt/qemu/4.2.0/bin/qemu-aarch64' all test + - cd test && make CCOMPOPTS=-static SIMU='qemu-aarch64' EXECUTE='qemu-aarch64' all test rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always -- cgit From 4583f92b71edf48b47fadba06e60137870ffd003 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 27 Mar 2020 18:29:01 +0100 Subject: build and execute tests on other architectures than aarch64 --- .gitlab-ci.yml | 15 +++++++++++---- 1 file changed, 11 insertions(+), 4 deletions(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index c57a5a84..1623d1ba 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -11,6 +11,7 @@ build_x86_64: script: - ./config_x86_64.sh - make -j "$NJOBS" + - cd test && make all test rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always @@ -30,6 +31,7 @@ build_ia32: script: - ./config_ia32.sh - make -j "$NJOBS" + - cd test && make all test rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always @@ -66,13 +68,14 @@ build_arm: image: "coqorg/coq" before_script: - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update - - sudo apt-get -y install gcc-arm-linux-gnueabihf + - sudo apt-get -y install gcc-arm-linux-gnueabihf qemu-user - opam switch 4.07.1+flambda - eval `opam config env` - opam install -y menhir script: - ./config_arm.sh - make -j "$NJOBS" + - cd test && make CCOMPOPTS=-static SIMU='qemu-arm' EXECUTE='qemu-arm' all test rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always @@ -87,13 +90,14 @@ build_ppc: image: "coqorg/coq" before_script: - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update - - sudo apt-get -y install gcc-powerpc-linux-gnu + - sudo apt-get -y install gcc-powerpc-linux-gnu qemu-user - opam switch 4.07.1+flambda - eval `opam config env` - opam install -y menhir script: - ./config_ppc.sh - make -j "$NJOBS" + - cd test && make CCOMPOPTS=-static SIMU='qemu-ppc' EXECUTE='qemu-ppc' all test rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always @@ -108,13 +112,14 @@ build_rv64: image: "coqorg/coq" before_script: - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update - - sudo apt-get -y install gcc-riscv64-linux-gnu + - sudo apt-get -y install gcc-riscv64-linux-gnu qemu-user - opam switch 4.07.1+flambda - eval `opam config env` - opam install -y menhir script: - ./config_rv64.sh - make -j "$NJOBS" + - cd test && make CCOMPOPTS=-static SIMU='qemu-riscv64' EXECUTE='qemu-riscv64' all test rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always @@ -129,13 +134,14 @@ build_rv32: image: "coqorg/coq" before_script: - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update - - sudo apt-get -y install gcc-riscv64-linux-gnu + - sudo apt-get -y install gcc-riscv64-linux-gnu qemu-user - opam switch 4.07.1+flambda - eval `opam config env` - opam install -y menhir script: - ./config_rv32.sh -no-runtime-lib - make -j "$NJOBS" + - cd test && make CCOMPOPTS=-static SIMU='qemu-riscv32' EXECUTE='qemu-riscv32' all test rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always @@ -163,3 +169,4 @@ build_k1c: - if: '$CI_COMMIT_BRANCH == "master"' when: always - when: manual + - cd test && make CCOMPOPTS=-static all -- cgit From 4799ad6121055c7a95c5a9c3c76d15706abafb6d Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 27 Mar 2020 18:30:16 +0100 Subject: wrong line --- .gitlab-ci.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 1623d1ba..aa215d38 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -161,6 +161,7 @@ build_k1c: script: - ./config_k1c.sh -no-runtime-lib - make -j "$NJOBS" + - cd test && make CCOMPOPTS=-static all rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always @@ -169,4 +170,3 @@ build_k1c: - if: '$CI_COMMIT_BRANCH == "master"' when: always - when: manual - - cd test && make CCOMPOPTS=-static all -- cgit From 2a92e61c7469e4ce4340b64fad59508a71b6efb1 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 27 Mar 2020 18:56:16 +0100 Subject: fix config for K1C PPC RV32 for CI --- .gitlab-ci.yml | 4 ++-- config_rv32.sh | 2 +- 2 files changed, 3 insertions(+), 3 deletions(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index aa215d38..1e392b2c 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -25,6 +25,7 @@ build_ia32: stage: build image: "coqorg/coq" before_script: + - sudo apt-get -y install gcc-multilib - opam switch 4.07.1+flambda - eval `opam config env` - opam install -y menhir @@ -90,7 +91,7 @@ build_ppc: image: "coqorg/coq" before_script: - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update - - sudo apt-get -y install gcc-powerpc-linux-gnu qemu-user + - sudo apt-get -y install gcc-multilib-powerpc qemu-user - opam switch 4.07.1+flambda - eval `opam config env` - opam install -y menhir @@ -161,7 +162,6 @@ build_k1c: script: - ./config_k1c.sh -no-runtime-lib - make -j "$NJOBS" - - cd test && make CCOMPOPTS=-static all rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always diff --git a/config_rv32.sh b/config_rv32.sh index 654cacfa..a5a5cf1c 100755 --- a/config_rv32.sh +++ b/config_rv32.sh @@ -1 +1 @@ -exec ./config_simple.sh rv32-linux --toolprefix riscv64-unknown-elf- "$@" +exec ./config_simple.sh rv32-linux --toolprefix riscv64-linux-gnu- "$@" -- cgit From 99f90f43d7942c8e9c2667c448a7bc876f5c72cc Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 27 Mar 2020 18:59:46 +0100 Subject: various fixes for multilib --- .gitlab-ci.yml | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 1e392b2c..0ac47449 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -25,6 +25,7 @@ build_ia32: stage: build image: "coqorg/coq" before_script: + - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update - sudo apt-get -y install gcc-multilib - opam switch 4.07.1+flambda - eval `opam config env` @@ -91,7 +92,7 @@ build_ppc: image: "coqorg/coq" before_script: - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update - - sudo apt-get -y install gcc-multilib-powerpc qemu-user + - sudo apt-get -y install gcc-multilib-powerpc-linux-gnu qemu-user - opam switch 4.07.1+flambda - eval `opam config env` - opam install -y menhir -- cgit From 7eb27df1b3f682bef18e58783f4ed866183d4303 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 27 Mar 2020 19:11:14 +0100 Subject: fixes --- .gitlab-ci.yml | 1 - 1 file changed, 1 deletion(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 0ac47449..6aca7a8c 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -143,7 +143,6 @@ build_rv32: script: - ./config_rv32.sh -no-runtime-lib - make -j "$NJOBS" - - cd test && make CCOMPOPTS=-static SIMU='qemu-riscv32' EXECUTE='qemu-riscv32' all test rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always -- cgit From f939375d5074fc9af004a4c3a3f51a7cb2b26caf Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 27 Mar 2020 19:29:34 +0100 Subject: more config --- .gitlab-ci.yml | 25 +++++++++++++++++++++++-- config_ppc64.sh | 1 + 2 files changed, 24 insertions(+), 2 deletions(-) create mode 100644 config_ppc64.sh diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 6aca7a8c..e67a7508 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -92,14 +92,35 @@ build_ppc: image: "coqorg/coq" before_script: - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update - - sudo apt-get -y install gcc-multilib-powerpc-linux-gnu qemu-user + - sudo apt-get -y install gcc-powerpc-linux-gnu qemu-user - opam switch 4.07.1+flambda - eval `opam config env` - opam install -y menhir script: - ./config_ppc.sh - make -j "$NJOBS" - - cd test && make CCOMPOPTS=-static SIMU='qemu-ppc' EXECUTE='qemu-ppc' all test + rules: + - if: '$CI_COMMIT_BRANCH == "mppa-work"' + when: always + - if: '$CI_COMMIT_BRANCH == "mppa-k1c"' + when: always + - if: '$CI_COMMIT_BRANCH == "master"' + when: always + - when: manual + +build_ppc64: + stage: build + image: "coqorg/coq" + before_script: + - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update + - sudo apt-get -y install gcc-powerpc64-linux-gnu qemu-user + - opam switch 4.07.1+flambda + - eval `opam config env` + - opam install -y menhir + script: + - ./config_ppc64.sh + - make -j "$NJOBS" + - cd test && make CCOMPOPTS=-static SIMU='qemu-ppc64' EXECUTE='qemu-ppc64' all test rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always diff --git a/config_ppc64.sh b/config_ppc64.sh new file mode 100644 index 00000000..df31c18f --- /dev/null +++ b/config_ppc64.sh @@ -0,0 +1 @@ +exec ./config_simple.sh ppc64-linux --toolprefix powerpc64-linux-gnu- "$@" -- cgit From ebab676bc779c533d408e65b3f8a42b77cb17f9e Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 27 Mar 2020 19:35:28 +0100 Subject: temporarily disable raytracer test on ARM https://github.com/AbsInt/CompCert/issues/342 --- .gitlab-ci.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index e67a7508..2dd70e21 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -77,7 +77,7 @@ build_arm: script: - ./config_arm.sh - make -j "$NJOBS" - - cd test && make CCOMPOPTS=-static SIMU='qemu-arm' EXECUTE='qemu-arm' all test + - cd test && make DIRS="c compression spass regression" CCOMPOPTS=-static SIMU='qemu-arm' EXECUTE='qemu-arm' all test rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always -- cgit From 9536534a099fb95ae8eeef37b3ba10e03e31e823 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 27 Mar 2020 19:38:56 +0100 Subject: +x --- config_ppc64.sh | 0 1 file changed, 0 insertions(+), 0 deletions(-) mode change 100644 => 100755 config_ppc64.sh diff --git a/config_ppc64.sh b/config_ppc64.sh old mode 100644 new mode 100755 -- cgit From cef775c724e819c01d25cc1d461112c7a12a3227 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 27 Mar 2020 19:49:45 +0100 Subject: disable testing on ppc64 --- .gitlab-ci.yml | 2 -- 1 file changed, 2 deletions(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 2dd70e21..9ee14292 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -113,14 +113,12 @@ build_ppc64: image: "coqorg/coq" before_script: - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update - - sudo apt-get -y install gcc-powerpc64-linux-gnu qemu-user - opam switch 4.07.1+flambda - eval `opam config env` - opam install -y menhir script: - ./config_ppc64.sh - make -j "$NJOBS" - - cd test && make CCOMPOPTS=-static SIMU='qemu-ppc64' EXECUTE='qemu-ppc64' all test rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always -- cgit From 63b547484630fc774be06cf02033e3f0ecbfc26f Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 27 Mar 2020 20:00:56 +0100 Subject: we still need a ppc64 compiler --- .gitlab-ci.yml | 1 + 1 file changed, 1 insertion(+) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 9ee14292..17228f76 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -113,6 +113,7 @@ build_ppc64: image: "coqorg/coq" before_script: - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update + - sudo apt-get -y install gcc-powerpc64-linux-gnu - opam switch 4.07.1+flambda - eval `opam config env` - opam install -y menhir -- cgit From 96482f7df095e6244c414a14b07771dbaef67aec Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 27 Mar 2020 20:13:49 +0100 Subject: Run tests on various targets in addition to compiling --- .gitlab-ci.yml | 38 +++++++++++++++++++++++++++++++++----- config_ppc64.sh | 1 + config_rv32.sh | 2 +- runtime/include/math.h | 7 +++++++ test/Makefile | 4 +++- test/c/mandelbrot.c | 2 -- test/regression/Makefile | 8 ++++---- 7 files changed, 49 insertions(+), 13 deletions(-) create mode 100755 config_ppc64.sh diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index ad684229..17228f76 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -11,6 +11,7 @@ build_x86_64: script: - ./config_x86_64.sh - make -j "$NJOBS" + - cd test && make all test rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always @@ -24,12 +25,15 @@ build_ia32: stage: build image: "coqorg/coq" before_script: + - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update + - sudo apt-get -y install gcc-multilib - opam switch 4.07.1+flambda - eval `opam config env` - opam install -y menhir script: - ./config_ia32.sh - make -j "$NJOBS" + - cd test && make all test rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always @@ -44,13 +48,14 @@ build_aarch64: image: "coqorg/coq" before_script: - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update - - sudo apt-get -y install gcc-aarch64-linux-gnu + - sudo apt-get -y install gcc-aarch64-linux-gnu qemu-user - opam switch 4.07.1+flambda - eval `opam config env` - opam install -y menhir script: - ./config_aarch64.sh - make -j "$NJOBS" + - cd test && make CCOMPOPTS=-static SIMU='qemu-aarch64' EXECUTE='qemu-aarch64' all test rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always @@ -65,13 +70,14 @@ build_arm: image: "coqorg/coq" before_script: - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update - - sudo apt-get -y install gcc-arm-linux-gnueabihf + - sudo apt-get -y install gcc-arm-linux-gnueabihf qemu-user - opam switch 4.07.1+flambda - eval `opam config env` - opam install -y menhir script: - ./config_arm.sh - make -j "$NJOBS" + - cd test && make DIRS="c compression spass regression" CCOMPOPTS=-static SIMU='qemu-arm' EXECUTE='qemu-arm' all test rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always @@ -86,7 +92,7 @@ build_ppc: image: "coqorg/coq" before_script: - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update - - sudo apt-get -y install gcc-powerpc-linux-gnu + - sudo apt-get -y install gcc-powerpc-linux-gnu qemu-user - opam switch 4.07.1+flambda - eval `opam config env` - opam install -y menhir @@ -102,18 +108,40 @@ build_ppc: when: always - when: manual +build_ppc64: + stage: build + image: "coqorg/coq" + before_script: + - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update + - sudo apt-get -y install gcc-powerpc64-linux-gnu + - opam switch 4.07.1+flambda + - eval `opam config env` + - opam install -y menhir + script: + - ./config_ppc64.sh + - make -j "$NJOBS" + rules: + - if: '$CI_COMMIT_BRANCH == "mppa-work"' + when: always + - if: '$CI_COMMIT_BRANCH == "mppa-k1c"' + when: always + - if: '$CI_COMMIT_BRANCH == "master"' + when: always + - when: manual + build_rv64: stage: build image: "coqorg/coq" before_script: - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update - - sudo apt-get -y install gcc-riscv64-linux-gnu + - sudo apt-get -y install gcc-riscv64-linux-gnu qemu-user - opam switch 4.07.1+flambda - eval `opam config env` - opam install -y menhir script: - ./config_rv64.sh - make -j "$NJOBS" + - cd test && make CCOMPOPTS=-static SIMU='qemu-riscv64' EXECUTE='qemu-riscv64' all test rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always @@ -128,7 +156,7 @@ build_rv32: image: "coqorg/coq" before_script: - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update - - sudo apt-get -y install gcc-riscv64-linux-gnu + - sudo apt-get -y install gcc-riscv64-linux-gnu qemu-user - opam switch 4.07.1+flambda - eval `opam config env` - opam install -y menhir diff --git a/config_ppc64.sh b/config_ppc64.sh new file mode 100755 index 00000000..df31c18f --- /dev/null +++ b/config_ppc64.sh @@ -0,0 +1 @@ +exec ./config_simple.sh ppc64-linux --toolprefix powerpc64-linux-gnu- "$@" diff --git a/config_rv32.sh b/config_rv32.sh index 654cacfa..a5a5cf1c 100755 --- a/config_rv32.sh +++ b/config_rv32.sh @@ -1 +1 @@ -exec ./config_simple.sh rv32-linux --toolprefix riscv64-unknown-elf- "$@" +exec ./config_simple.sh rv32-linux --toolprefix riscv64-linux-gnu- "$@" diff --git a/runtime/include/math.h b/runtime/include/math.h index d6475df1..01b8d8d8 100644 --- a/runtime/include/math.h +++ b/runtime/include/math.h @@ -1,6 +1,8 @@ #ifndef _COMPCERT_MATH_H #define _COMPCERT_MATH_H +#ifdef __K1C__ + #define isfinite(__y) (fpclassify((__y)) >= FP_ZERO) #include_next @@ -16,4 +18,9 @@ #define fmaf(x, y, z) __builtin_fmaf((x),(y),(z)) #endif +#else + +#include_next + +#endif #endif diff --git a/test/Makefile b/test/Makefile index 7efcd8f1..e9c5d6a1 100644 --- a/test/Makefile +++ b/test/Makefile @@ -4,7 +4,9 @@ include ../Makefile.config # Kalray note - removing compression, raytracer and spass that cannot be executed by the simulator in reasonable time ifeq ($(ARCH),mppa_k1c) - DIRS:=c regression + DIRS=c regression +else + DIRS=c compression raytracer spass regression endif ifeq ($(CLIGHTGEN),true) diff --git a/test/c/mandelbrot.c b/test/c/mandelbrot.c index fb8b929c..548c3ffa 100644 --- a/test/c/mandelbrot.c +++ b/test/c/mandelbrot.c @@ -59,7 +59,6 @@ int main (int argc, char **argv) if(bit_num == 8) { - printf("%c", byte_acc); putc(byte_acc,stdout); #ifdef __K1C__ // stdout isn't flushed enough when --syscall=libstd_scalls.so is passed to the simulator k1-cluster fflush(stdout); @@ -70,7 +69,6 @@ int main (int argc, char **argv) else if(x == w-1) { byte_acc <<= (8-w%8); - printf("%c", byte_acc); putc(byte_acc,stdout); #ifdef __K1C__ // stdout isn't flushed enough when --syscall=libstd_scalls.so is passed to the simulator k1-cluster fflush(stdout); diff --git a/test/regression/Makefile b/test/regression/Makefile index 3447d6a5..ad3ffd99 100644 --- a/test/regression/Makefile +++ b/test/regression/Makefile @@ -26,18 +26,18 @@ TESTS_COMP?=attribs1 bitfields1 bitfields2 bitfields3 bitfields4 \ varargs1 varargs2 varargs3 sections alias aligned\ packedstruct1 packedstruct2 -ifeq ($(ARCH),mppa_k1c) +# FIXME ifeq ($(ARCH),mppa_k1c) TESTS_COMP:=$(filter-out packedstruct1,$(TESTS_COMP)) TESTS_COMP:=$(filter-out packedstruct2,$(TESTS_COMP)) -endif +# endif # Can run, both in compiled mode and in interpreter mode, # but produce processor-dependent results, so no reference output in Results TESTS_DIFF=NaNs -ifeq ($(ARCH),mppa_k1c) +# FIXME ifeq ($(ARCH),mppa_k1c) TESTS_DIFF:=$(filter-out NaNs,$(TESTS_DIFF)) -endif +# endif endif # Other tests: should compile to .s without errors (but expect warnings) -- cgit From e0843c84d1aa69ed0176cf44109b3bbdd4019504 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Fri, 27 Mar 2020 20:37:03 +0100 Subject: disable leaf function removal of return address restoration due to memcpy overwriting the return address register --- aarch64/Asmgen.v | 7 +++++-- aarch64/Asmgenproof.v | 4 ++-- aarch64/Asmgenproof1.v | 9 ++++++--- 3 files changed, 13 insertions(+), 7 deletions(-) diff --git a/aarch64/Asmgen.v b/aarch64/Asmgen.v index f4446696..024c9a17 100644 --- a/aarch64/Asmgen.v +++ b/aarch64/Asmgen.v @@ -1061,9 +1061,12 @@ Definition storeptr (src: ireg) (base: iregsp) (ofs: ptrofs) (k: code) := (** Function epilogue *) Definition make_epilogue (f: Mach.function) (k: code) := - if is_leaf_function f + (* FIXME + Cannot be used because memcpy destroys X30; + issue being discussed with X. Leroy *) + (* if is_leaf_function f then Pfreeframe f.(fn_stacksize) f.(fn_link_ofs) :: k - else loadptr XSP f.(fn_retaddr_ofs) RA + else*) loadptr XSP f.(fn_retaddr_ofs) RA (Pfreeframe f.(fn_stacksize) f.(fn_link_ofs) :: k). (** Translation of a Mach instruction. *) diff --git a/aarch64/Asmgenproof.v b/aarch64/Asmgenproof.v index 0d3179d4..6831509f 100644 --- a/aarch64/Asmgenproof.v +++ b/aarch64/Asmgenproof.v @@ -338,8 +338,8 @@ Remark make_epilogue_label: forall f k, tail_nolabel k (make_epilogue f k). Proof. unfold make_epilogue; intros. - destruct is_leaf_function. - { TailNoLabel. } + (* FIXME destruct is_leaf_function. + { TailNoLabel. } *) eapply tail_nolabel_trans. apply loadptr_label. TailNoLabel. diff --git a/aarch64/Asmgenproof1.v b/aarch64/Asmgenproof1.v index 2b89723f..0e36bd05 100644 --- a/aarch64/Asmgenproof1.v +++ b/aarch64/Asmgenproof1.v @@ -2070,7 +2070,7 @@ Lemma make_epilogue_correct: forall ge0 f m stk soff cs m' ms rs k tm, (is_leaf_function f = true -> rs # (IR RA) = parent_ra cs) -> load_stack m (Vptr stk soff) Tptr f.(fn_link_ofs) = Some (parent_sp cs) -> - (is_leaf_function f = false -> load_stack m (Vptr stk soff) Tptr f.(fn_retaddr_ofs) = Some (parent_ra cs)) -> + ((* FIXME is_leaf_function f = false -> *) 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 -> @@ -2084,6 +2084,9 @@ Lemma make_epilogue_correct: /\ (forall r, r <> PC -> r <> SP -> r <> RA -> r <> X16 -> rs'#r = rs#r). Proof. intros until tm; intros LEAF_RA LP LRA FREE AG MEXT MCS. + + (* FIXME + Cannot be used at this point destruct (is_leaf_function f) eqn:IS_LEAF. { exploit Mem.loadv_extends. eauto. eexact LP. auto. simpl. intros (parent' & LP' & LDP'). @@ -2106,14 +2109,14 @@ Proof. intros. Simpl. } lapply LRA. 2: reflexivity. - clear LRA. intro LRA. + clear LRA. intro LRA. *) 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 IS_LEAF. + (* FIXME rewrite IS_LEAF. *) exploit (loadptr_correct XSP (fn_retaddr_ofs f)). instantiate (2 := rs). simpl. rewrite <- (sp_val _ _ _ AG). simpl. eexact LRA'. simpl; congruence. intros (rs1 & A1 & B1 & C1). -- cgit From b5deca576e000cb8cabd9c3c036e8de83cbe2e37 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sat, 28 Mar 2020 08:08:47 +0100 Subject: Makefile for CI --- test/monniaux/yarpgen/Makefile | 65 ++++++++++++++++++++++++-------------- test/monniaux/yarpgen/Makefile.old | 52 ++++++++++++++++++++++++++++++ 2 files changed, 93 insertions(+), 24 deletions(-) create mode 100644 test/monniaux/yarpgen/Makefile.old diff --git a/test/monniaux/yarpgen/Makefile b/test/monniaux/yarpgen/Makefile index 9da82deb..a320a7e8 100644 --- a/test/monniaux/yarpgen/Makefile +++ b/test/monniaux/yarpgen/Makefile @@ -1,37 +1,54 @@ +TARGET_CCOMP=../../../ccomp + YARPGEN=yarpgen MAX=300 PREFIX=ran%06.f -include ../rules.mk -K1C_CCOMPFLAGS += -funprototyped -fbitfields -CCOMPFLAGS += -funprototyped -fbitfields +CCOMPOPTS += -funprototyped -fbitfields -TARGETS_C=$(shell seq --format $(PREFIX)/func.c 0 $(MAX)) \ +TESTS_C=$(shell seq --format $(PREFIX)/func.c 0 $(MAX)) \ $(shell seq --format $(PREFIX)/driver.c 0 $(MAX)) \ $(shell seq --format $(PREFIX)/init.h 0 $(MAX)) -TARGETS_CCOMP_K1C_S=$(shell seq --format $(PREFIX)/func.ccomp.k1c.s 0 $(MAX)) \ - $(shell seq --format $(PREFIX)/driver.ccomp.k1c.s 0 $(MAX)) -TARGETS_GCC_K1C_S=$(shell seq --format $(PREFIX)/func.gcc.k1c.s 0 $(MAX)) \ - $(shell seq --format $(PREFIX)/driver.gcc.k1c.s 0 $(MAX)) -TARGETS_CCOMP_HOST_S=$(shell seq --format $(PREFIX)/func.ccomp.host.s 0 $(MAX)) \ - $(shell seq --format $(PREFIX)/driver.ccomp.host.s 0 $(MAX)) -TARGETS_GCC_HOST_S=$(shell seq --format $(PREFIX)/func.gcc.host.s 0 $(MAX)) \ +TESTS_CCOMP_TARGET_S=$(shell seq --format $(PREFIX)/func.ccomp.target.s 0 $(MAX)) \ + $(shell seq --format $(PREFIX)/driver.ccomp.target.s 0 $(MAX)) +TESTS_GCC_TARGET_S=$(shell seq --format $(PREFIX)/func.gcc.target.s 0 $(MAX)) \ + $(shell seq --format $(PREFIX)/driver.gcc.target.s 0 $(MAX)) +TESTS_GCC_HOST_S=$(shell seq --format $(PREFIX)/func.gcc.host.s 0 $(MAX)) \ $(shell seq --format $(PREFIX)/driver.gcc.host.s 0 $(MAX)) -TARGETS_CCOMP_K1C_OUT=$(shell seq --format $(PREFIX)/example.ccomp.k1c.out 0 $(MAX)) -TARGETS_GCC_K1C_OUT=$(shell seq --format $(PREFIX)/example.gcc.k1c.out 0 $(MAX)) -TARGETS_GCC_HOST_OUT=$(shell seq --format $(PREFIX)/example.gcc.host.out 0 $(MAX)) -TARGETS_CCOMP_HOST_OUT=$(shell seq --format $(PREFIX)/example.ccomp.host.out 0 $(MAX)) -TARGETS_CMP=$(shell seq --format $(PREFIX)/example.k1c.cmp 0 $(MAX)) +TESTS_CCOMP_TARGET_OUT=$(shell seq --format $(PREFIX)/example.ccomp.target.out 0 $(MAX)) +TESTS_GCC_TARGET_OUT=$(shell seq --format $(PREFIX)/example.gcc.target.out 0 $(MAX)) +TESTS_GCC_HOST_OUT=$(shell seq --format $(PREFIX)/example.gcc.host.out 0 $(MAX)) +TESTS_CMP=$(shell seq --format $(PREFIX)/example.target.cmp 0 $(MAX)) + +all: $(TESTS_CCOMP_TARGET_OUT) $(TESTS_GCC_TARGET_OUT) $(TESTS_GCC_HOST_OUT) $(TESTS_CCOMP_TARGET_S) $(TESTS_GCC_TARGET_S) $(TESTS_GCC_HOST_S) $(TESTS_CMP) $(TESTS_C) + +tests_c: $(TESTS_C) + +%.ccomp.target.s : %.c + $(TARGET_CCOMP) $(CCOMPOPTS) -S -o $@ $< + +%.gcc.target.s : %.c + $(TARGET_CC) $(CCOMPOPTS) -S -o $@ $< + +%.gcc.host.s : %.c + $(CC) $(CFLAGS) -S -o $@ $< + +%.target.o : %.target.s + $(TARGET_CCOMP) $(CCOMPOPTS) -c -o $@ $< + +%.target.out : %.target + $(EXECUTE) $< > $@ -all: $(TARGETS_CCOMP_K1C_OUT) $(TARGETS_GCC_K1C_OUT) $(TARGETS_GCC_HOST_OUT) $(TARGETS_CCOMP_HOST_OUT) $(TARGETS_CCOMP_K1C_S) $(TARGETS_GCC_K1C_S) $(TARGETS_GCC_HOST_S) $(TARGETS_CCOMP_HOST_S) $(TARGETS_CMP) $(TARGETS_C) +%.host.out : %.host + ./$< > $@ -ran%/func.ccomp.k1c.s ran%/func.gcc.k1c.s ran%/func.ccomp.host.s ran%/func.gcc.host.s : ran%/init.h +ran%/func.ccomp.target.s ran%/func.gcc.target.s ran%/func.ccomp.host.s ran%/func.gcc.host.s : ran%/init.h -ran%/example.ccomp.k1c: ran%/func.ccomp.k1c.o ran%/driver.ccomp.k1c.o - $(K1C_CCOMP) $(K1C_CCOMPFLAGS) $+ -o $@ +ran%/example.ccomp.target: ran%/func.ccomp.target.o ran%/driver.ccomp.target.o + $(TARGET_CCOMP) $(CCOMPOPTS) $+ -o $@ -ran%/example.gcc.k1c: ran%/func.gcc.k1c.o ran%/driver.gcc.k1c.o - $(K1C_CC) $(K1C_CFLAGS) $+ -o $@ +ran%/example.gcc.target: ran%/func.gcc.target.o ran%/driver.gcc.target.o + $(TARGET_CC) $(TARGET_CFLAGS) $+ -o $@ ran%/example.gcc.host: ran%/func.gcc.host.o ran%/driver.gcc.host.o $(CC) $(CFLAGS) $+ -o $@ @@ -43,10 +60,10 @@ ran%/driver.c ran%/func.c ran%/init.h: -mkdir ran$* $(YARPGEN) --seed=$* --out-dir=ran$*/ --std=c99 -ran%/example.k1c.cmp : ran%/example.gcc.k1c.out ran%/example.ccomp.k1c.out +ran%/example.target.cmp : ran%/example.gcc.target.out ran%/example.ccomp.target.out cmp $+ > $@ -.PHONY: all clean +.PHONY: all clean tests_c clean: -rm -rf ran* diff --git a/test/monniaux/yarpgen/Makefile.old b/test/monniaux/yarpgen/Makefile.old new file mode 100644 index 00000000..9da82deb --- /dev/null +++ b/test/monniaux/yarpgen/Makefile.old @@ -0,0 +1,52 @@ +YARPGEN=yarpgen +MAX=300 +PREFIX=ran%06.f +include ../rules.mk + +K1C_CCOMPFLAGS += -funprototyped -fbitfields +CCOMPFLAGS += -funprototyped -fbitfields + +TARGETS_C=$(shell seq --format $(PREFIX)/func.c 0 $(MAX)) \ + $(shell seq --format $(PREFIX)/driver.c 0 $(MAX)) \ + $(shell seq --format $(PREFIX)/init.h 0 $(MAX)) +TARGETS_CCOMP_K1C_S=$(shell seq --format $(PREFIX)/func.ccomp.k1c.s 0 $(MAX)) \ + $(shell seq --format $(PREFIX)/driver.ccomp.k1c.s 0 $(MAX)) +TARGETS_GCC_K1C_S=$(shell seq --format $(PREFIX)/func.gcc.k1c.s 0 $(MAX)) \ + $(shell seq --format $(PREFIX)/driver.gcc.k1c.s 0 $(MAX)) +TARGETS_CCOMP_HOST_S=$(shell seq --format $(PREFIX)/func.ccomp.host.s 0 $(MAX)) \ + $(shell seq --format $(PREFIX)/driver.ccomp.host.s 0 $(MAX)) +TARGETS_GCC_HOST_S=$(shell seq --format $(PREFIX)/func.gcc.host.s 0 $(MAX)) \ + $(shell seq --format $(PREFIX)/driver.gcc.host.s 0 $(MAX)) +TARGETS_CCOMP_K1C_OUT=$(shell seq --format $(PREFIX)/example.ccomp.k1c.out 0 $(MAX)) +TARGETS_GCC_K1C_OUT=$(shell seq --format $(PREFIX)/example.gcc.k1c.out 0 $(MAX)) +TARGETS_GCC_HOST_OUT=$(shell seq --format $(PREFIX)/example.gcc.host.out 0 $(MAX)) +TARGETS_CCOMP_HOST_OUT=$(shell seq --format $(PREFIX)/example.ccomp.host.out 0 $(MAX)) +TARGETS_CMP=$(shell seq --format $(PREFIX)/example.k1c.cmp 0 $(MAX)) + +all: $(TARGETS_CCOMP_K1C_OUT) $(TARGETS_GCC_K1C_OUT) $(TARGETS_GCC_HOST_OUT) $(TARGETS_CCOMP_HOST_OUT) $(TARGETS_CCOMP_K1C_S) $(TARGETS_GCC_K1C_S) $(TARGETS_GCC_HOST_S) $(TARGETS_CCOMP_HOST_S) $(TARGETS_CMP) $(TARGETS_C) + +ran%/func.ccomp.k1c.s ran%/func.gcc.k1c.s ran%/func.ccomp.host.s ran%/func.gcc.host.s : ran%/init.h + +ran%/example.ccomp.k1c: ran%/func.ccomp.k1c.o ran%/driver.ccomp.k1c.o + $(K1C_CCOMP) $(K1C_CCOMPFLAGS) $+ -o $@ + +ran%/example.gcc.k1c: ran%/func.gcc.k1c.o ran%/driver.gcc.k1c.o + $(K1C_CC) $(K1C_CFLAGS) $+ -o $@ + +ran%/example.gcc.host: ran%/func.gcc.host.o ran%/driver.gcc.host.o + $(CC) $(CFLAGS) $+ -o $@ + +ran%/example.ccomp.host: ran%/func.ccomp.host.o ran%/driver.ccomp.host.o + $(CCOMP) $(CCOMPFLAGS) $+ -o $@ + +ran%/driver.c ran%/func.c ran%/init.h: + -mkdir ran$* + $(YARPGEN) --seed=$* --out-dir=ran$*/ --std=c99 + +ran%/example.k1c.cmp : ran%/example.gcc.k1c.out ran%/example.ccomp.k1c.out + cmp $+ > $@ + +.PHONY: all clean + +clean: + -rm -rf ran* -- cgit From 575a56ce631090624041db36ddb2747be907d091 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sat, 28 Mar 2020 08:12:12 +0100 Subject: cleaner make invocation --- .gitlab-ci.yml | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 17228f76..16055618 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -11,7 +11,7 @@ build_x86_64: script: - ./config_x86_64.sh - make -j "$NJOBS" - - cd test && make all test + - make -C test all test rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always @@ -33,7 +33,7 @@ build_ia32: script: - ./config_ia32.sh - make -j "$NJOBS" - - cd test && make all test + - make -C test all test rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always @@ -55,7 +55,7 @@ build_aarch64: script: - ./config_aarch64.sh - make -j "$NJOBS" - - cd test && make CCOMPOPTS=-static SIMU='qemu-aarch64' EXECUTE='qemu-aarch64' all test + - make -C test CCOMPOPTS=-static SIMU='qemu-aarch64' EXECUTE='qemu-aarch64' all test rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always @@ -77,7 +77,7 @@ build_arm: script: - ./config_arm.sh - make -j "$NJOBS" - - cd test && make DIRS="c compression spass regression" CCOMPOPTS=-static SIMU='qemu-arm' EXECUTE='qemu-arm' all test + - make -C test DIRS="c compression spass regression" CCOMPOPTS=-static SIMU='qemu-arm' EXECUTE='qemu-arm' all test rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always @@ -141,7 +141,7 @@ build_rv64: script: - ./config_rv64.sh - make -j "$NJOBS" - - cd test && make CCOMPOPTS=-static SIMU='qemu-riscv64' EXECUTE='qemu-riscv64' all test + - make -C test CCOMPOPTS=-static SIMU='qemu-riscv64' EXECUTE='qemu-riscv64' all test rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always -- cgit From 292241f51df20ee0f057f5d3f7cc00f1425f1727 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sat, 28 Mar 2020 09:23:51 +0100 Subject: set up for autogeneration of yarpgen --- test/monniaux/yarpgen/Makefile | 29 ++++++++++++++++++++++------- 1 file changed, 22 insertions(+), 7 deletions(-) diff --git a/test/monniaux/yarpgen/Makefile b/test/monniaux/yarpgen/Makefile index a320a7e8..02564fef 100644 --- a/test/monniaux/yarpgen/Makefile +++ b/test/monniaux/yarpgen/Makefile @@ -1,14 +1,21 @@ TARGET_CCOMP=../../../ccomp -YARPGEN=yarpgen +ifndef YARPGEN +YARPGEN=./yarpgen +generator: yarpgen +endif + MAX=300 PREFIX=ran%06.f -CCOMPOPTS += -funprototyped -fbitfields +CCOMPFLAGS+=-funprototyped -fbitfields TESTS_C=$(shell seq --format $(PREFIX)/func.c 0 $(MAX)) \ $(shell seq --format $(PREFIX)/driver.c 0 $(MAX)) \ $(shell seq --format $(PREFIX)/init.h 0 $(MAX)) + +$(TESTS_C): generator + TESTS_CCOMP_TARGET_S=$(shell seq --format $(PREFIX)/func.ccomp.target.s 0 $(MAX)) \ $(shell seq --format $(PREFIX)/driver.ccomp.target.s 0 $(MAX)) TESTS_GCC_TARGET_S=$(shell seq --format $(PREFIX)/func.gcc.target.s 0 $(MAX)) \ @@ -24,8 +31,10 @@ all: $(TESTS_CCOMP_TARGET_OUT) $(TESTS_GCC_TARGET_OUT) $(TESTS_GCC_HOST_OUT) $(T tests_c: $(TESTS_C) +tests_s: $(TESTS_CCOMP_TARGET_S) + %.ccomp.target.s : %.c - $(TARGET_CCOMP) $(CCOMPOPTS) -S -o $@ $< + $(TARGET_CCOMP) $(CCOMPOPTS) $(CCOMPFLAGS) -S -o $@ $< %.gcc.target.s : %.c $(TARGET_CC) $(CCOMPOPTS) -S -o $@ $< @@ -34,7 +43,7 @@ tests_c: $(TESTS_C) $(CC) $(CFLAGS) -S -o $@ $< %.target.o : %.target.s - $(TARGET_CCOMP) $(CCOMPOPTS) -c -o $@ $< + $(TARGET_CCOMP) $(CCOMPOPTS) $(CCOMPFLAGS) -c -o $@ $< %.target.out : %.target $(EXECUTE) $< > $@ @@ -45,7 +54,7 @@ tests_c: $(TESTS_C) ran%/func.ccomp.target.s ran%/func.gcc.target.s ran%/func.ccomp.host.s ran%/func.gcc.host.s : ran%/init.h ran%/example.ccomp.target: ran%/func.ccomp.target.o ran%/driver.ccomp.target.o - $(TARGET_CCOMP) $(CCOMPOPTS) $+ -o $@ + $(TARGET_CCOMP) $(CCOMPOPTS) $(CCOMPFLAGS) $+ -o $@ ran%/example.gcc.target: ran%/func.gcc.target.o ran%/driver.gcc.target.o $(TARGET_CC) $(TARGET_CFLAGS) $+ -o $@ @@ -57,13 +66,19 @@ ran%/example.ccomp.host: ran%/func.ccomp.host.o ran%/driver.ccomp.host.o $(CCOMP) $(CCOMPFLAGS) $+ -o $@ ran%/driver.c ran%/func.c ran%/init.h: - -mkdir ran$* + mkdir -p ran$* $(YARPGEN) --seed=$* --out-dir=ran$*/ --std=c99 ran%/example.target.cmp : ran%/example.gcc.target.out ran%/example.ccomp.target.out cmp $+ > $@ -.PHONY: all clean tests_c +yarpgen: + curl -L -o yarpgen_v1.1.tar.gz https://github.com/intel/yarpgen/archive/v1.1.tar.gz + tar xfz yarpgen_v1.1.tar.gz + $(MAKE) CXX=g++ -C yarpgen-1.1 + cp yarpgen-1.1/yarpgen $@ + +.PHONY: all clean tests_c tests_c generator clean: -rm -rf ran* -- cgit From 6dc7548e1e8dad708ad3348ecc324e02cd5f3472 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sat, 28 Mar 2020 09:34:22 +0100 Subject: fix Makefile for not remaking the generator --- test/monniaux/yarpgen/Makefile | 26 +++++++++++++------------- 1 file changed, 13 insertions(+), 13 deletions(-) diff --git a/test/monniaux/yarpgen/Makefile b/test/monniaux/yarpgen/Makefile index 02564fef..7a62ef61 100644 --- a/test/monniaux/yarpgen/Makefile +++ b/test/monniaux/yarpgen/Makefile @@ -2,7 +2,7 @@ TARGET_CCOMP=../../../ccomp ifndef YARPGEN YARPGEN=./yarpgen -generator: yarpgen +GENERATOR=yarpgen endif MAX=300 @@ -14,18 +14,18 @@ TESTS_C=$(shell seq --format $(PREFIX)/func.c 0 $(MAX)) \ $(shell seq --format $(PREFIX)/driver.c 0 $(MAX)) \ $(shell seq --format $(PREFIX)/init.h 0 $(MAX)) -$(TESTS_C): generator +$(TESTS_C): $(GENERATOR) -TESTS_CCOMP_TARGET_S=$(shell seq --format $(PREFIX)/func.ccomp.target.s 0 $(MAX)) \ - $(shell seq --format $(PREFIX)/driver.ccomp.target.s 0 $(MAX)) -TESTS_GCC_TARGET_S=$(shell seq --format $(PREFIX)/func.gcc.target.s 0 $(MAX)) \ - $(shell seq --format $(PREFIX)/driver.gcc.target.s 0 $(MAX)) -TESTS_GCC_HOST_S=$(shell seq --format $(PREFIX)/func.gcc.host.s 0 $(MAX)) \ - $(shell seq --format $(PREFIX)/driver.gcc.host.s 0 $(MAX)) -TESTS_CCOMP_TARGET_OUT=$(shell seq --format $(PREFIX)/example.ccomp.target.out 0 $(MAX)) -TESTS_GCC_TARGET_OUT=$(shell seq --format $(PREFIX)/example.gcc.target.out 0 $(MAX)) -TESTS_GCC_HOST_OUT=$(shell seq --format $(PREFIX)/example.gcc.host.out 0 $(MAX)) -TESTS_CMP=$(shell seq --format $(PREFIX)/example.target.cmp 0 $(MAX)) +TESTS_CCOMP_TARGET_S=$(shell seq --format $(PREFIX)/func.ccomp.target.s 1 $(MAX)) \ + $(shell seq --format $(PREFIX)/driver.ccomp.target.s 1 $(MAX)) +TESTS_GCC_TARGET_S=$(shell seq --format $(PREFIX)/func.gcc.target.s 1 $(MAX)) \ + $(shell seq --format $(PREFIX)/driver.gcc.target.s 1 $(MAX)) +TESTS_GCC_HOST_S=$(shell seq --format $(PREFIX)/func.gcc.host.s 1 $(MAX)) \ + $(shell seq --format $(PREFIX)/driver.gcc.host.s 1 $(MAX)) +TESTS_CCOMP_TARGET_OUT=$(shell seq --format $(PREFIX)/example.ccomp.target.out 1 $(MAX)) +TESTS_GCC_TARGET_OUT=$(shell seq --format $(PREFIX)/example.gcc.target.out 1 $(MAX)) +TESTS_GCC_HOST_OUT=$(shell seq --format $(PREFIX)/example.gcc.host.out 1 $(MAX)) +TESTS_CMP=$(shell seq --format $(PREFIX)/example.target.cmp 1 $(MAX)) all: $(TESTS_CCOMP_TARGET_OUT) $(TESTS_GCC_TARGET_OUT) $(TESTS_GCC_HOST_OUT) $(TESTS_CCOMP_TARGET_S) $(TESTS_GCC_TARGET_S) $(TESTS_GCC_HOST_S) $(TESTS_CMP) $(TESTS_C) @@ -78,7 +78,7 @@ yarpgen: $(MAKE) CXX=g++ -C yarpgen-1.1 cp yarpgen-1.1/yarpgen $@ -.PHONY: all clean tests_c tests_c generator +.PHONY: all clean tests_c tests_c clean: -rm -rf ran* -- cgit From 3e5fcc7e1bea051e2f14f7b3a20d4e78cb23e539 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sat, 28 Mar 2020 10:03:26 +0100 Subject: fix Makefile (again) --- test/monniaux/yarpgen/Makefile | 19 ++++++++++--------- 1 file changed, 10 insertions(+), 9 deletions(-) diff --git a/test/monniaux/yarpgen/Makefile b/test/monniaux/yarpgen/Makefile index 7a62ef61..ffa58172 100644 --- a/test/monniaux/yarpgen/Makefile +++ b/test/monniaux/yarpgen/Makefile @@ -1,4 +1,7 @@ TARGET_CCOMP=../../../ccomp +TARGET_CC=gcc + +all: ifndef YARPGEN YARPGEN=./yarpgen @@ -10,9 +13,10 @@ PREFIX=ran%06.f CCOMPFLAGS+=-funprototyped -fbitfields -TESTS_C=$(shell seq --format $(PREFIX)/func.c 0 $(MAX)) \ - $(shell seq --format $(PREFIX)/driver.c 0 $(MAX)) \ - $(shell seq --format $(PREFIX)/init.h 0 $(MAX)) +TESTS_C=$(shell seq --format $(PREFIX)/func.c 1 $(MAX)) \ + $(shell seq --format $(PREFIX)/driver.c 1 $(MAX)) \ + $(shell seq --format $(PREFIX)/init.c 1 $(MAX)) \ + $(shell seq --format $(PREFIX)/init.h 1 $(MAX)) $(TESTS_C): $(GENERATOR) @@ -53,18 +57,15 @@ tests_s: $(TESTS_CCOMP_TARGET_S) ran%/func.ccomp.target.s ran%/func.gcc.target.s ran%/func.ccomp.host.s ran%/func.gcc.host.s : ran%/init.h -ran%/example.ccomp.target: ran%/func.ccomp.target.o ran%/driver.ccomp.target.o +ran%/example.ccomp.target: ran%/func.ccomp.target.o ran%/driver.ccomp.target.o ran%/init.ccomp.target.o ran%/check.ccomp.target.o ran%/hash.ccomp.target.o $(TARGET_CCOMP) $(CCOMPOPTS) $(CCOMPFLAGS) $+ -o $@ -ran%/example.gcc.target: ran%/func.gcc.target.o ran%/driver.gcc.target.o +ran%/example.gcc.target: ran%/func.gcc.target.o ran%/driver.gcc.target.o ran%/init.ccomp.target.o ran%/check.gcc.target.o ran%/hash.gcc.target.o $(TARGET_CC) $(TARGET_CFLAGS) $+ -o $@ -ran%/example.gcc.host: ran%/func.gcc.host.o ran%/driver.gcc.host.o +ran%/example.gcc.host: ran%/func.gcc.host.o ran%/driver.gcc.host.o ran%/init.gcc.host.o ran%/check.gcc.host.o ran%/hash.gcc.host.o $(CC) $(CFLAGS) $+ -o $@ -ran%/example.ccomp.host: ran%/func.ccomp.host.o ran%/driver.ccomp.host.o - $(CCOMP) $(CCOMPFLAGS) $+ -o $@ - ran%/driver.c ran%/func.c ran%/init.h: mkdir -p ran$* $(YARPGEN) --seed=$* --out-dir=ran$*/ --std=c99 -- cgit From 0024edef4ce251a154733a241868ecc3119c8adf Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sat, 28 Mar 2020 10:19:25 +0100 Subject: fix targets for proper generation --- test/monniaux/yarpgen/Makefile | 7 +++++-- 1 file changed, 5 insertions(+), 2 deletions(-) diff --git a/test/monniaux/yarpgen/Makefile b/test/monniaux/yarpgen/Makefile index ffa58172..8023902b 100644 --- a/test/monniaux/yarpgen/Makefile +++ b/test/monniaux/yarpgen/Makefile @@ -29,7 +29,7 @@ TESTS_GCC_HOST_S=$(shell seq --format $(PREFIX)/func.gcc.host.s 1 $(MAX)) \ TESTS_CCOMP_TARGET_OUT=$(shell seq --format $(PREFIX)/example.ccomp.target.out 1 $(MAX)) TESTS_GCC_TARGET_OUT=$(shell seq --format $(PREFIX)/example.gcc.target.out 1 $(MAX)) TESTS_GCC_HOST_OUT=$(shell seq --format $(PREFIX)/example.gcc.host.out 1 $(MAX)) -TESTS_CMP=$(shell seq --format $(PREFIX)/example.target.cmp 1 $(MAX)) +TESTS_CMP=$(shell seq --format $(PREFIX)/example.target.cmp 1 $(MAX)) $(shell seq --format $(PREFIX)/example.host_target.cmp 1 $(MAX)) all: $(TESTS_CCOMP_TARGET_OUT) $(TESTS_GCC_TARGET_OUT) $(TESTS_GCC_HOST_OUT) $(TESTS_CCOMP_TARGET_S) $(TESTS_GCC_TARGET_S) $(TESTS_GCC_HOST_S) $(TESTS_CMP) $(TESTS_C) @@ -66,13 +66,16 @@ ran%/example.gcc.target: ran%/func.gcc.target.o ran%/driver.gcc.target.o ran%/in ran%/example.gcc.host: ran%/func.gcc.host.o ran%/driver.gcc.host.o ran%/init.gcc.host.o ran%/check.gcc.host.o ran%/hash.gcc.host.o $(CC) $(CFLAGS) $+ -o $@ -ran%/driver.c ran%/func.c ran%/init.h: +ran%/driver.c ran%/func.c ran%/init.c ran%/check.c ran%/hash.c ran%/init.h: mkdir -p ran$* $(YARPGEN) --seed=$* --out-dir=ran$*/ --std=c99 ran%/example.target.cmp : ran%/example.gcc.target.out ran%/example.ccomp.target.out cmp $+ > $@ +ran%/example.host_target.cmp : ran%/example.gcc.host.out ran%/example.gcc.target.out + cmp $+ > $@ + yarpgen: curl -L -o yarpgen_v1.1.tar.gz https://github.com/intel/yarpgen/archive/v1.1.tar.gz tar xfz yarpgen_v1.1.tar.gz -- cgit From e1fb8ff636881054d7223129d61b214cce36acf7 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sat, 28 Mar 2020 10:48:51 +0100 Subject: some more Makefile fixes (disable cse2 it's too slow) --- test/monniaux/yarpgen/Makefile | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/test/monniaux/yarpgen/Makefile b/test/monniaux/yarpgen/Makefile index 8023902b..f94dffce 100644 --- a/test/monniaux/yarpgen/Makefile +++ b/test/monniaux/yarpgen/Makefile @@ -11,7 +11,7 @@ endif MAX=300 PREFIX=ran%06.f -CCOMPFLAGS+=-funprototyped -fbitfields +CCOMPFLAGS+=-funprototyped -fbitfields -fno-cse2 # FIXME TESTS_C=$(shell seq --format $(PREFIX)/func.c 1 $(MAX)) \ $(shell seq --format $(PREFIX)/driver.c 1 $(MAX)) \ -- cgit From 5fe8b78a18f3bd8b4ad80c7318115d5d2ebe932f Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sat, 28 Mar 2020 11:07:09 +0100 Subject: run yarpgen test on aarch64 --- .gitlab-ci.yml | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 16055618..ee50b751 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -55,7 +55,8 @@ build_aarch64: script: - ./config_aarch64.sh - make -j "$NJOBS" - - make -C test CCOMPOPTS=-static SIMU='qemu-aarch64' EXECUTE='qemu-aarch64' all test + - make -C test CCOMPOPTS='-static' SIMU='qemu-aarch64' EXECUTE='qemu-aarch64' all test + - make -C test/monniaux/yarpgen -k TARGET_CC='aarch64-linux-gnu-gcc' EXECUTE='qemu-aarch64' CCOMPOPTS='-static' TARGET_CFLAGS='-static' rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always -- cgit From 2569e84cd235045e4419e8d65c0e69bb3f2ffb60 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sat, 28 Mar 2020 11:26:22 +0100 Subject: stdlib path --- test/monniaux/yarpgen/Makefile | 8 +++++--- 1 file changed, 5 insertions(+), 3 deletions(-) diff --git a/test/monniaux/yarpgen/Makefile b/test/monniaux/yarpgen/Makefile index f94dffce..bdb7cb63 100644 --- a/test/monniaux/yarpgen/Makefile +++ b/test/monniaux/yarpgen/Makefile @@ -11,7 +11,8 @@ endif MAX=300 PREFIX=ran%06.f -CCOMPFLAGS+=-funprototyped -fbitfields -fno-cse2 # FIXME +CCOMPOPTS=-static +CCOMPFLAGS+=-funprototyped -fbitfields -fno-cse2 -stdlib ../../../runtime # FIXME TESTS_C=$(shell seq --format $(PREFIX)/func.c 1 $(MAX)) \ $(shell seq --format $(PREFIX)/driver.c 1 $(MAX)) \ @@ -29,7 +30,7 @@ TESTS_GCC_HOST_S=$(shell seq --format $(PREFIX)/func.gcc.host.s 1 $(MAX)) \ TESTS_CCOMP_TARGET_OUT=$(shell seq --format $(PREFIX)/example.ccomp.target.out 1 $(MAX)) TESTS_GCC_TARGET_OUT=$(shell seq --format $(PREFIX)/example.gcc.target.out 1 $(MAX)) TESTS_GCC_HOST_OUT=$(shell seq --format $(PREFIX)/example.gcc.host.out 1 $(MAX)) -TESTS_CMP=$(shell seq --format $(PREFIX)/example.target.cmp 1 $(MAX)) $(shell seq --format $(PREFIX)/example.host_target.cmp 1 $(MAX)) +TESTS_CMP=$(shell seq --format $(PREFIX)/example.host_target.cmp 1 $(MAX)) $(shell seq --format $(PREFIX)/example.target.cmp 1 $(MAX)) all: $(TESTS_CCOMP_TARGET_OUT) $(TESTS_GCC_TARGET_OUT) $(TESTS_GCC_HOST_OUT) $(TESTS_CCOMP_TARGET_S) $(TESTS_GCC_TARGET_S) $(TESTS_GCC_HOST_S) $(TESTS_CMP) $(TESTS_C) @@ -73,7 +74,7 @@ ran%/driver.c ran%/func.c ran%/init.c ran%/check.c ran%/hash.c ran%/init.h: ran%/example.target.cmp : ran%/example.gcc.target.out ran%/example.ccomp.target.out cmp $+ > $@ -ran%/example.host_target.cmp : ran%/example.gcc.host.out ran%/example.gcc.target.out +ran%/example.host_target.cmp : ran%/example.gcc.host.out ran%/example.ccomp.target.out cmp $+ > $@ yarpgen: @@ -83,6 +84,7 @@ yarpgen: cp yarpgen-1.1/yarpgen $@ .PHONY: all clean tests_c tests_c +.SECONDARY: .s .target .out clean: -rm -rf ran* -- cgit From aa52be6e57fb0627113cfcaf26340e323f58b3bc Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sat, 28 Mar 2020 11:55:04 +0100 Subject: fix limitxy --- test/monniaux/yarpgen/Makefile | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/test/monniaux/yarpgen/Makefile b/test/monniaux/yarpgen/Makefile index bdb7cb63..2aafd394 100644 --- a/test/monniaux/yarpgen/Makefile +++ b/test/monniaux/yarpgen/Makefile @@ -8,7 +8,7 @@ YARPGEN=./yarpgen GENERATOR=yarpgen endif -MAX=300 +MAX=170 PREFIX=ran%06.f CCOMPOPTS=-static -- cgit From dd76ee0e14acbff78b6cd575e53d9c9d59fa6747 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sat, 28 Mar 2020 12:17:41 +0100 Subject: better assemble with gcc --- test/monniaux/yarpgen/Makefile | 5 +++-- 1 file changed, 3 insertions(+), 2 deletions(-) diff --git a/test/monniaux/yarpgen/Makefile b/test/monniaux/yarpgen/Makefile index 2aafd394..9800d9f0 100644 --- a/test/monniaux/yarpgen/Makefile +++ b/test/monniaux/yarpgen/Makefile @@ -17,6 +17,8 @@ CCOMPFLAGS+=-funprototyped -fbitfields -fno-cse2 -stdlib ../../../runtime # FIXM TESTS_C=$(shell seq --format $(PREFIX)/func.c 1 $(MAX)) \ $(shell seq --format $(PREFIX)/driver.c 1 $(MAX)) \ $(shell seq --format $(PREFIX)/init.c 1 $(MAX)) \ + $(shell seq --format $(PREFIX)/hash.c 1 $(MAX)) \ + $(shell seq --format $(PREFIX)/check.c 1 $(MAX)) \ $(shell seq --format $(PREFIX)/init.h 1 $(MAX)) $(TESTS_C): $(GENERATOR) @@ -48,7 +50,7 @@ tests_s: $(TESTS_CCOMP_TARGET_S) $(CC) $(CFLAGS) -S -o $@ $< %.target.o : %.target.s - $(TARGET_CCOMP) $(CCOMPOPTS) $(CCOMPFLAGS) -c -o $@ $< + $(TARGET_CC) -c -o $@ $< %.target.out : %.target $(EXECUTE) $< > $@ @@ -84,7 +86,6 @@ yarpgen: cp yarpgen-1.1/yarpgen $@ .PHONY: all clean tests_c tests_c -.SECONDARY: .s .target .out clean: -rm -rf ran* -- cgit From 8beb9b73085b2ad49dfa68a74ac3812b78d6e558 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sat, 28 Mar 2020 12:33:54 +0100 Subject: fix inconsistency --- test/monniaux/yarpgen/Makefile | 15 +++++++-------- 1 file changed, 7 insertions(+), 8 deletions(-) diff --git a/test/monniaux/yarpgen/Makefile b/test/monniaux/yarpgen/Makefile index 9800d9f0..af586ac4 100644 --- a/test/monniaux/yarpgen/Makefile +++ b/test/monniaux/yarpgen/Makefile @@ -3,6 +3,8 @@ TARGET_CC=gcc all: +.SECONDARY: + ifndef YARPGEN YARPGEN=./yarpgen GENERATOR=yarpgen @@ -23,12 +25,9 @@ TESTS_C=$(shell seq --format $(PREFIX)/func.c 1 $(MAX)) \ $(TESTS_C): $(GENERATOR) -TESTS_CCOMP_TARGET_S=$(shell seq --format $(PREFIX)/func.ccomp.target.s 1 $(MAX)) \ - $(shell seq --format $(PREFIX)/driver.ccomp.target.s 1 $(MAX)) -TESTS_GCC_TARGET_S=$(shell seq --format $(PREFIX)/func.gcc.target.s 1 $(MAX)) \ - $(shell seq --format $(PREFIX)/driver.gcc.target.s 1 $(MAX)) -TESTS_GCC_HOST_S=$(shell seq --format $(PREFIX)/func.gcc.host.s 1 $(MAX)) \ - $(shell seq --format $(PREFIX)/driver.gcc.host.s 1 $(MAX)) +TESTS_CCOMP_TARGET_S=$(TEST_C:.c=.ccomp.target.s) +TESTS_GCC_TARGET_S=$(TEST_C:.c=.gcc.target.s) +TESTS_GCC_HOST_S=$(TEST_C:.c=.gcc.host.s) TESTS_CCOMP_TARGET_OUT=$(shell seq --format $(PREFIX)/example.ccomp.target.out 1 $(MAX)) TESTS_GCC_TARGET_OUT=$(shell seq --format $(PREFIX)/example.gcc.target.out 1 $(MAX)) TESTS_GCC_HOST_OUT=$(shell seq --format $(PREFIX)/example.gcc.host.out 1 $(MAX)) @@ -58,12 +57,12 @@ tests_s: $(TESTS_CCOMP_TARGET_S) %.host.out : %.host ./$< > $@ -ran%/func.ccomp.target.s ran%/func.gcc.target.s ran%/func.ccomp.host.s ran%/func.gcc.host.s : ran%/init.h +ran%/func.ccomp.target.s ran%/func.gcc.target.s ran%/func.ccomp.host.s ran%/func.gcc.host.s ran%/init.gcc.host.s : ran%/init.h ran%/example.ccomp.target: ran%/func.ccomp.target.o ran%/driver.ccomp.target.o ran%/init.ccomp.target.o ran%/check.ccomp.target.o ran%/hash.ccomp.target.o $(TARGET_CCOMP) $(CCOMPOPTS) $(CCOMPFLAGS) $+ -o $@ -ran%/example.gcc.target: ran%/func.gcc.target.o ran%/driver.gcc.target.o ran%/init.ccomp.target.o ran%/check.gcc.target.o ran%/hash.gcc.target.o +ran%/example.gcc.target: ran%/func.gcc.target.o ran%/driver.gcc.target.o ran%/init.gcc.target.o ran%/check.gcc.target.o ran%/hash.gcc.target.o $(TARGET_CC) $(TARGET_CFLAGS) $+ -o $@ ran%/example.gcc.host: ran%/func.gcc.host.o ran%/driver.gcc.host.o ran%/init.gcc.host.o ran%/check.gcc.host.o ran%/hash.gcc.host.o -- cgit From 520ac43fdfaa6ecef14fc93b68034e03112e12fd Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sat, 28 Mar 2020 13:23:51 +0100 Subject: run yarpgen on other architectures --- .gitlab-ci.yml | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index ee50b751..bf7361f8 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -12,6 +12,7 @@ build_x86_64: - ./config_x86_64.sh - make -j "$NJOBS" - make -C test all test + - make -C test/monniaux/yarpgen rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always @@ -34,6 +35,7 @@ build_ia32: - ./config_ia32.sh - make -j "$NJOBS" - make -C test all test + - make -C test/monniaux/yarpgen rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always @@ -79,6 +81,7 @@ build_arm: - ./config_arm.sh - make -j "$NJOBS" - make -C test DIRS="c compression spass regression" CCOMPOPTS=-static SIMU='qemu-arm' EXECUTE='qemu-arm' all test + - make -C test/monniaux/yarpgen -k TARGET_CC='arm-linux-gnu-gcc' EXECUTE='qemu-arm' CCOMPOPTS='-static' TARGET_CFLAGS='-static' rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always @@ -143,6 +146,7 @@ build_rv64: - ./config_rv64.sh - make -j "$NJOBS" - make -C test CCOMPOPTS=-static SIMU='qemu-riscv64' EXECUTE='qemu-riscv64' all test + - make -C test/monniaux/yarpgen -k TARGET_CC='riscv64-linux-gnu-gcc' EXECUTE='qemu-riscv64' CCOMPOPTS='-static' TARGET_CFLAGS='-static' rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always -- cgit From d0b8ed2ab2979bfef689c5c801a73434b0abab51 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sat, 28 Mar 2020 13:56:19 +0100 Subject: more fixes for CI --- .gitlab-ci.yml | 4 ++-- test/monniaux/yarpgen/Makefile | 2 +- 2 files changed, 3 insertions(+), 3 deletions(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index bf7361f8..119fa44e 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -35,7 +35,7 @@ build_ia32: - ./config_ia32.sh - make -j "$NJOBS" - make -C test all test - - make -C test/monniaux/yarpgen +# - make -C test/monniaux/yarpgen disabled due to -m32/-m64 issues rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always @@ -81,7 +81,7 @@ build_arm: - ./config_arm.sh - make -j "$NJOBS" - make -C test DIRS="c compression spass regression" CCOMPOPTS=-static SIMU='qemu-arm' EXECUTE='qemu-arm' all test - - make -C test/monniaux/yarpgen -k TARGET_CC='arm-linux-gnu-gcc' EXECUTE='qemu-arm' CCOMPOPTS='-static' TARGET_CFLAGS='-static' + - make -C test/monniaux/yarpgen -k TARGET_CC='arm-linux-gnueabihf-gcc' EXECUTE='qemu-arm' CCOMPOPTS='-static' TARGET_CFLAGS='-static' rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always diff --git a/test/monniaux/yarpgen/Makefile b/test/monniaux/yarpgen/Makefile index af586ac4..fc524d92 100644 --- a/test/monniaux/yarpgen/Makefile +++ b/test/monniaux/yarpgen/Makefile @@ -10,7 +10,7 @@ YARPGEN=./yarpgen GENERATOR=yarpgen endif -MAX=170 +MAX=169 PREFIX=ran%06.f CCOMPOPTS=-static -- cgit From 69247e1024f99f628f9b00eb9ecc30ef30e51d3f Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sat, 28 Mar 2020 14:27:13 +0100 Subject: limit due to stack overflows --- test/monniaux/yarpgen/Makefile | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/test/monniaux/yarpgen/Makefile b/test/monniaux/yarpgen/Makefile index fc524d92..dbd6ae75 100644 --- a/test/monniaux/yarpgen/Makefile +++ b/test/monniaux/yarpgen/Makefile @@ -10,7 +10,7 @@ YARPGEN=./yarpgen GENERATOR=yarpgen endif -MAX=169 +MAX=129 PREFIX=ran%06.f CCOMPOPTS=-static -- cgit From 567b0fae43cfe39cfbc15adaf5c31c62a02190ae Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sat, 28 Mar 2020 15:14:02 +0100 Subject: rm yarpgen on arm --- .gitlab-ci.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 119fa44e..5f696257 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -81,7 +81,7 @@ build_arm: - ./config_arm.sh - make -j "$NJOBS" - make -C test DIRS="c compression spass regression" CCOMPOPTS=-static SIMU='qemu-arm' EXECUTE='qemu-arm' all test - - make -C test/monniaux/yarpgen -k TARGET_CC='arm-linux-gnueabihf-gcc' EXECUTE='qemu-arm' CCOMPOPTS='-static' TARGET_CFLAGS='-static' +# - make -C test/monniaux/yarpgen -k TARGET_CC='arm-linux-gnueabihf-gcc' EXECUTE='qemu-arm' CCOMPOPTS='-static' TARGET_CFLAGS='-static' disabled; mysterious differences between gcc/clang and compcert rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always -- cgit From 35c545b3b2711645452b747b3d75b4f46a078776 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sat, 28 Mar 2020 15:27:25 +0100 Subject: Run Yarpgen in CI --- .gitlab-ci.yml | 15 +++-- test/monniaux/yarpgen/Makefile | 122 ++++++++++++++++++++++++------------- test/monniaux/yarpgen/Makefile.old | 52 ++++++++++++++++ 3 files changed, 142 insertions(+), 47 deletions(-) create mode 100644 test/monniaux/yarpgen/Makefile.old diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 17228f76..5f696257 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -11,7 +11,8 @@ build_x86_64: script: - ./config_x86_64.sh - make -j "$NJOBS" - - cd test && make all test + - make -C test all test + - make -C test/monniaux/yarpgen rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always @@ -33,7 +34,8 @@ build_ia32: script: - ./config_ia32.sh - make -j "$NJOBS" - - cd test && make all test + - make -C test all test +# - make -C test/monniaux/yarpgen disabled due to -m32/-m64 issues rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always @@ -55,7 +57,8 @@ build_aarch64: script: - ./config_aarch64.sh - make -j "$NJOBS" - - cd test && make CCOMPOPTS=-static SIMU='qemu-aarch64' EXECUTE='qemu-aarch64' all test + - make -C test CCOMPOPTS='-static' SIMU='qemu-aarch64' EXECUTE='qemu-aarch64' all test + - make -C test/monniaux/yarpgen -k TARGET_CC='aarch64-linux-gnu-gcc' EXECUTE='qemu-aarch64' CCOMPOPTS='-static' TARGET_CFLAGS='-static' rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always @@ -77,7 +80,8 @@ build_arm: script: - ./config_arm.sh - make -j "$NJOBS" - - cd test && make DIRS="c compression spass regression" CCOMPOPTS=-static SIMU='qemu-arm' EXECUTE='qemu-arm' all test + - make -C test DIRS="c compression spass regression" CCOMPOPTS=-static SIMU='qemu-arm' EXECUTE='qemu-arm' all test +# - make -C test/monniaux/yarpgen -k TARGET_CC='arm-linux-gnueabihf-gcc' EXECUTE='qemu-arm' CCOMPOPTS='-static' TARGET_CFLAGS='-static' disabled; mysterious differences between gcc/clang and compcert rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always @@ -141,7 +145,8 @@ build_rv64: script: - ./config_rv64.sh - make -j "$NJOBS" - - cd test && make CCOMPOPTS=-static SIMU='qemu-riscv64' EXECUTE='qemu-riscv64' all test + - make -C test CCOMPOPTS=-static SIMU='qemu-riscv64' EXECUTE='qemu-riscv64' all test + - make -C test/monniaux/yarpgen -k TARGET_CC='riscv64-linux-gnu-gcc' EXECUTE='qemu-riscv64' CCOMPOPTS='-static' TARGET_CFLAGS='-static' rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always diff --git a/test/monniaux/yarpgen/Makefile b/test/monniaux/yarpgen/Makefile index 9da82deb..dbd6ae75 100644 --- a/test/monniaux/yarpgen/Makefile +++ b/test/monniaux/yarpgen/Makefile @@ -1,52 +1,90 @@ -YARPGEN=yarpgen -MAX=300 +TARGET_CCOMP=../../../ccomp +TARGET_CC=gcc + +all: + +.SECONDARY: + +ifndef YARPGEN +YARPGEN=./yarpgen +GENERATOR=yarpgen +endif + +MAX=129 PREFIX=ran%06.f -include ../rules.mk - -K1C_CCOMPFLAGS += -funprototyped -fbitfields -CCOMPFLAGS += -funprototyped -fbitfields - -TARGETS_C=$(shell seq --format $(PREFIX)/func.c 0 $(MAX)) \ - $(shell seq --format $(PREFIX)/driver.c 0 $(MAX)) \ - $(shell seq --format $(PREFIX)/init.h 0 $(MAX)) -TARGETS_CCOMP_K1C_S=$(shell seq --format $(PREFIX)/func.ccomp.k1c.s 0 $(MAX)) \ - $(shell seq --format $(PREFIX)/driver.ccomp.k1c.s 0 $(MAX)) -TARGETS_GCC_K1C_S=$(shell seq --format $(PREFIX)/func.gcc.k1c.s 0 $(MAX)) \ - $(shell seq --format $(PREFIX)/driver.gcc.k1c.s 0 $(MAX)) -TARGETS_CCOMP_HOST_S=$(shell seq --format $(PREFIX)/func.ccomp.host.s 0 $(MAX)) \ - $(shell seq --format $(PREFIX)/driver.ccomp.host.s 0 $(MAX)) -TARGETS_GCC_HOST_S=$(shell seq --format $(PREFIX)/func.gcc.host.s 0 $(MAX)) \ - $(shell seq --format $(PREFIX)/driver.gcc.host.s 0 $(MAX)) -TARGETS_CCOMP_K1C_OUT=$(shell seq --format $(PREFIX)/example.ccomp.k1c.out 0 $(MAX)) -TARGETS_GCC_K1C_OUT=$(shell seq --format $(PREFIX)/example.gcc.k1c.out 0 $(MAX)) -TARGETS_GCC_HOST_OUT=$(shell seq --format $(PREFIX)/example.gcc.host.out 0 $(MAX)) -TARGETS_CCOMP_HOST_OUT=$(shell seq --format $(PREFIX)/example.ccomp.host.out 0 $(MAX)) -TARGETS_CMP=$(shell seq --format $(PREFIX)/example.k1c.cmp 0 $(MAX)) - -all: $(TARGETS_CCOMP_K1C_OUT) $(TARGETS_GCC_K1C_OUT) $(TARGETS_GCC_HOST_OUT) $(TARGETS_CCOMP_HOST_OUT) $(TARGETS_CCOMP_K1C_S) $(TARGETS_GCC_K1C_S) $(TARGETS_GCC_HOST_S) $(TARGETS_CCOMP_HOST_S) $(TARGETS_CMP) $(TARGETS_C) - -ran%/func.ccomp.k1c.s ran%/func.gcc.k1c.s ran%/func.ccomp.host.s ran%/func.gcc.host.s : ran%/init.h - -ran%/example.ccomp.k1c: ran%/func.ccomp.k1c.o ran%/driver.ccomp.k1c.o - $(K1C_CCOMP) $(K1C_CCOMPFLAGS) $+ -o $@ - -ran%/example.gcc.k1c: ran%/func.gcc.k1c.o ran%/driver.gcc.k1c.o - $(K1C_CC) $(K1C_CFLAGS) $+ -o $@ - -ran%/example.gcc.host: ran%/func.gcc.host.o ran%/driver.gcc.host.o - $(CC) $(CFLAGS) $+ -o $@ -ran%/example.ccomp.host: ran%/func.ccomp.host.o ran%/driver.ccomp.host.o - $(CCOMP) $(CCOMPFLAGS) $+ -o $@ +CCOMPOPTS=-static +CCOMPFLAGS+=-funprototyped -fbitfields -fno-cse2 -stdlib ../../../runtime # FIXME + +TESTS_C=$(shell seq --format $(PREFIX)/func.c 1 $(MAX)) \ + $(shell seq --format $(PREFIX)/driver.c 1 $(MAX)) \ + $(shell seq --format $(PREFIX)/init.c 1 $(MAX)) \ + $(shell seq --format $(PREFIX)/hash.c 1 $(MAX)) \ + $(shell seq --format $(PREFIX)/check.c 1 $(MAX)) \ + $(shell seq --format $(PREFIX)/init.h 1 $(MAX)) + +$(TESTS_C): $(GENERATOR) + +TESTS_CCOMP_TARGET_S=$(TEST_C:.c=.ccomp.target.s) +TESTS_GCC_TARGET_S=$(TEST_C:.c=.gcc.target.s) +TESTS_GCC_HOST_S=$(TEST_C:.c=.gcc.host.s) +TESTS_CCOMP_TARGET_OUT=$(shell seq --format $(PREFIX)/example.ccomp.target.out 1 $(MAX)) +TESTS_GCC_TARGET_OUT=$(shell seq --format $(PREFIX)/example.gcc.target.out 1 $(MAX)) +TESTS_GCC_HOST_OUT=$(shell seq --format $(PREFIX)/example.gcc.host.out 1 $(MAX)) +TESTS_CMP=$(shell seq --format $(PREFIX)/example.host_target.cmp 1 $(MAX)) $(shell seq --format $(PREFIX)/example.target.cmp 1 $(MAX)) + +all: $(TESTS_CCOMP_TARGET_OUT) $(TESTS_GCC_TARGET_OUT) $(TESTS_GCC_HOST_OUT) $(TESTS_CCOMP_TARGET_S) $(TESTS_GCC_TARGET_S) $(TESTS_GCC_HOST_S) $(TESTS_CMP) $(TESTS_C) + +tests_c: $(TESTS_C) + +tests_s: $(TESTS_CCOMP_TARGET_S) + +%.ccomp.target.s : %.c + $(TARGET_CCOMP) $(CCOMPOPTS) $(CCOMPFLAGS) -S -o $@ $< + +%.gcc.target.s : %.c + $(TARGET_CC) $(CCOMPOPTS) -S -o $@ $< -ran%/driver.c ran%/func.c ran%/init.h: - -mkdir ran$* +%.gcc.host.s : %.c + $(CC) $(CFLAGS) -S -o $@ $< + +%.target.o : %.target.s + $(TARGET_CC) -c -o $@ $< + +%.target.out : %.target + $(EXECUTE) $< > $@ + +%.host.out : %.host + ./$< > $@ + +ran%/func.ccomp.target.s ran%/func.gcc.target.s ran%/func.ccomp.host.s ran%/func.gcc.host.s ran%/init.gcc.host.s : ran%/init.h + +ran%/example.ccomp.target: ran%/func.ccomp.target.o ran%/driver.ccomp.target.o ran%/init.ccomp.target.o ran%/check.ccomp.target.o ran%/hash.ccomp.target.o + $(TARGET_CCOMP) $(CCOMPOPTS) $(CCOMPFLAGS) $+ -o $@ + +ran%/example.gcc.target: ran%/func.gcc.target.o ran%/driver.gcc.target.o ran%/init.gcc.target.o ran%/check.gcc.target.o ran%/hash.gcc.target.o + $(TARGET_CC) $(TARGET_CFLAGS) $+ -o $@ + +ran%/example.gcc.host: ran%/func.gcc.host.o ran%/driver.gcc.host.o ran%/init.gcc.host.o ran%/check.gcc.host.o ran%/hash.gcc.host.o + $(CC) $(CFLAGS) $+ -o $@ + +ran%/driver.c ran%/func.c ran%/init.c ran%/check.c ran%/hash.c ran%/init.h: + mkdir -p ran$* $(YARPGEN) --seed=$* --out-dir=ran$*/ --std=c99 -ran%/example.k1c.cmp : ran%/example.gcc.k1c.out ran%/example.ccomp.k1c.out +ran%/example.target.cmp : ran%/example.gcc.target.out ran%/example.ccomp.target.out cmp $+ > $@ -.PHONY: all clean +ran%/example.host_target.cmp : ran%/example.gcc.host.out ran%/example.ccomp.target.out + cmp $+ > $@ + +yarpgen: + curl -L -o yarpgen_v1.1.tar.gz https://github.com/intel/yarpgen/archive/v1.1.tar.gz + tar xfz yarpgen_v1.1.tar.gz + $(MAKE) CXX=g++ -C yarpgen-1.1 + cp yarpgen-1.1/yarpgen $@ + +.PHONY: all clean tests_c tests_c clean: -rm -rf ran* diff --git a/test/monniaux/yarpgen/Makefile.old b/test/monniaux/yarpgen/Makefile.old new file mode 100644 index 00000000..9da82deb --- /dev/null +++ b/test/monniaux/yarpgen/Makefile.old @@ -0,0 +1,52 @@ +YARPGEN=yarpgen +MAX=300 +PREFIX=ran%06.f +include ../rules.mk + +K1C_CCOMPFLAGS += -funprototyped -fbitfields +CCOMPFLAGS += -funprototyped -fbitfields + +TARGETS_C=$(shell seq --format $(PREFIX)/func.c 0 $(MAX)) \ + $(shell seq --format $(PREFIX)/driver.c 0 $(MAX)) \ + $(shell seq --format $(PREFIX)/init.h 0 $(MAX)) +TARGETS_CCOMP_K1C_S=$(shell seq --format $(PREFIX)/func.ccomp.k1c.s 0 $(MAX)) \ + $(shell seq --format $(PREFIX)/driver.ccomp.k1c.s 0 $(MAX)) +TARGETS_GCC_K1C_S=$(shell seq --format $(PREFIX)/func.gcc.k1c.s 0 $(MAX)) \ + $(shell seq --format $(PREFIX)/driver.gcc.k1c.s 0 $(MAX)) +TARGETS_CCOMP_HOST_S=$(shell seq --format $(PREFIX)/func.ccomp.host.s 0 $(MAX)) \ + $(shell seq --format $(PREFIX)/driver.ccomp.host.s 0 $(MAX)) +TARGETS_GCC_HOST_S=$(shell seq --format $(PREFIX)/func.gcc.host.s 0 $(MAX)) \ + $(shell seq --format $(PREFIX)/driver.gcc.host.s 0 $(MAX)) +TARGETS_CCOMP_K1C_OUT=$(shell seq --format $(PREFIX)/example.ccomp.k1c.out 0 $(MAX)) +TARGETS_GCC_K1C_OUT=$(shell seq --format $(PREFIX)/example.gcc.k1c.out 0 $(MAX)) +TARGETS_GCC_HOST_OUT=$(shell seq --format $(PREFIX)/example.gcc.host.out 0 $(MAX)) +TARGETS_CCOMP_HOST_OUT=$(shell seq --format $(PREFIX)/example.ccomp.host.out 0 $(MAX)) +TARGETS_CMP=$(shell seq --format $(PREFIX)/example.k1c.cmp 0 $(MAX)) + +all: $(TARGETS_CCOMP_K1C_OUT) $(TARGETS_GCC_K1C_OUT) $(TARGETS_GCC_HOST_OUT) $(TARGETS_CCOMP_HOST_OUT) $(TARGETS_CCOMP_K1C_S) $(TARGETS_GCC_K1C_S) $(TARGETS_GCC_HOST_S) $(TARGETS_CCOMP_HOST_S) $(TARGETS_CMP) $(TARGETS_C) + +ran%/func.ccomp.k1c.s ran%/func.gcc.k1c.s ran%/func.ccomp.host.s ran%/func.gcc.host.s : ran%/init.h + +ran%/example.ccomp.k1c: ran%/func.ccomp.k1c.o ran%/driver.ccomp.k1c.o + $(K1C_CCOMP) $(K1C_CCOMPFLAGS) $+ -o $@ + +ran%/example.gcc.k1c: ran%/func.gcc.k1c.o ran%/driver.gcc.k1c.o + $(K1C_CC) $(K1C_CFLAGS) $+ -o $@ + +ran%/example.gcc.host: ran%/func.gcc.host.o ran%/driver.gcc.host.o + $(CC) $(CFLAGS) $+ -o $@ + +ran%/example.ccomp.host: ran%/func.ccomp.host.o ran%/driver.ccomp.host.o + $(CCOMP) $(CCOMPFLAGS) $+ -o $@ + +ran%/driver.c ran%/func.c ran%/init.h: + -mkdir ran$* + $(YARPGEN) --seed=$* --out-dir=ran$*/ --std=c99 + +ran%/example.k1c.cmp : ran%/example.gcc.k1c.out ran%/example.ccomp.k1c.out + cmp $+ > $@ + +.PHONY: all clean + +clean: + -rm -rf ran* -- cgit From 5e84a1aea751e8c4c46a2899a2901bb59a1f049b Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sat, 28 Mar 2020 16:29:34 +0100 Subject: run yarpgen correctly on arm --- .gitlab-ci.yml | 2 +- test/monniaux/yarpgen/Makefile | 14 +++++++++++--- 2 files changed, 12 insertions(+), 4 deletions(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 119fa44e..bf83a026 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -81,7 +81,7 @@ build_arm: - ./config_arm.sh - make -j "$NJOBS" - make -C test DIRS="c compression spass regression" CCOMPOPTS=-static SIMU='qemu-arm' EXECUTE='qemu-arm' all test - - make -C test/monniaux/yarpgen -k TARGET_CC='arm-linux-gnueabihf-gcc' EXECUTE='qemu-arm' CCOMPOPTS='-static' TARGET_CFLAGS='-static' + - make -C test/monniaux/yarpgen -k TARGET_CC='arm-linux-gnueabihf-gcc' EXECUTE='qemu-arm' CCOMPOPTS='-static' TARGET_CFLAGS='-static' BITS=32 rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always diff --git a/test/monniaux/yarpgen/Makefile b/test/monniaux/yarpgen/Makefile index dbd6ae75..339d6808 100644 --- a/test/monniaux/yarpgen/Makefile +++ b/test/monniaux/yarpgen/Makefile @@ -10,6 +10,11 @@ YARPGEN=./yarpgen GENERATOR=yarpgen endif +ifdef BITS +YARPGEN+=-m $(BITS) +CFLAGS+=-m$(BITS) +endif + MAX=129 PREFIX=ran%06.f @@ -49,13 +54,16 @@ tests_s: $(TESTS_CCOMP_TARGET_S) $(CC) $(CFLAGS) -S -o $@ $< %.target.o : %.target.s - $(TARGET_CC) -c -o $@ $< + $(TARGET_CC) $(CFLAGS) -c -o $@ $< + +%.host.o : %.host.s + $(CC) $(CFLAGS) -c -o $@ $< %.target.out : %.target - $(EXECUTE) $< > $@ + $(EXECUTE) $< | tee $@ %.host.out : %.host - ./$< > $@ + ./$< | tee $@ ran%/func.ccomp.target.s ran%/func.gcc.target.s ran%/func.ccomp.host.s ran%/func.gcc.host.s ran%/init.gcc.host.s : ran%/init.h -- cgit From ad5c72c2bf72e11eeb58e95842879c272077e669 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sat, 28 Mar 2020 16:30:34 +0100 Subject: run also on IA32 remove -k --- .gitlab-ci.yml | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index bf83a026..7de12153 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -35,7 +35,7 @@ build_ia32: - ./config_ia32.sh - make -j "$NJOBS" - make -C test all test -# - make -C test/monniaux/yarpgen disabled due to -m32/-m64 issues + - make -C test/monniaux/yarpgen BITS=32 rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always @@ -58,7 +58,7 @@ build_aarch64: - ./config_aarch64.sh - make -j "$NJOBS" - make -C test CCOMPOPTS='-static' SIMU='qemu-aarch64' EXECUTE='qemu-aarch64' all test - - make -C test/monniaux/yarpgen -k TARGET_CC='aarch64-linux-gnu-gcc' EXECUTE='qemu-aarch64' CCOMPOPTS='-static' TARGET_CFLAGS='-static' + - make -C test/monniaux/yarpgen TARGET_CC='aarch64-linux-gnu-gcc' EXECUTE='qemu-aarch64' CCOMPOPTS='-static' TARGET_CFLAGS='-static' rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always @@ -81,7 +81,7 @@ build_arm: - ./config_arm.sh - make -j "$NJOBS" - make -C test DIRS="c compression spass regression" CCOMPOPTS=-static SIMU='qemu-arm' EXECUTE='qemu-arm' all test - - make -C test/monniaux/yarpgen -k TARGET_CC='arm-linux-gnueabihf-gcc' EXECUTE='qemu-arm' CCOMPOPTS='-static' TARGET_CFLAGS='-static' BITS=32 + - make -C test/monniaux/yarpgen TARGET_CC='arm-linux-gnueabihf-gcc' EXECUTE='qemu-arm' CCOMPOPTS='-static' TARGET_CFLAGS='-static' BITS=32 rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always @@ -146,7 +146,7 @@ build_rv64: - ./config_rv64.sh - make -j "$NJOBS" - make -C test CCOMPOPTS=-static SIMU='qemu-riscv64' EXECUTE='qemu-riscv64' all test - - make -C test/monniaux/yarpgen -k TARGET_CC='riscv64-linux-gnu-gcc' EXECUTE='qemu-riscv64' CCOMPOPTS='-static' TARGET_CFLAGS='-static' + - make -C test/monniaux/yarpgen TARGET_CC='riscv64-linux-gnu-gcc' EXECUTE='qemu-riscv64' CCOMPOPTS='-static' TARGET_CFLAGS='-static' rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always -- cgit From d3fe2c1d8a05b5124395cca3de0cf91470424e55 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sat, 28 Mar 2020 16:43:10 +0100 Subject: enlarge stack size --- .gitlab-ci.yml | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 7de12153..2294d090 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -12,7 +12,7 @@ build_x86_64: - ./config_x86_64.sh - make -j "$NJOBS" - make -C test all test - - make -C test/monniaux/yarpgen + - ulimit -s65536 && make -C test/monniaux/yarpgen rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always @@ -35,7 +35,7 @@ build_ia32: - ./config_ia32.sh - make -j "$NJOBS" - make -C test all test - - make -C test/monniaux/yarpgen BITS=32 + - ulimit -s65536 && make -C test/monniaux/yarpgen BITS=32 rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always @@ -58,7 +58,7 @@ build_aarch64: - ./config_aarch64.sh - make -j "$NJOBS" - make -C test CCOMPOPTS='-static' SIMU='qemu-aarch64' EXECUTE='qemu-aarch64' all test - - make -C test/monniaux/yarpgen TARGET_CC='aarch64-linux-gnu-gcc' EXECUTE='qemu-aarch64' CCOMPOPTS='-static' TARGET_CFLAGS='-static' + - ulimit -s65536 && make -C test/monniaux/yarpgen TARGET_CC='aarch64-linux-gnu-gcc' EXECUTE='qemu-aarch64' CCOMPOPTS='-static' TARGET_CFLAGS='-static' rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always @@ -81,7 +81,7 @@ build_arm: - ./config_arm.sh - make -j "$NJOBS" - make -C test DIRS="c compression spass regression" CCOMPOPTS=-static SIMU='qemu-arm' EXECUTE='qemu-arm' all test - - make -C test/monniaux/yarpgen TARGET_CC='arm-linux-gnueabihf-gcc' EXECUTE='qemu-arm' CCOMPOPTS='-static' TARGET_CFLAGS='-static' BITS=32 + - ulimit -s65536 && make -C test/monniaux/yarpgen TARGET_CC='arm-linux-gnueabihf-gcc' EXECUTE='qemu-arm' CCOMPOPTS='-static' TARGET_CFLAGS='-static' BITS=32 rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always @@ -146,7 +146,7 @@ build_rv64: - ./config_rv64.sh - make -j "$NJOBS" - make -C test CCOMPOPTS=-static SIMU='qemu-riscv64' EXECUTE='qemu-riscv64' all test - - make -C test/monniaux/yarpgen TARGET_CC='riscv64-linux-gnu-gcc' EXECUTE='qemu-riscv64' CCOMPOPTS='-static' TARGET_CFLAGS='-static' + - ulimit -s65536 && make -C test/monniaux/yarpgen TARGET_CC='riscv64-linux-gnu-gcc' EXECUTE='qemu-riscv64' CCOMPOPTS='-static' TARGET_CFLAGS='-static' rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always -- cgit From 55b7716fa59b7b3c1b16f64b5b9debbd1736a974 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sat, 28 Mar 2020 17:01:49 +0100 Subject: fixup for arm --- test/monniaux/yarpgen/Makefile | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/test/monniaux/yarpgen/Makefile b/test/monniaux/yarpgen/Makefile index 339d6808..a9f62eb7 100644 --- a/test/monniaux/yarpgen/Makefile +++ b/test/monniaux/yarpgen/Makefile @@ -54,7 +54,7 @@ tests_s: $(TESTS_CCOMP_TARGET_S) $(CC) $(CFLAGS) -S -o $@ $< %.target.o : %.target.s - $(TARGET_CC) $(CFLAGS) -c -o $@ $< + $(TARGET_CC) -c -o $@ $< %.host.o : %.host.s $(CC) $(CFLAGS) -c -o $@ $< -- cgit From dc9a19346eb6bb69b30381beb703a4b8c2a7c59b Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sat, 28 Mar 2020 17:33:53 +0100 Subject: yet another problem with 32-bit arm --- .gitlab-ci.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 2294d090..9d2103b1 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -73,7 +73,7 @@ build_arm: image: "coqorg/coq" before_script: - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update - - sudo apt-get -y install gcc-arm-linux-gnueabihf qemu-user + - sudo apt-get -y install gcc-multilib gcc-multilib-arm-linux-gnueabihf qemu-user - opam switch 4.07.1+flambda - eval `opam config env` - opam install -y menhir -- cgit From 6008cab1fad50f61cf76075664e6c8bada818509 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sat, 28 Mar 2020 17:38:38 +0100 Subject: Debian is not like Ubuntu on multilib! --- .gitlab-ci.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 9d2103b1..d4b1385e 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -73,7 +73,7 @@ build_arm: image: "coqorg/coq" before_script: - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update - - sudo apt-get -y install gcc-multilib gcc-multilib-arm-linux-gnueabihf qemu-user + - sudo apt-get -y install gcc-multilib gcc-arm-linux-gnueabihf qemu-user - opam switch 4.07.1+flambda - eval `opam config env` - opam install -y menhir -- cgit From 2bf8878a4b424e0481e9931c9047f6450e7ba0fd Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sat, 28 Mar 2020 17:49:30 +0100 Subject: remove tests wrt host --- .gitlab-ci.yml | 2 +- test/monniaux/yarpgen/Makefile | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index d4b1385e..2294d090 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -73,7 +73,7 @@ build_arm: image: "coqorg/coq" before_script: - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update - - sudo apt-get -y install gcc-multilib gcc-arm-linux-gnueabihf qemu-user + - sudo apt-get -y install gcc-arm-linux-gnueabihf qemu-user - opam switch 4.07.1+flambda - eval `opam config env` - opam install -y menhir diff --git a/test/monniaux/yarpgen/Makefile b/test/monniaux/yarpgen/Makefile index a9f62eb7..65759f1e 100644 --- a/test/monniaux/yarpgen/Makefile +++ b/test/monniaux/yarpgen/Makefile @@ -36,7 +36,7 @@ TESTS_GCC_HOST_S=$(TEST_C:.c=.gcc.host.s) TESTS_CCOMP_TARGET_OUT=$(shell seq --format $(PREFIX)/example.ccomp.target.out 1 $(MAX)) TESTS_GCC_TARGET_OUT=$(shell seq --format $(PREFIX)/example.gcc.target.out 1 $(MAX)) TESTS_GCC_HOST_OUT=$(shell seq --format $(PREFIX)/example.gcc.host.out 1 $(MAX)) -TESTS_CMP=$(shell seq --format $(PREFIX)/example.host_target.cmp 1 $(MAX)) $(shell seq --format $(PREFIX)/example.target.cmp 1 $(MAX)) +TESTS_CMP=$(shell seq --format $(PREFIX)/example.target.cmp 1 $(MAX)) # $(shell seq --format $(PREFIX)/example.host_target.cmp 1 $(MAX)) all: $(TESTS_CCOMP_TARGET_OUT) $(TESTS_GCC_TARGET_OUT) $(TESTS_GCC_HOST_OUT) $(TESTS_CCOMP_TARGET_S) $(TESTS_GCC_TARGET_S) $(TESTS_GCC_HOST_S) $(TESTS_CMP) $(TESTS_C) -- cgit From bd432c88df63a21dfd3cfbf038f25781fa8b0ee5 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sat, 28 Mar 2020 18:06:16 +0100 Subject: use gcc -m32 on ia32 --- .gitlab-ci.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 2294d090..fd5dffe5 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -35,7 +35,7 @@ build_ia32: - ./config_ia32.sh - make -j "$NJOBS" - make -C test all test - - ulimit -s65536 && make -C test/monniaux/yarpgen BITS=32 + - ulimit -s65536 && make -C test/monniaux/yarpgen BITS=32 TARGET_CC='gcc -m32' rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always -- cgit From 80078dcafb07578878a3d0ed2a52c08b4ee12e4a Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sat, 28 Mar 2020 18:25:25 +0100 Subject: remove host .s generation --- test/monniaux/yarpgen/Makefile | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/test/monniaux/yarpgen/Makefile b/test/monniaux/yarpgen/Makefile index 65759f1e..28bd5ae0 100644 --- a/test/monniaux/yarpgen/Makefile +++ b/test/monniaux/yarpgen/Makefile @@ -38,7 +38,7 @@ TESTS_GCC_TARGET_OUT=$(shell seq --format $(PREFIX)/example.gcc.target.out 1 $(M TESTS_GCC_HOST_OUT=$(shell seq --format $(PREFIX)/example.gcc.host.out 1 $(MAX)) TESTS_CMP=$(shell seq --format $(PREFIX)/example.target.cmp 1 $(MAX)) # $(shell seq --format $(PREFIX)/example.host_target.cmp 1 $(MAX)) -all: $(TESTS_CCOMP_TARGET_OUT) $(TESTS_GCC_TARGET_OUT) $(TESTS_GCC_HOST_OUT) $(TESTS_CCOMP_TARGET_S) $(TESTS_GCC_TARGET_S) $(TESTS_GCC_HOST_S) $(TESTS_CMP) $(TESTS_C) +all: $(TESTS_CCOMP_TARGET_OUT) $(TESTS_GCC_TARGET_OUT) $(TESTS_CCOMP_TARGET_S) $(TESTS_GCC_TARGET_S) $(TESTS_CMP) $(TESTS_C) tests_c: $(TESTS_C) -- cgit From 39d53bf1d57c6ce9c78cbc42521ad96a783e6896 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sat, 28 Mar 2020 19:09:55 +0100 Subject: More Yarpgen --- .gitlab-ci.yml | 10 +++++----- test/monniaux/yarpgen/Makefile | 25 +++++++++++++++++++++++++ 2 files changed, 30 insertions(+), 5 deletions(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 5f696257..fd5dffe5 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -12,7 +12,7 @@ build_x86_64: - ./config_x86_64.sh - make -j "$NJOBS" - make -C test all test - - make -C test/monniaux/yarpgen + - ulimit -s65536 && make -C test/monniaux/yarpgen rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always @@ -35,7 +35,7 @@ build_ia32: - ./config_ia32.sh - make -j "$NJOBS" - make -C test all test -# - make -C test/monniaux/yarpgen disabled due to -m32/-m64 issues + - ulimit -s65536 && make -C test/monniaux/yarpgen BITS=32 TARGET_CC='gcc -m32' rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always @@ -58,7 +58,7 @@ build_aarch64: - ./config_aarch64.sh - make -j "$NJOBS" - make -C test CCOMPOPTS='-static' SIMU='qemu-aarch64' EXECUTE='qemu-aarch64' all test - - make -C test/monniaux/yarpgen -k TARGET_CC='aarch64-linux-gnu-gcc' EXECUTE='qemu-aarch64' CCOMPOPTS='-static' TARGET_CFLAGS='-static' + - ulimit -s65536 && make -C test/monniaux/yarpgen TARGET_CC='aarch64-linux-gnu-gcc' EXECUTE='qemu-aarch64' CCOMPOPTS='-static' TARGET_CFLAGS='-static' rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always @@ -81,7 +81,7 @@ build_arm: - ./config_arm.sh - make -j "$NJOBS" - make -C test DIRS="c compression spass regression" CCOMPOPTS=-static SIMU='qemu-arm' EXECUTE='qemu-arm' all test -# - make -C test/monniaux/yarpgen -k TARGET_CC='arm-linux-gnueabihf-gcc' EXECUTE='qemu-arm' CCOMPOPTS='-static' TARGET_CFLAGS='-static' disabled; mysterious differences between gcc/clang and compcert + - ulimit -s65536 && make -C test/monniaux/yarpgen TARGET_CC='arm-linux-gnueabihf-gcc' EXECUTE='qemu-arm' CCOMPOPTS='-static' TARGET_CFLAGS='-static' BITS=32 rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always @@ -146,7 +146,7 @@ build_rv64: - ./config_rv64.sh - make -j "$NJOBS" - make -C test CCOMPOPTS=-static SIMU='qemu-riscv64' EXECUTE='qemu-riscv64' all test - - make -C test/monniaux/yarpgen -k TARGET_CC='riscv64-linux-gnu-gcc' EXECUTE='qemu-riscv64' CCOMPOPTS='-static' TARGET_CFLAGS='-static' + - ulimit -s65536 && make -C test/monniaux/yarpgen TARGET_CC='riscv64-linux-gnu-gcc' EXECUTE='qemu-riscv64' CCOMPOPTS='-static' TARGET_CFLAGS='-static' rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always diff --git a/test/monniaux/yarpgen/Makefile b/test/monniaux/yarpgen/Makefile index dbd6ae75..30080855 100644 --- a/test/monniaux/yarpgen/Makefile +++ b/test/monniaux/yarpgen/Makefile @@ -10,6 +10,14 @@ YARPGEN=./yarpgen GENERATOR=yarpgen endif +<<<<<<< HEAD +======= +ifdef BITS +YARPGEN+=-m $(BITS) +CFLAGS+=-m$(BITS) +endif + +>>>>>>> origin/mppa-ci MAX=129 PREFIX=ran%06.f @@ -31,9 +39,15 @@ TESTS_GCC_HOST_S=$(TEST_C:.c=.gcc.host.s) TESTS_CCOMP_TARGET_OUT=$(shell seq --format $(PREFIX)/example.ccomp.target.out 1 $(MAX)) TESTS_GCC_TARGET_OUT=$(shell seq --format $(PREFIX)/example.gcc.target.out 1 $(MAX)) TESTS_GCC_HOST_OUT=$(shell seq --format $(PREFIX)/example.gcc.host.out 1 $(MAX)) +<<<<<<< HEAD TESTS_CMP=$(shell seq --format $(PREFIX)/example.host_target.cmp 1 $(MAX)) $(shell seq --format $(PREFIX)/example.target.cmp 1 $(MAX)) all: $(TESTS_CCOMP_TARGET_OUT) $(TESTS_GCC_TARGET_OUT) $(TESTS_GCC_HOST_OUT) $(TESTS_CCOMP_TARGET_S) $(TESTS_GCC_TARGET_S) $(TESTS_GCC_HOST_S) $(TESTS_CMP) $(TESTS_C) +======= +TESTS_CMP=$(shell seq --format $(PREFIX)/example.target.cmp 1 $(MAX)) # $(shell seq --format $(PREFIX)/example.host_target.cmp 1 $(MAX)) + +all: $(TESTS_CCOMP_TARGET_OUT) $(TESTS_GCC_TARGET_OUT) $(TESTS_CCOMP_TARGET_S) $(TESTS_GCC_TARGET_S) $(TESTS_CMP) $(TESTS_C) +>>>>>>> origin/mppa-ci tests_c: $(TESTS_C) @@ -51,11 +65,22 @@ tests_s: $(TESTS_CCOMP_TARGET_S) %.target.o : %.target.s $(TARGET_CC) -c -o $@ $< +<<<<<<< HEAD %.target.out : %.target $(EXECUTE) $< > $@ %.host.out : %.host ./$< > $@ +======= +%.host.o : %.host.s + $(CC) $(CFLAGS) -c -o $@ $< + +%.target.out : %.target + $(EXECUTE) $< | tee $@ + +%.host.out : %.host + ./$< | tee $@ +>>>>>>> origin/mppa-ci ran%/func.ccomp.target.s ran%/func.gcc.target.s ran%/func.ccomp.host.s ran%/func.gcc.host.s ran%/init.gcc.host.s : ran%/init.h -- cgit From e0fcc8d9023d962ec3921d8c6f09c8baa69bca32 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sat, 28 Mar 2020 19:33:37 +0100 Subject: fix Makefile --- test/monniaux/yarpgen/Makefile | 17 ----------------- 1 file changed, 17 deletions(-) diff --git a/test/monniaux/yarpgen/Makefile b/test/monniaux/yarpgen/Makefile index 30080855..28bd5ae0 100644 --- a/test/monniaux/yarpgen/Makefile +++ b/test/monniaux/yarpgen/Makefile @@ -10,14 +10,11 @@ YARPGEN=./yarpgen GENERATOR=yarpgen endif -<<<<<<< HEAD -======= ifdef BITS YARPGEN+=-m $(BITS) CFLAGS+=-m$(BITS) endif ->>>>>>> origin/mppa-ci MAX=129 PREFIX=ran%06.f @@ -39,15 +36,9 @@ TESTS_GCC_HOST_S=$(TEST_C:.c=.gcc.host.s) TESTS_CCOMP_TARGET_OUT=$(shell seq --format $(PREFIX)/example.ccomp.target.out 1 $(MAX)) TESTS_GCC_TARGET_OUT=$(shell seq --format $(PREFIX)/example.gcc.target.out 1 $(MAX)) TESTS_GCC_HOST_OUT=$(shell seq --format $(PREFIX)/example.gcc.host.out 1 $(MAX)) -<<<<<<< HEAD -TESTS_CMP=$(shell seq --format $(PREFIX)/example.host_target.cmp 1 $(MAX)) $(shell seq --format $(PREFIX)/example.target.cmp 1 $(MAX)) - -all: $(TESTS_CCOMP_TARGET_OUT) $(TESTS_GCC_TARGET_OUT) $(TESTS_GCC_HOST_OUT) $(TESTS_CCOMP_TARGET_S) $(TESTS_GCC_TARGET_S) $(TESTS_GCC_HOST_S) $(TESTS_CMP) $(TESTS_C) -======= TESTS_CMP=$(shell seq --format $(PREFIX)/example.target.cmp 1 $(MAX)) # $(shell seq --format $(PREFIX)/example.host_target.cmp 1 $(MAX)) all: $(TESTS_CCOMP_TARGET_OUT) $(TESTS_GCC_TARGET_OUT) $(TESTS_CCOMP_TARGET_S) $(TESTS_GCC_TARGET_S) $(TESTS_CMP) $(TESTS_C) ->>>>>>> origin/mppa-ci tests_c: $(TESTS_C) @@ -65,13 +56,6 @@ tests_s: $(TESTS_CCOMP_TARGET_S) %.target.o : %.target.s $(TARGET_CC) -c -o $@ $< -<<<<<<< HEAD -%.target.out : %.target - $(EXECUTE) $< > $@ - -%.host.out : %.host - ./$< > $@ -======= %.host.o : %.host.s $(CC) $(CFLAGS) -c -o $@ $< @@ -80,7 +64,6 @@ tests_s: $(TESTS_CCOMP_TARGET_S) %.host.out : %.host ./$< | tee $@ ->>>>>>> origin/mppa-ci ran%/func.ccomp.target.s ran%/func.gcc.target.s ran%/func.ccomp.host.s ran%/func.gcc.host.s ran%/init.gcc.host.s : ran%/init.h -- cgit From 14f89bf9c397a4268d2b47418de234992b008d6c Mon Sep 17 00:00:00 2001 From: Bernhard Schommer Date: Sun, 29 Mar 2020 11:35:58 +0200 Subject: Explicit error messages for ill-formed section attributes (#232) Introduce an error message for section attributes with non string arguments,and another for multiple, ambiguous section attributes. This is more consistent with the handling of other attributes, like packed, than the old behavior of silently ignoring them. --- cfrontend/C2C.ml | 4 ++-- common/Sections.ml | 29 +++++++++++++++++++++-------- common/Sections.mli | 4 ++-- 3 files changed, 25 insertions(+), 12 deletions(-) diff --git a/cfrontend/C2C.ml b/cfrontend/C2C.ml index 293e79f0..7f796fe3 100644 --- a/cfrontend/C2C.ml +++ b/cfrontend/C2C.ml @@ -1234,7 +1234,7 @@ let convertFundef loc env fd = { a_storage = fd.fd_storage; a_alignment = None; a_size = None; - a_sections = Sections.for_function env id' fd.fd_attrib; + a_sections = Sections.for_function env loc id' fd.fd_attrib; a_access = Sections.Access_default; a_inline = inline; a_loc = loc }; @@ -1311,7 +1311,7 @@ let convertGlobvar loc env (sto, id, ty, optinit) = | Some i -> convertInitializer env ty i in let (section, access) = - Sections.for_variable env id' ty (optinit <> None) in + Sections.for_variable env loc id' ty (optinit <> None) in if Z.gt sz (Z.of_uint64 0xFFFF_FFFFL) then error "'%s' is too big (%s bytes)" id.name (Z.to_string sz); diff --git a/common/Sections.ml b/common/Sections.ml index 30be9e69..839128a5 100644 --- a/common/Sections.ml +++ b/common/Sections.ml @@ -160,9 +160,22 @@ let gcc_section name readonly exec = sec_writable = not readonly; sec_executable = exec; sec_access = Access_default } +(* Check and extract whether a section was given as attribute *) + +let get_attr_section loc attr = + match Cutil.find_custom_attributes ["section"; "__section__"] attr with + | [] -> None + | [[C.AString name]] -> Some name + | [[_]] -> + Diagnostics.error loc "'section' attribute requires a string"; + None + | _ -> + Diagnostics.error loc "ambiguous 'section' attribute"; + None + (* Determine section for a variable definition *) -let for_variable env id ty init = +let for_variable env loc id ty init = let attr = Cutil.attributes_of_type env ty in let readonly = List.mem C.AConst attr && not(List.mem C.AVolatile attr) in let si = @@ -170,11 +183,11 @@ let for_variable env id ty init = (* 1- Section explicitly associated with #use_section *) Hashtbl.find use_section_table id with Not_found -> - match Cutil.find_custom_attributes ["section"; "__section__"] attr with - | [[C.AString name]] -> + match get_attr_section loc attr with + | Some name -> (* 2- Section given as an attribute, gcc-style *) gcc_section name readonly false - | _ -> + | None -> (* 3- Default section appropriate for size and const-ness *) let size = match Cutil.sizeof env ty with Some sz -> sz | None -> max_int in @@ -190,17 +203,17 @@ let for_variable env id ty init = (* Determine sections for a function definition *) -let for_function env id attr = +let for_function env loc id attr = let si_code = try (* 1- Section explicitly associated with #use_section *) Hashtbl.find use_section_table id with Not_found -> - match Cutil.find_custom_attributes ["section"; "__section__"] attr with - | [[C.AString name]] -> + match get_attr_section loc attr with + | Some name -> (* 2- Section given as an attribute, gcc-style *) gcc_section name true true - | _ -> + | None -> (* 3- Default section *) try Hashtbl.find current_section_table "CODE" diff --git a/common/Sections.mli b/common/Sections.mli index bc97814d..d9fd9239 100644 --- a/common/Sections.mli +++ b/common/Sections.mli @@ -46,7 +46,7 @@ val define_section: -> ?writable:bool -> ?executable:bool -> ?access:access_mode -> unit -> unit val use_section_for: AST.ident -> string -> bool -val for_variable: Env.t -> AST.ident -> C.typ -> bool -> +val for_variable: Env.t -> C.location -> AST.ident -> C.typ -> bool -> section_name * access_mode -val for_function: Env.t -> AST.ident -> C.attributes -> section_name list +val for_function: Env.t -> C.location -> AST.ident -> C.attributes -> section_name list val for_stringlit: unit -> section_name -- cgit From ac16e9fd0fde714464fbef4719e1a341208d5a93 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sun, 29 Mar 2020 12:29:03 +0200 Subject: fix mismatch between hardware FP and software FP on ARM --- .gitlab-ci.yml | 28 ++++++++++++++++++++++++++-- config_arm.sh | 2 +- config_armhf.sh | 1 + 3 files changed, 28 insertions(+), 3 deletions(-) create mode 100755 config_armhf.sh diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index fd5dffe5..eaa313f1 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -73,14 +73,38 @@ build_arm: image: "coqorg/coq" before_script: - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update - - sudo apt-get -y install gcc-arm-linux-gnueabihf qemu-user + - sudo apt-get -y install gcc-arm-linux-gnueabi qemu-user - opam switch 4.07.1+flambda - eval `opam config env` - opam install -y menhir script: - ./config_arm.sh - make -j "$NJOBS" - - make -C test DIRS="c compression spass regression" CCOMPOPTS=-static SIMU='qemu-arm' EXECUTE='qemu-arm' all test + - make -C test CCOMPOPTS=-static SIMU='qemu-arm' EXECUTE='qemu-arm' all test + - ulimit -s65536 && make -C test/monniaux/yarpgen TARGET_CC='arm-linux-gnueabihf-gcc' EXECUTE='qemu-arm' CCOMPOPTS='-static' TARGET_CFLAGS='-static' BITS=32 + rules: + - if: '$CI_COMMIT_BRANCH == "mppa-work"' + when: always + - if: '$CI_COMMIT_BRANCH == "mppa-k1c"' + when: always + - if: '$CI_COMMIT_BRANCH == "master"' + when: always + - when: manual + + +build_armhf: + stage: build + image: "coqorg/coq" + before_script: + - sudo apt-get -o Acquire::Check-Valid-Until=false -o Acquire::Check-Date=false update + - sudo apt-get -y install gcc-arm-linux-gnueabihf qemu-user + - opam switch 4.07.1+flambda + - eval `opam config env` + - opam install -y menhir + script: + - ./config_armhf.sh + - make -j "$NJOBS" + - make -C test CCOMPOPTS=-static SIMU='qemu-arm' EXECUTE='qemu-arm' all test - ulimit -s65536 && make -C test/monniaux/yarpgen TARGET_CC='arm-linux-gnueabihf-gcc' EXECUTE='qemu-arm' CCOMPOPTS='-static' TARGET_CFLAGS='-static' BITS=32 rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' diff --git a/config_arm.sh b/config_arm.sh index eed55fab..1861e029 100755 --- a/config_arm.sh +++ b/config_arm.sh @@ -1 +1 @@ -exec ./config_simple.sh arm-linux --toolprefix arm-linux-gnueabihf- "$@" +exec ./config_simple.sh arm-linux --toolprefix arm-linux-gnueabi- "$@" diff --git a/config_armhf.sh b/config_armhf.sh new file mode 100755 index 00000000..8a1302bd --- /dev/null +++ b/config_armhf.sh @@ -0,0 +1 @@ +exec ./config_simple.sh arm-eabihf --toolprefix arm-linux-gnueabihf- "$@" -- cgit From aa4dcf296521ebe5be4aee5ee29aa678b8325c46 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Sun, 29 Mar 2020 12:46:56 +0200 Subject: fix typo in hf --- .gitlab-ci.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index eaa313f1..79a32b25 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -81,7 +81,7 @@ build_arm: - ./config_arm.sh - make -j "$NJOBS" - make -C test CCOMPOPTS=-static SIMU='qemu-arm' EXECUTE='qemu-arm' all test - - ulimit -s65536 && make -C test/monniaux/yarpgen TARGET_CC='arm-linux-gnueabihf-gcc' EXECUTE='qemu-arm' CCOMPOPTS='-static' TARGET_CFLAGS='-static' BITS=32 + - ulimit -s65536 && make -C test/monniaux/yarpgen TARGET_CC='arm-linux-gnueabi-gcc' EXECUTE='qemu-arm' CCOMPOPTS='-static' TARGET_CFLAGS='-static' BITS=32 rules: - if: '$CI_COMMIT_BRANCH == "mppa-work"' when: always -- cgit From b52ebb5aaaf9838310d1b7e68b9198c388cab74a Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Mon, 30 Mar 2020 15:24:40 +0200 Subject: Fixing heuristics too sure of themselves --- backend/Duplicateaux.ml | 21 +++++++++++++++------ 1 file changed, 15 insertions(+), 6 deletions(-) diff --git a/backend/Duplicateaux.ml b/backend/Duplicateaux.ml index 54d60d24..28ad4266 100644 --- a/backend/Duplicateaux.ml +++ b/backend/Duplicateaux.ml @@ -159,8 +159,11 @@ let do_call_heuristic code cond ifso ifnot is_loop_header = let predicate n = (function | Icall _ -> true | _ -> false) @@ get_some @@ PTree.get n code - in if (look_ahead code ifso is_loop_header predicate) then Some false - else if (look_ahead code ifnot is_loop_header predicate) then Some true + in let ifso_call = look_ahead code ifso is_loop_header predicate + in let ifnot_call = look_ahead code ifnot is_loop_header predicate + in if ifso_call && ifnot_call then None + else if ifso_call then Some false + else if ifnot_call then Some true else None end @@ -176,8 +179,11 @@ let do_return_heuristic code cond ifso ifnot is_loop_header = let predicate n = (function | Ireturn _ -> true | _ -> false) @@ get_some @@ PTree.get n code - in if (look_ahead code ifso is_loop_header predicate) then Some false - else if (look_ahead code ifnot is_loop_header predicate) then Some true + in let ifso_return = look_ahead code ifso is_loop_header predicate + in let ifnot_return = look_ahead code ifnot is_loop_header predicate + in if ifso_return && ifnot_return then None + else if ifso_return then Some false + else if ifnot_return then Some true else None end @@ -187,8 +193,11 @@ let do_store_heuristic code cond ifso ifnot is_loop_header = let predicate n = (function | Istore _ -> true | _ -> false) @@ get_some @@ PTree.get n code - in if (look_ahead code ifso is_loop_header predicate) then Some false - else if (look_ahead code ifnot is_loop_header predicate) then Some true + in let ifso_store = look_ahead code ifso is_loop_header predicate + in let ifnot_store = look_ahead code ifnot is_loop_header predicate + in if ifso_store && ifnot_store then None + else if ifso_store then Some false + else if ifnot_store then Some true else None end -- cgit From f2de2518509f198c5ce958ec06c18e78e896f814 Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Mon, 30 Mar 2020 12:59:24 +0200 Subject: Add a test for int64 -> float32 conversion This is a special value that causes double rounding with the naive conversion schema int64 -> float64 -> float32. --- test/regression/Results/int64 | 874 ++++++++++++++++++++++++++++++++++++++++-- test/regression/int64.c | 3 +- 2 files changed, 838 insertions(+), 39 deletions(-) diff --git a/test/regression/Results/int64 b/test/regression/Results/int64 index af444cf6..ae8a3cc1 100644 --- a/test/regression/Results/int64 +++ b/test/regression/Results/int64 @@ -334,6 +334,48 @@ dtos f = 0 utof x = 0 stof x = 0 +x = 0 +y = 52ce6b4000000063 +-x = 0 +x + y = 52ce6b4000000063 +x - y = ad3194bfffffff9d +x * y = 0 +x /u y = 0 +x %u y = 0 +x /s y = 0 +x %s y = 0 +x /u y2 = 0 +x %u y2 = 0 +x /s y3 = 0 +x %s y3 = 0 +x /u 3 = 0 +x %u 3 = 0 +x /s 3 = 0 +x %s 3 = 0 +x /u 5 = 0 +x %u 5 = 0 +x /s 5 = 0 +x %s 5 = 0 +x /u 11 = 0 +x %u 11 = 0 +x /s 11 = 0 +x %s 11 = 0 +~x = ffffffffffffffff +x & y = 0 +x | y = 52ce6b4000000063 +x ^ y = 52ce6b4000000063 +x << i = 0 +x >>u i = 0 +x >>s i = 0 +x cmpu y = lt +x cmps y = lt +utod x = 0 +dtou f = 0 +stod x = 0 +dtos f = 0 +utof x = 0 +stof x = 0 + x = 0 y = 14057b7ef767814f -x = 0 @@ -754,6 +796,48 @@ dtos f = 0 utof x = 3f800000 stof x = 3f800000 +x = 1 +y = 52ce6b4000000063 +-x = ffffffffffffffff +x + y = 52ce6b4000000064 +x - y = ad3194bfffffff9e +x * y = 52ce6b4000000063 +x /u y = 0 +x %u y = 1 +x /s y = 0 +x %s y = 1 +x /u y2 = 0 +x %u y2 = 1 +x /s y3 = 0 +x %s y3 = 1 +x /u 3 = 0 +x %u 3 = 1 +x /s 3 = 0 +x %s 3 = 1 +x /u 5 = 0 +x %u 5 = 1 +x /s 5 = 0 +x %s 5 = 1 +x /u 11 = 0 +x %u 11 = 1 +x /s 11 = 0 +x %s 11 = 1 +~x = fffffffffffffffe +x & y = 1 +x | y = 52ce6b4000000063 +x ^ y = 52ce6b4000000062 +x << i = 800000000 +x >>u i = 0 +x >>s i = 0 +x cmpu y = lt +x cmps y = lt +utod x = 3ff0000000000000 +dtou f = 0 +stod x = 3ff0000000000000 +dtos f = 0 +utof x = 3f800000 +stof x = 3f800000 + x = 1 y = 9af678222e728119 -x = ffffffffffffffff @@ -1174,6 +1258,48 @@ dtos f = 0 utof x = 5f800000 stof x = bf800000 +x = ffffffffffffffff +y = 52ce6b4000000063 +-x = 1 +x + y = 52ce6b4000000062 +x - y = ad3194bfffffff9c +x * y = ad3194bfffffff9d +x /u y = 3 +x %u y = 794be3ffffffed6 +x /s y = 0 +x %s y = ffffffffffffffff +x /u y2 = 3176fe836 +x %u y2 = 3683607f +x /s y3 = 0 +x %s y3 = ffffffffffffffff +x /u 3 = 5555555555555555 +x %u 3 = 0 +x /s 3 = 0 +x %s 3 = ffffffffffffffff +x /u 5 = 3333333333333333 +x %u 5 = 0 +x /s 5 = 0 +x %s 5 = ffffffffffffffff +x /u 11 = 1745d1745d1745d1 +x %u 11 = 4 +x /s 11 = 0 +x %s 11 = ffffffffffffffff +~x = 0 +x & y = 52ce6b4000000063 +x | y = ffffffffffffffff +x ^ y = ad3194bfffffff9c +x << i = fffffff800000000 +x >>u i = 1fffffff +x >>s i = ffffffffffffffff +x cmpu y = gt +x cmps y = lt +utod x = 43f0000000000000 +dtou f = 68db8bac710cb +stod x = bff0000000000000 +dtos f = 0 +utof x = 5f800000 +stof x = bf800000 + x = ffffffffffffffff y = 62354cda6226d1f3 -x = 1 @@ -1594,6 +1720,48 @@ dtos f = 346dc utof x = 4f000000 stof x = 4f000000 +x = 7fffffff +y = 52ce6b4000000063 +-x = ffffffff80000001 +x + y = 52ce6b4080000062 +x - y = ad3194c07fffff9c +x * y = ad3194f17fffff9d +x /u y = 0 +x %u y = 7fffffff +x /s y = 0 +x %s y = 7fffffff +x /u y2 = 1 +x %u y2 = 2d3194bf +x /s y3 = 1 +x %s y3 = 2d3194bf +x /u 3 = 2aaaaaaa +x %u 3 = 1 +x /s 3 = 2aaaaaaa +x %s 3 = 1 +x /u 5 = 19999999 +x %u 5 = 2 +x /s 5 = 19999999 +x %s 5 = 2 +x /u 11 = ba2e8ba +x %u 11 = 1 +x /s 11 = ba2e8ba +x %s 11 = 1 +~x = ffffffff80000000 +x & y = 63 +x | y = 52ce6b407fffffff +x ^ y = 52ce6b407fffff9c +x << i = fffffff800000000 +x >>u i = 0 +x >>s i = 0 +x cmpu y = lt +x cmps y = lt +utod x = 41dfffffffc00000 +dtou f = 346dc +stod x = 41dfffffffc00000 +dtos f = 346dc +utof x = 4f000000 +stof x = 4f000000 + x = 7fffffff y = 144093704fadba5d -x = ffffffff80000001 @@ -2014,6 +2182,48 @@ dtos f = 346dc utof x = 4f000000 stof x = 4f000000 +x = 80000000 +y = 52ce6b4000000063 +-x = ffffffff80000000 +x + y = 52ce6b4080000063 +x - y = ad3194c07fffff9d +x * y = 3180000000 +x /u y = 0 +x %u y = 80000000 +x /s y = 0 +x %s y = 80000000 +x /u y2 = 1 +x %u y2 = 2d3194c0 +x /s y3 = 1 +x %s y3 = 2d3194c0 +x /u 3 = 2aaaaaaa +x %u 3 = 2 +x /s 3 = 2aaaaaaa +x %s 3 = 2 +x /u 5 = 19999999 +x %u 5 = 3 +x /s 5 = 19999999 +x %s 5 = 3 +x /u 11 = ba2e8ba +x %u 11 = 2 +x /s 11 = ba2e8ba +x %s 11 = 2 +~x = ffffffff7fffffff +x & y = 0 +x | y = 52ce6b4080000063 +x ^ y = 52ce6b4080000063 +x << i = 0 +x >>u i = 0 +x >>s i = 0 +x cmpu y = lt +x cmps y = lt +utod x = 41e0000000000000 +dtou f = 346dc +stod x = 41e0000000000000 +dtos f = 346dc +utof x = 4f000000 +stof x = 4f000000 + x = 80000000 y = 7b985bc1e7bce4d7 -x = ffffffff80000000 @@ -2434,6 +2644,48 @@ dtos f = 346dc5d638865 utof x = 5f000000 stof x = 5f000000 +x = 7fffffffffffffff +y = 52ce6b4000000063 +-x = 8000000000000001 +x + y = d2ce6b4000000062 +x - y = 2d3194bfffffff9c +x * y = 2d3194bfffffff9d +x /u y = 1 +x %u y = 2d3194bfffffff9c +x /s y = 1 +x %s y = 2d3194bfffffff9c +x /u y2 = 18bb7f41b +x %u y2 = 1b41b03f +x /s y3 = 18bb7f41b +x %s y3 = 1b41b03f +x /u 3 = 2aaaaaaaaaaaaaaa +x %u 3 = 1 +x /s 3 = 2aaaaaaaaaaaaaaa +x %s 3 = 1 +x /u 5 = 1999999999999999 +x %u 5 = 2 +x /s 5 = 1999999999999999 +x %s 5 = 2 +x /u 11 = ba2e8ba2e8ba2e8 +x %u 11 = 7 +x /s 11 = ba2e8ba2e8ba2e8 +x %s 11 = 7 +~x = 8000000000000000 +x & y = 52ce6b4000000063 +x | y = 7fffffffffffffff +x ^ y = 2d3194bfffffff9c +x << i = fffffff800000000 +x >>u i = fffffff +x >>s i = fffffff +x cmpu y = gt +x cmps y = gt +utod x = 43e0000000000000 +dtou f = 346dc5d638865 +stod x = 43e0000000000000 +dtos f = 346dc5d638865 +utof x = 5f000000 +stof x = 5f000000 + x = 7fffffffffffffff y = a220229ec164ffe1 -x = 8000000000000001 @@ -2854,6 +3106,48 @@ dtos f = fffcb923a29c779b utof x = 5f000000 stof x = df000000 +x = 8000000000000000 +y = 52ce6b4000000063 +-x = 8000000000000000 +x + y = d2ce6b4000000063 +x - y = 2d3194bfffffff9d +x * y = 8000000000000000 +x /u y = 1 +x %u y = 2d3194bfffffff9d +x /s y = ffffffffffffffff +x %s y = d2ce6b4000000063 +x /u y2 = 18bb7f41b +x %u y2 = 1b41b040 +x /s y3 = fffffffe74480be5 +x %s y3 = ffffffffe4be4fc0 +x /u 3 = 2aaaaaaaaaaaaaaa +x %u 3 = 2 +x /s 3 = d555555555555556 +x %s 3 = fffffffffffffffe +x /u 5 = 1999999999999999 +x %u 5 = 3 +x /s 5 = e666666666666667 +x %s 5 = fffffffffffffffd +x /u 11 = ba2e8ba2e8ba2e8 +x %u 11 = 8 +x /s 11 = f45d1745d1745d18 +x %s 11 = fffffffffffffff8 +~x = 7fffffffffffffff +x & y = 0 +x | y = d2ce6b4000000063 +x ^ y = d2ce6b4000000063 +x << i = 0 +x >>u i = 10000000 +x >>s i = fffffffff0000000 +x cmpu y = gt +x cmps y = lt +utod x = 43e0000000000000 +dtou f = 346dc5d638865 +stod x = c3e0000000000000 +dtos f = fffcb923a29c779b +utof x = 5f000000 +stof x = df000000 + x = 8000000000000000 y = c73aa0d9a415dfb -x = 8000000000000000 @@ -3274,6 +3568,48 @@ dtos f = 68db8 utof x = 4f800000 stof x = 4f800000 +x = 100000003 +y = 52ce6b4000000063 +-x = fffffffefffffffd +x + y = 52ce6b4100000066 +x - y = ad3194c0ffffffa0 +x * y = f86b422300000129 +x /u y = 0 +x %u y = 100000003 +x /s y = 0 +x %s y = 100000003 +x /u y2 = 3 +x %u y2 = 794be43 +x /s y3 = 3 +x %s y3 = 794be43 +x /u 3 = 55555556 +x %u 3 = 1 +x /s 3 = 55555556 +x %s 3 = 1 +x /u 5 = 33333333 +x %u 5 = 4 +x /s 5 = 33333333 +x %s 5 = 4 +x /u 11 = 1745d174 +x %u 11 = 7 +x /s 11 = 1745d174 +x %s 11 = 7 +~x = fffffffefffffffc +x & y = 3 +x | y = 52ce6b4100000063 +x ^ y = 52ce6b4100000060 +x << i = 1800000000 +x >>u i = 0 +x >>s i = 0 +x cmpu y = lt +x cmps y = lt +utod x = 41f0000000300000 +dtou f = 68db8 +stod x = 41f0000000300000 +dtos f = 68db8 +utof x = 4f800000 +stof x = 4f800000 + x = 100000003 y = e9bcd26890f095a5 -x = fffffffefffffffd @@ -3358,47 +3694,467 @@ dtos f = 14bb101261e18 utof x = 5e4a72c9 stof x = 5e4a72c9 -x = 8362aa9340fe215f -y = f986342416ec8002 --x = 7c9d556cbf01dea1 -x + y = 7ce8deb757eaa161 -x - y = 89dc766f2a11a15d -x * y = e4a2b426803fc2be +x = 52ce6b4000000063 +y = 0 +-x = ad3194bfffffff9d +x + y = 52ce6b4000000063 +x - y = 52ce6b4000000063 +x * y = 0 x /u y = 0 -x %u y = 8362aa9340fe215f -x /s y = 13 -x %s y = fe6ccbe58d70a139 -x /u y2 = 86cb918b -x %u y2 = 910b6dd3 -x /s y3 = 133e437097 -x %s y3 = fffffffffe99a023 -x /u 3 = 2bcb8e3115aa0b1f -x %u 3 = 2 -x /s 3 = d67638dbc054b5cb -x %s 3 = fffffffffffffffe -x /u 5 = 1a46eeea4032d379 -x %u 5 = 2 -x /s 5 = e713bbb70cffa047 -x %s 5 = fffffffffffffffc -x /u 11 = bf1b26a7a45a5f1 -x %u 11 = 4 -x /s 11 = f4abe0f61d2e6020 -x %s 11 = ffffffffffffffff -~x = 7c9d556cbf01dea0 -x & y = 8102200000ec0002 -x | y = fbe6beb756fea15f -x ^ y = 7ae49eb75612a15d -x << i = d8aaa4d03f8857c -x >>u i = 20d8aaa4d03f8857 -x >>s i = e0d8aaa4d03f8857 +x %u y = 0 +x /s y = 0 +x %s y = 0 +x /u y2 = 0 +x %u y2 = 0 +x /s y3 = 0 +x %s y3 = 0 +x /u 3 = 1b9a23c000000021 +x %u 3 = 0 +x /s 3 = 1b9a23c000000021 +x %s 3 = 0 +x /u 5 = 108faf0ccccccce0 +x %u 5 = 3 +x /s 5 = 108faf0ccccccce0 +x %s 5 = 3 +x /u 11 = 7872105d1745d20 +x %u 11 = 3 +x /s 11 = 7872105d1745d20 +x %s 11 = 3 +~x = ad3194bfffffff9c +x & y = 0 +x | y = 52ce6b4000000063 +x ^ y = 52ce6b4000000063 +x << i = 52ce6b4000000063 +x >>u i = 52ce6b4000000063 +x >>s i = 52ce6b4000000063 +x cmpu y = gt +x cmps y = gt +utod x = 43d4b39ad0000000 +dtou f = 21eadf559b3d0 +stod x = 43d4b39ad0000000 +dtos f = 21eadf559b3d0 +utof x = 5ea59cd7 +stof x = 5ea59cd7 + +x = 52ce6b4000000063 +y = 1 +-x = ad3194bfffffff9d +x + y = 52ce6b4000000064 +x - y = 52ce6b4000000062 +x * y = 52ce6b4000000063 +x /u y = 52ce6b4000000063 +x %u y = 0 +x /s y = 52ce6b4000000063 +x %s y = 0 +x /u y2 = 0 +x %u y2 = 0 +x /s y3 = 0 +x %s y3 = 0 +x /u 3 = 1b9a23c000000021 +x %u 3 = 0 +x /s 3 = 1b9a23c000000021 +x %s 3 = 0 +x /u 5 = 108faf0ccccccce0 +x %u 5 = 3 +x /s 5 = 108faf0ccccccce0 +x %s 5 = 3 +x /u 11 = 7872105d1745d20 +x %u 11 = 3 +x /s 11 = 7872105d1745d20 +x %s 11 = 3 +~x = ad3194bfffffff9c +x & y = 1 +x | y = 52ce6b4000000063 +x ^ y = 52ce6b4000000062 +x << i = a59cd680000000c6 +x >>u i = 296735a000000031 +x >>s i = 296735a000000031 +x cmpu y = gt +x cmps y = gt +utod x = 43d4b39ad0000000 +dtou f = 21eadf559b3d0 +stod x = 43d4b39ad0000000 +dtos f = 21eadf559b3d0 +utof x = 5ea59cd7 +stof x = 5ea59cd7 + +x = 52ce6b4000000063 +y = ffffffffffffffff +-x = ad3194bfffffff9d +x + y = 52ce6b4000000062 +x - y = 52ce6b4000000064 +x * y = ad3194bfffffff9d +x /u y = 0 +x %u y = 52ce6b4000000063 +x /s y = ad3194bfffffff9d +x %s y = 0 +x /u y2 = 52ce6b40 +x %u y2 = 52ce6ba3 +x /s y3 = ad3194bfffffff9d +x %s y3 = 0 +x /u 3 = 1b9a23c000000021 +x %u 3 = 0 +x /s 3 = 1b9a23c000000021 +x %s 3 = 0 +x /u 5 = 108faf0ccccccce0 +x %u 5 = 3 +x /s 5 = 108faf0ccccccce0 +x %s 5 = 3 +x /u 11 = 7872105d1745d20 +x %u 11 = 3 +x /s 11 = 7872105d1745d20 +x %s 11 = 3 +~x = ad3194bfffffff9c +x & y = 52ce6b4000000063 +x | y = ffffffffffffffff +x ^ y = ad3194bfffffff9c +x << i = 8000000000000000 +x >>u i = 0 +x >>s i = 0 +x cmpu y = lt +x cmps y = gt +utod x = 43d4b39ad0000000 +dtou f = 21eadf559b3d0 +stod x = 43d4b39ad0000000 +dtos f = 21eadf559b3d0 +utof x = 5ea59cd7 +stof x = 5ea59cd7 + +x = 52ce6b4000000063 +y = 7fffffff +-x = ad3194bfffffff9d +x + y = 52ce6b4080000062 +x - y = 52ce6b3f80000064 +x * y = ad3194f17fffff9d +x /u y = a59cd681 +x %u y = 259cd6e4 +x /s y = a59cd681 +x %s y = 259cd6e4 +x /u y2 = 0 +x %u y2 = 0 +x /s y3 = 0 +x %s y3 = 0 +x /u 3 = 1b9a23c000000021 +x %u 3 = 0 +x /s 3 = 1b9a23c000000021 +x %s 3 = 0 +x /u 5 = 108faf0ccccccce0 +x %u 5 = 3 +x /s 5 = 108faf0ccccccce0 +x %s 5 = 3 +x /u 11 = 7872105d1745d20 +x %u 11 = 3 +x /s 11 = 7872105d1745d20 +x %s 11 = 3 +~x = ad3194bfffffff9c +x & y = 63 +x | y = 52ce6b407fffffff +x ^ y = 52ce6b407fffff9c +x << i = 8000000000000000 +x >>u i = 0 +x >>s i = 0 +x cmpu y = gt +x cmps y = gt +utod x = 43d4b39ad0000000 +dtou f = 21eadf559b3d0 +stod x = 43d4b39ad0000000 +dtos f = 21eadf559b3d0 +utof x = 5ea59cd7 +stof x = 5ea59cd7 + +x = 52ce6b4000000063 +y = 80000000 +-x = ad3194bfffffff9d +x + y = 52ce6b4080000063 +x - y = 52ce6b3f80000063 +x * y = 3180000000 +x /u y = a59cd680 +x %u y = 63 +x /s y = a59cd680 +x %s y = 63 +x /u y2 = 0 +x %u y2 = 0 +x /s y3 = 0 +x %s y3 = 0 +x /u 3 = 1b9a23c000000021 +x %u 3 = 0 +x /s 3 = 1b9a23c000000021 +x %s 3 = 0 +x /u 5 = 108faf0ccccccce0 +x %u 5 = 3 +x /s 5 = 108faf0ccccccce0 +x %s 5 = 3 +x /u 11 = 7872105d1745d20 +x %u 11 = 3 +x /s 11 = 7872105d1745d20 +x %s 11 = 3 +~x = ad3194bfffffff9c +x & y = 0 +x | y = 52ce6b4080000063 +x ^ y = 52ce6b4080000063 +x << i = 52ce6b4000000063 +x >>u i = 52ce6b4000000063 +x >>s i = 52ce6b4000000063 +x cmpu y = gt +x cmps y = gt +utod x = 43d4b39ad0000000 +dtou f = 21eadf559b3d0 +stod x = 43d4b39ad0000000 +dtos f = 21eadf559b3d0 +utof x = 5ea59cd7 +stof x = 5ea59cd7 + +x = 52ce6b4000000063 +y = 7fffffffffffffff +-x = ad3194bfffffff9d +x + y = d2ce6b4000000062 +x - y = d2ce6b4000000064 +x * y = 2d3194bfffffff9d +x /u y = 0 +x %u y = 52ce6b4000000063 +x /s y = 0 +x %s y = 52ce6b4000000063 +x /u y2 = a59cd681 +x %u y2 = 259cd6e4 +x /s y3 = a59cd681 +x %s y3 = 259cd6e4 +x /u 3 = 1b9a23c000000021 +x %u 3 = 0 +x /s 3 = 1b9a23c000000021 +x %s 3 = 0 +x /u 5 = 108faf0ccccccce0 +x %u 5 = 3 +x /s 5 = 108faf0ccccccce0 +x %s 5 = 3 +x /u 11 = 7872105d1745d20 +x %u 11 = 3 +x /s 11 = 7872105d1745d20 +x %s 11 = 3 +~x = ad3194bfffffff9c +x & y = 52ce6b4000000063 +x | y = 7fffffffffffffff +x ^ y = 2d3194bfffffff9c +x << i = 8000000000000000 +x >>u i = 0 +x >>s i = 0 x cmpu y = lt x cmps y = lt -utod x = 43e06c5552681fc4 -dtou f = 35d0c262d14d7 -stod x = c3df27555b2fc078 -dtos f = fffccf536b66040d -utof x = 5f0362ab -stof x = def93aab +utod x = 43d4b39ad0000000 +dtou f = 21eadf559b3d0 +stod x = 43d4b39ad0000000 +dtos f = 21eadf559b3d0 +utof x = 5ea59cd7 +stof x = 5ea59cd7 + +x = 52ce6b4000000063 +y = 8000000000000000 +-x = ad3194bfffffff9d +x + y = d2ce6b4000000063 +x - y = d2ce6b4000000063 +x * y = 8000000000000000 +x /u y = 0 +x %u y = 52ce6b4000000063 +x /s y = 0 +x %s y = 52ce6b4000000063 +x /u y2 = a59cd680 +x %u y2 = 63 +x /s y3 = ffffffff5a632980 +x %s y3 = 63 +x /u 3 = 1b9a23c000000021 +x %u 3 = 0 +x /s 3 = 1b9a23c000000021 +x %s 3 = 0 +x /u 5 = 108faf0ccccccce0 +x %u 5 = 3 +x /s 5 = 108faf0ccccccce0 +x %s 5 = 3 +x /u 11 = 7872105d1745d20 +x %u 11 = 3 +x /s 11 = 7872105d1745d20 +x %s 11 = 3 +~x = ad3194bfffffff9c +x & y = 0 +x | y = d2ce6b4000000063 +x ^ y = d2ce6b4000000063 +x << i = 52ce6b4000000063 +x >>u i = 52ce6b4000000063 +x >>s i = 52ce6b4000000063 +x cmpu y = lt +x cmps y = gt +utod x = 43d4b39ad0000000 +dtou f = 21eadf559b3d0 +stod x = 43d4b39ad0000000 +dtos f = 21eadf559b3d0 +utof x = 5ea59cd7 +stof x = 5ea59cd7 + +x = 52ce6b4000000063 +y = 100000003 +-x = ad3194bfffffff9d +x + y = 52ce6b4100000066 +x - y = 52ce6b3f00000060 +x * y = f86b422300000129 +x /u y = 52ce6b3f +x %u y = 794bea6 +x /s y = 52ce6b3f +x %s y = 794bea6 +x /u y2 = 52ce6b4000000063 +x %u y2 = 0 +x /s y3 = 52ce6b4000000063 +x %s y3 = 0 +x /u 3 = 1b9a23c000000021 +x %u 3 = 0 +x /s 3 = 1b9a23c000000021 +x %s 3 = 0 +x /u 5 = 108faf0ccccccce0 +x %u 5 = 3 +x /s 5 = 108faf0ccccccce0 +x %s 5 = 3 +x /u 11 = 7872105d1745d20 +x %u 11 = 3 +x /s 11 = 7872105d1745d20 +x %s 11 = 3 +~x = ad3194bfffffff9c +x & y = 3 +x | y = 52ce6b4100000063 +x ^ y = 52ce6b4100000060 +x << i = 96735a0000000318 +x >>u i = a59cd680000000c +x >>s i = a59cd680000000c +x cmpu y = gt +x cmps y = gt +utod x = 43d4b39ad0000000 +dtou f = 21eadf559b3d0 +stod x = 43d4b39ad0000000 +dtos f = 21eadf559b3d0 +utof x = 5ea59cd7 +stof x = 5ea59cd7 + +x = 52ce6b4000000063 +y = 52ce6b4000000063 +-x = ad3194bfffffff9d +x + y = a59cd680000000c6 +x - y = 0 +x * y = ba6f38000002649 +x /u y = 1 +x %u y = 0 +x /s y = 1 +x %s y = 0 +x /u y2 = 100000000 +x %u y2 = 63 +x /s y3 = 100000000 +x %s y3 = 63 +x /u 3 = 1b9a23c000000021 +x %u 3 = 0 +x /s 3 = 1b9a23c000000021 +x %s 3 = 0 +x /u 5 = 108faf0ccccccce0 +x %u 5 = 3 +x /s 5 = 108faf0ccccccce0 +x %s 5 = 3 +x /u 11 = 7872105d1745d20 +x %u 11 = 3 +x /s 11 = 7872105d1745d20 +x %s 11 = 3 +~x = ad3194bfffffff9c +x & y = 52ce6b4000000063 +x | y = 52ce6b4000000063 +x ^ y = 0 +x << i = 31800000000 +x >>u i = a59cd68 +x >>s i = a59cd68 +x cmpu y = eq +x cmps y = eq +utod x = 43d4b39ad0000000 +dtou f = 21eadf559b3d0 +stod x = 43d4b39ad0000000 +dtos f = 21eadf559b3d0 +utof x = 5ea59cd7 +stof x = 5ea59cd7 + +x = 52ce6b4000000063 +y = 8362aa9340fe215f +-x = ad3194bfffffff9d +x + y = d63115d340fe21c2 +x - y = cf6bc0acbf01df04 +x * y = 8f1503b22246e7bd +x /u y = 0 +x %u y = 52ce6b4000000063 +x /s y = 0 +x %s y = 52ce6b4000000063 +x /u y2 = a158656f +x %u y2 = 5640ba6 +x /s y3 = ffffffff55e35d11 +x %s y3 = 5f2245a0 +x /u 3 = 1b9a23c000000021 +x %u 3 = 0 +x /s 3 = 1b9a23c000000021 +x %s 3 = 0 +x /u 5 = 108faf0ccccccce0 +x %u 5 = 3 +x /s 5 = 108faf0ccccccce0 +x %s 5 = 3 +x /u 11 = 7872105d1745d20 +x %u 11 = 3 +x /s 11 = 7872105d1745d20 +x %s 11 = 3 +~x = ad3194bfffffff9c +x & y = 2422a0000000043 +x | y = d3eeebd340fe217f +x ^ y = d1acc1d340fe213c +x << i = 3180000000 +x >>u i = a59cd680 +x >>s i = a59cd680 +x cmpu y = lt +x cmps y = gt +utod x = 43d4b39ad0000000 +dtou f = 21eadf559b3d0 +stod x = 43d4b39ad0000000 +dtos f = 21eadf559b3d0 +utof x = 5ea59cd7 +stof x = 5ea59cd7 + +x = f986342416ec8002 +y = 52ce6b4000000063 +-x = 679cbdbe9137ffe +x + y = 4c549f6416ec8065 +x - y = a6b7c8e416ec7f9f +x * y = b9230074dd7580c6 +x /u y = 3 +x %u y = 11af26416ec7ed9 +x /s y = 0 +x %s y = f986342416ec8002 +x /u y2 = 3036abea3 +x %u y2 = 164b642 +x /s y3 = ffffffffebfad66d +x %s y3 = ffffffffcae155c2 +x /u 3 = 532cbc0c07a42aab +x %u 3 = 1 +x /s 3 = fdd766b6b24ed556 +x %s 3 = 0 +x /u 5 = 31e7a40737c8e666 +x %u 5 = 4 +x /s 5 = feb470d40495b334 +x %s 5 = fffffffffffffffe +x /u 11 = 16af1c0347e6f45d +x %u 11 = 3 +x /s 11 = ff694a8eeacfae8c +x %s 11 = fffffffffffffffe +~x = 679cbdbe9137ffd +x & y = 5086200000000002 +x | y = fbce7f6416ec8063 +x ^ y = ab485f6416ec8061 +x << i = b764001000000000 +x >>u i = 1f30c684 +x >>s i = ffffffffff30c684 +x cmpu y = gt +x cmps y = lt +utod x = 43ef30c68482dd90 +dtou f = 6634832136daf +stod x = c399e72f6fa44e00 +dtos f = ffffd58f774c5ce4 +utof x = 5f798634 +stof x = dccf397b x = 368083376ba4ffa9 y = 6912b247b79a4904 @@ -7558,3 +8314,45 @@ dtos f = b3fdf698d581 utof x = 5ddbb784 stof x = 5ddbb784 +x = ca9a47c1649d27a7 +y = d56d650045e652aa +-x = 3565b83e9b62d859 +x + y = a007acc1aa837a51 +x - y = f52ce2c11eb6d4fd +x * y = 630e3c88ca19d2e6 +x /u y = 0 +x %u y = ca9a47c1649d27a7 +x /s y = 1 +x %s y = f52ce2c11eb6d4fd +x /u y2 = f3042098 +x %u y2 = 6b092fa7 +x /s y3 = 141176486 +x %s y3 = ffffffffdee649a7 +x /u 3 = 4388c295cc34628d +x %u 3 = 0 +x /s 3 = ee336d4076df0d38 +x %s 3 = ffffffffffffffff +x /u 5 = 2885418d141f6e54 +x %u 5 = 3 +x /s 5 = f5520e59e0ec3b22 +x %s 5 = fffffffffffffffd +x /u 11 = 126b1dcbc3541ae0 +x %u 11 = 7 +x /s 11 = fb254c57663cd510 +x %s 11 = fffffffffffffff7 +~x = 3565b83e9b62d858 +x & y = c0084500448402a2 +x | y = dfff67c165ff77af +x ^ y = 1ff722c1217b750d +x << i = 749e9c0000000000 +x >>u i = 32a691 +x >>s i = fffffffffff2a691 +x cmpu y = lt +x cmps y = lt +utod x = 43e95348f82c93a5 +dtou f = 52fc6dac31674 +stod x = c3cab2dc1f4db16c +dtos f = fffea20e1ffc05aa +utof x = 5f4a9a48 +stof x = de5596e1 + diff --git a/test/regression/int64.c b/test/regression/int64.c index d9785e95..0da9602d 100644 --- a/test/regression/int64.c +++ b/test/regression/int64.c @@ -103,7 +103,8 @@ u64 special_values[] = { 0x80000000LLU, 0x7FFFFFFFFFFFFFFFLLU, 0x8000000000000000LLU, - 0x100000003LLU + 0x100000003LLU, + 0x52ce6b4000000063LLU }; #define NUM_SPECIAL_VALUES (sizeof(special_values) / sizeof(u64)) -- cgit From f1abe04e503f1c54c5a50f7b3f3906beca15a760 Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Sun, 29 Mar 2020 11:20:01 +0200 Subject: Double rounding error in int64->float32 conversions on PowerPC and ARM The "stof" and "utof" runtime functions contain a round-to-odd step that avoids double rounding. However, this step was incorrectly coded on PowerPC (stof and utof), PowerPC64 (utof), and ARM (stof), making round-to-odd ineffective and causing double rounding. Closes: #343 --- runtime/arm/i64_stof.S | 9 ++++----- runtime/powerpc/i64_stof.s | 17 ++++++++--------- runtime/powerpc/i64_utof.s | 10 +++++----- runtime/powerpc64/i64_utof.s | 10 +++++----- 4 files changed, 22 insertions(+), 24 deletions(-) diff --git a/runtime/arm/i64_stof.S b/runtime/arm/i64_stof.S index bcfa471c..11e00a2a 100644 --- a/runtime/arm/i64_stof.S +++ b/runtime/arm/i64_stof.S @@ -39,12 +39,11 @@ @@@ Conversion from signed 64-bit integer to single float FUNCTION(__compcert_i64_stof) - @ Check whether -2^53 <= X < 2^53 - ASR r2, Reg0HI, #21 - ASR r3, Reg0HI, #31 @ (r2,r3) = X >> 53 + @ Check whether -2^53 <= X < 2^53 + ASR r2, Reg0HI, #21 @ r2 = high 32 bits of X >> 53 + @ -2^53 <= X < 2^53 iff r2 is -1 or 0, that is, iff r2 + 1 is 0 or 1 adds r2, r2, #1 - adc r3, r3, #0 @ (r2,r3) = X >> 53 + 1 - cmp r3, #2 + cmp r2, #2 blo 1f @ X is large enough that double rounding can occur. @ Avoid it by nudging X away from the points where double rounding diff --git a/runtime/powerpc/i64_stof.s b/runtime/powerpc/i64_stof.s index 97fa6bb8..ea23a1c8 100644 --- a/runtime/powerpc/i64_stof.s +++ b/runtime/powerpc/i64_stof.s @@ -43,20 +43,19 @@ __compcert_i64_stof: mflr r9 # Check whether -2^53 <= X < 2^53 - srawi r5, r3, 31 - srawi r6, r3, 21 # (r5,r6) = X >> 53 - addic r6, r6, 1 - addze r5, r5 # (r5,r6) = (X >> 53) + 1 + srawi r5, r3, 21 # r5 = high 32 bits of X >> 53 + # -2^53 <= X < 2^53 iff r5 is -1 or 0, that is, iff r5 + 1 is 0 or 1 + addi r5, r5, 1 cmplwi r5, 2 blt 1f # X is large enough that double rounding can occur. # Avoid it by nudging X away from the points where double rounding # occurs (the "round to odd" technique) - rlwinm r0, r4, 0, 21, 31 # extract bits 0 to 11 of X - addi r0, r0, 0x7FF # r0 = (X & 0x7FF) + 0x7FF - # bit 12 of r0 is 0 if all low 12 bits of X are 0, 1 otherwise - # bits 13-31 of r0 are 0 - or r4, r4, r0 # correct bit number 12 of X + rlwinm r5, r4, 0, 21, 31 # extract bits 0 to 11 of X + addi r5, r5, 0x7FF # r5 = (X & 0x7FF) + 0x7FF + # bit 12 of r5 is 0 if all low 12 bits of X are 0, 1 otherwise + # bits 13-31 of r5 are 0 + or r4, r4, r5 # correct bit number 12 of X rlwinm r4, r4, 0, 0, 20 # set to 0 bits 0 to 11 of X # Convert to double, then round to single 1: bl __compcert_i64_stod diff --git a/runtime/powerpc/i64_utof.s b/runtime/powerpc/i64_utof.s index cdb2f867..4a2a172b 100644 --- a/runtime/powerpc/i64_utof.s +++ b/runtime/powerpc/i64_utof.s @@ -48,11 +48,11 @@ __compcert_i64_utof: # X is large enough that double rounding can occur. # Avoid it by nudging X away from the points where double rounding # occurs (the "round to odd" technique) - rlwinm r0, r4, 0, 21, 31 # extract bits 0 to 11 of X - addi r0, r0, 0x7FF # r0 = (X & 0x7FF) + 0x7FF - # bit 12 of r0 is 0 if all low 12 bits of X are 0, 1 otherwise - # bits 13-31 of r0 are 0 - or r4, r4, r0 # correct bit number 12 of X + rlwinm r5, r4, 0, 21, 31 # extract bits 0 to 11 of X + addi r5, r5, 0x7FF # r5 = (X & 0x7FF) + 0x7FF + # bit 12 of r5 is 0 if all low 12 bits of X are 0, 1 otherwise + # bits 13-31 of r5 are 0 + or r4, r4, r5 # correct bit number 12 of X rlwinm r4, r4, 0, 0, 20 # set to 0 bits 0 to 11 of X # Convert to double, then round to single 1: bl __compcert_i64_utod diff --git a/runtime/powerpc64/i64_utof.s b/runtime/powerpc64/i64_utof.s index cdb2f867..4a2a172b 100644 --- a/runtime/powerpc64/i64_utof.s +++ b/runtime/powerpc64/i64_utof.s @@ -48,11 +48,11 @@ __compcert_i64_utof: # X is large enough that double rounding can occur. # Avoid it by nudging X away from the points where double rounding # occurs (the "round to odd" technique) - rlwinm r0, r4, 0, 21, 31 # extract bits 0 to 11 of X - addi r0, r0, 0x7FF # r0 = (X & 0x7FF) + 0x7FF - # bit 12 of r0 is 0 if all low 12 bits of X are 0, 1 otherwise - # bits 13-31 of r0 are 0 - or r4, r4, r0 # correct bit number 12 of X + rlwinm r5, r4, 0, 21, 31 # extract bits 0 to 11 of X + addi r5, r5, 0x7FF # r5 = (X & 0x7FF) + 0x7FF + # bit 12 of r5 is 0 if all low 12 bits of X are 0, 1 otherwise + # bits 13-31 of r5 are 0 + or r4, r4, r5 # correct bit number 12 of X rlwinm r4, r4, 0, 0, 20 # set to 0 bits 0 to 11 of X # Convert to double, then round to single 1: bl __compcert_i64_utod -- cgit From 144f466e3baa41e67d1fa908836a74536d52c201 Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Tue, 31 Mar 2020 11:06:51 +0200 Subject: Update Changelog --- Changelog | 37 ++++++++++++++++++++++++++++++++++--- 1 file changed, 34 insertions(+), 3 deletions(-) diff --git a/Changelog b/Changelog index 08586da5..9dc858ca 100644 --- a/Changelog +++ b/Changelog @@ -1,7 +1,38 @@ -Coq development: -- Compatibility with Coq version 8.11.0 (#316) +ISO C conformance: +- Functions declared `extern` then implemented `inline` remain `extern` +- The type of a wide char constant is `wchar_t`, not `int` +- Support vertical tabs and treat them as whitespace +- Define the semantics of `free(NULL)` + +Bug fixing: +- Take sign into account for conversions from 32-bit integers to 64-bit pointers +- PowerPC: more precise determination of small data accesses +- AArch64: when addressing global variables, check for correct alignment +- PowerPC, ARM: double rounding error in int64->float32 conversions + +ABI conformance: +- x86, AArch64: re-normalize values of small integer types returned by + function calls +- PowerPC: `float` arguments passed on stack are passed in 64-bit format +- RISC-V: use the new ELF psABI instead of the old ABI from ISA 2.1 + +Usability and diagnostics: +- Unknown builtin functions trigger a specific error message +- Improved error messages + +Coq formalization: +- Revised modeling of the PowerPC/EREF `isel` instruction +- Weaker `ec_readonly` condition over external calls + (permissions can be dropped on read-only locations) + +Coq and OCaml development: +- Compatibility with Coq version 8.10.1, 8.10.2, 8.11.0 +- Compatibility with OCaml 4.10 and up +- Compatibility with Menhir 20200123 and up +- Coq versions prior to 8.8.0 are no longer supported +- OCaml versions prior to 4.05.0 are no longer supported + - Release 3.6, 2019-09-17 ======================= -- cgit From 6dace9be5f4760882f879d3026c168cc9112e150 Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Tue, 31 Mar 2020 17:48:34 +0200 Subject: Updates for release 3.7 --- Changelog | 3 +++ doc/index.html | 4 +++- 2 files changed, 6 insertions(+), 1 deletion(-) diff --git a/Changelog b/Changelog index 9dc858ca..8cf4e548 100644 --- a/Changelog +++ b/Changelog @@ -1,3 +1,6 @@ +Release 3.7, 2020-03-31 +======================= + ISO C conformance: - Functions declared `extern` then implemented `inline` remain `extern` - The type of a wide char constant is `wchar_t`, not `int` diff --git a/doc/index.html b/doc/index.html index 3a4cf6ba..631c5d99 100644 --- a/doc/index.html +++ b/doc/index.html @@ -24,7 +24,7 @@ a:active {color : Red; text-decoration : underline; }

The CompCert verified compiler

Commented Coq development

-

Version 3.6, 2019-09-17

+

Version 3.7, 2020-03-31

Introduction

@@ -101,6 +101,8 @@ See also: Memdata (in-memory rep
  • Determinism: determinism properties of small-step semantics.
  • Op: operators, addressing modes and their semantics. +
  • Builtins: semantics of built-in functions.
    +See also: Builtins0 (target-independent part), Builtins1 (target-dependent part).
  • Unityping: a solver for atomic unification constraints. -- cgit From 76a4ff8f5b37429a614a2a97f628d9d862c93f46 Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Tue, 31 Mar 2020 18:24:12 +0200 Subject: Updates for release 3.7 --- VERSION | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/VERSION b/VERSION index 92686b06..b60e8d9b 100644 --- a/VERSION +++ b/VERSION @@ -1,3 +1,3 @@ -version=3.6 +version=3.7 buildnr= tag= -- cgit From 47f10ace41f8fb8ef818dbf1ca1d846b91b753c9 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 31 Mar 2020 19:10:14 +0200 Subject: forgot an 'Admitted' --- backend/CSE2proof.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/backend/CSE2proof.v b/backend/CSE2proof.v index fc980fb4..309ccce1 100644 --- a/backend/CSE2proof.v +++ b/backend/CSE2proof.v @@ -1680,7 +1680,7 @@ Proof. econstructor; split. eapply exec_return; eauto. constructor; auto. -Admitted. +Qed. Lemma transf_initial_states: -- cgit From aad09479f2ae3e008cbe38ca4fa0e8e3e04daec9 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 31 Mar 2020 19:18:49 +0200 Subject: add check-admitted --- .gitlab-ci.yml | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 79a32b25..4c00848f 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -1,6 +1,12 @@ stages: + - check-admitted - build +check-admitted: + stage: check-admitted + script: + - make check-admitted + build_x86_64: stage: build image: "coqorg/coq" -- cgit From 00b000ac303bf41434fae2765d10b0b719260d0c Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 31 Mar 2020 19:20:05 +0200 Subject: forgot image --- .gitlab-ci.yml | 1 + 1 file changed, 1 insertion(+) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 4c00848f..7d4578ef 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -4,6 +4,7 @@ stages: check-admitted: stage: check-admitted + image: "coqorg/coq" script: - make check-admitted -- cgit From 746cee9d2ea090c0b91dc358844f3456b8e91da8 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 31 Mar 2020 19:22:26 +0200 Subject: move check-admitted elsewhere --- .gitlab-ci.yml | 8 ++++++-- 1 file changed, 6 insertions(+), 2 deletions(-) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 7d4578ef..ed3cb261 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -1,11 +1,15 @@ stages: - - check-admitted - build check-admitted: - stage: check-admitted + stage: build image: "coqorg/coq" + before_script: + - opam switch 4.07.1+flambda + - eval `opam config env` + - opam install -y menhir script: + - ./config_x86_64.sh - make check-admitted build_x86_64: -- cgit From 6bfbd278803e14ddd8c74ae582e76607cca69591 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Tue, 31 Mar 2020 20:28:50 +0200 Subject: do not run check-admitted always --- .gitlab-ci.yml | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index ed3cb261..1f854fc3 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -11,6 +11,14 @@ check-admitted: script: - ./config_x86_64.sh - make check-admitted + rules: + - if: '$CI_COMMIT_BRANCH == "mppa-work"' + when: always + - if: '$CI_COMMIT_BRANCH == "mppa-k1c"' + when: always + - if: '$CI_COMMIT_BRANCH == "master"' + when: always + - when: manual build_x86_64: stage: build -- cgit From c974b25682251da237dbbe8ef3af218c6d175ae2 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 1 Apr 2020 10:34:57 +0200 Subject: Removing 8.8.* versions of coq in configure --- configure | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/configure b/configure index f13d1af3..f790281c 100755 --- a/configure +++ b/configure @@ -568,7 +568,7 @@ missingtools=false echo "Testing Coq... " | tr -d '\n' coq_ver=$(${COQBIN}coqc -v 2>/dev/null | sed -n -e 's/The Coq Proof Assistant, version \([^ ]*\).*$/\1/p') case "$coq_ver" in - 8.8.0|8.8.1|8.8.2|8.9.0|8.9.1|8.10.0|8.10.1|8.10.2|8.11.0) + 8.9.0|8.9.1|8.10.0|8.10.1|8.10.2|8.11.0) echo "version $coq_ver -- good!";; ?*) echo "version $coq_ver -- UNSUPPORTED" -- cgit From 027c5f9b643c554bef742bf907e725f8ad949429 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 1 Apr 2020 10:35:16 +0200 Subject: Fix cutrewrite deprecated --- mppa_k1c/PostpassSchedulingproof.v | 7 ++++--- 1 file changed, 4 insertions(+), 3 deletions(-) diff --git a/mppa_k1c/PostpassSchedulingproof.v b/mppa_k1c/PostpassSchedulingproof.v index 3b123c75..8cc7f0ab 100644 --- a/mppa_k1c/PostpassSchedulingproof.v +++ b/mppa_k1c/PostpassSchedulingproof.v @@ -61,9 +61,9 @@ Proof. - subst. repeat (rewrite Pregmap.gss); auto. destruct v; simpl; auto. rewrite Ptrofs.add_assoc. - cutrewrite (Ptrofs.repr (x + y) = Ptrofs.add (Ptrofs.repr x) (Ptrofs.repr y)); auto. + enough (Ptrofs.repr (x + y) = Ptrofs.add (Ptrofs.repr x) (Ptrofs.repr y)) as ->; auto. unfold Ptrofs.add. - cutrewrite (x + y = Ptrofs.unsigned (Ptrofs.repr x) + Ptrofs.unsigned (Ptrofs.repr y)); auto. + enough (x + y = Ptrofs.unsigned (Ptrofs.repr x) + Ptrofs.unsigned (Ptrofs.repr y)) as ->; auto. repeat (rewrite Ptrofs.unsigned_repr); auto. - repeat (rewrite Pregmap.gso; auto). Qed. @@ -220,7 +220,8 @@ Proof. destruct (zeq pos 0). + inv H. exists lbb. constructor; auto. + apply IHlbb in H. destruct H as (c & TAIL). exists c. - cutrewrite (pos = pos - size a + size a). apply code_tail_S; auto. + enough (pos = pos - size a + size a) as ->. + apply code_tail_S; auto. omega. Qed. -- cgit From 7ad6991534ba4ab10fe29d5456393f45cb4e5605 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 1 Apr 2020 10:35:42 +0200 Subject: -fduplicate -1 really desactivates the pass in Coq now --- driver/Compiler.v | 10 +++++----- driver/Compopts.v | 3 +-- extraction/extraction.v | 2 ++ 3 files changed, 8 insertions(+), 7 deletions(-) diff --git a/driver/Compiler.v b/driver/Compiler.v index da19a0b9..499feff2 100644 --- a/driver/Compiler.v +++ b/driver/Compiler.v @@ -134,7 +134,7 @@ Definition transf_rtl_program (f: RTL.program) : res Asm.program := @@ print (print_RTL 2) @@ time "Renumbering" Renumber.transf_program @@ print (print_RTL 3) - @@@ time "Tail-duplicating" Duplicate.transf_program + @@@ partial_if Compopts.optim_duplicate (time "Tail-duplicating" Duplicate.transf_program) @@ print (print_RTL 4) @@ total_if Compopts.optim_constprop (time "Constant propagation" Constprop.transf_program) @@ print (print_RTL 5) @@ -254,7 +254,7 @@ Definition CompCert's_passes := ::: mkpass (match_if Compopts.optim_tailcalls Tailcallproof.match_prog) ::: mkpass Inliningproof.match_prog ::: mkpass Renumberproof.match_prog - ::: mkpass Duplicateproof.match_prog + ::: mkpass (match_if Compopts.optim_duplicate Duplicateproof.match_prog) ::: mkpass (match_if Compopts.optim_constprop Constpropproof.match_prog) ::: mkpass (match_if Compopts.optim_constprop Renumberproof.match_prog) ::: mkpass (match_if Compopts.optim_CSE CSEproof.match_prog) @@ -301,7 +301,7 @@ Proof. set (p7 := total_if optim_tailcalls Tailcall.transf_program p6) in *. destruct (Inlining.transf_program p7) as [p8|e] eqn:P8; simpl in T; try discriminate. set (p9 := Renumber.transf_program p8) in *. - destruct (Duplicate.transf_program p9) as [p10|e] eqn:P10; simpl in T; try discriminate. + destruct (partial_if optim_duplicate Duplicate.transf_program p9) as [p10|e] eqn:P10; simpl in T; try discriminate. set (p11 := total_if optim_constprop Constprop.transf_program p10) in *. set (p12 := total_if optim_constprop Renumber.transf_program p11) in *. destruct (partial_if optim_CSE CSE.transf_program p12) as [p13|e] eqn:P13; simpl in T; try discriminate. @@ -326,7 +326,7 @@ Proof. exists p7; split. apply total_if_match. apply Tailcallproof.transf_program_match. exists p8; split. apply Inliningproof.transf_program_match; auto. exists p9; split. apply Renumberproof.transf_program_match; auto. - exists p10; split. apply Duplicateproof.transf_program_match; auto. + exists p10; split. eapply partial_if_match; eauto. apply Duplicateproof.transf_program_match; auto. exists p11; split. apply total_if_match. apply Constpropproof.transf_program_match. exists p12; split. apply total_if_match. apply Renumberproof.transf_program_match. exists p13; split. eapply partial_if_match; eauto. apply CSEproof.transf_program_match. @@ -412,7 +412,7 @@ Ltac DestructM := eapply Inliningproof.transf_program_correct; eassumption. eapply compose_forward_simulations. eapply Renumberproof.transf_program_correct; eassumption. eapply compose_forward_simulations. - eapply Duplicateproof.transf_program_correct; eassumption. + eapply match_if_simulation. eassumption. exact Duplicateproof.transf_program_correct. eapply compose_forward_simulations. eapply match_if_simulation. eassumption. exact Constpropproof.transf_program_correct. eapply compose_forward_simulations. diff --git a/driver/Compopts.v b/driver/Compopts.v index b4b9f30d..848657e5 100644 --- a/driver/Compopts.v +++ b/driver/Compopts.v @@ -27,8 +27,7 @@ Parameter generate_float_constants: unit -> bool. (** For value analysis. Currently always false. *) Parameter va_strict: unit -> bool. -(** Flag -fduplicate. For tail duplication optimization. Necessary to have - * bigger superblocks *) +(** Flag -fduplicate. Branch prediction annotation + tail duplication *) Parameter optim_duplicate: unit -> bool. (** Flag -ftailcalls. For tail call optimization. *) diff --git a/extraction/extraction.v b/extraction/extraction.v index ba6b080b..9b568951 100644 --- a/extraction/extraction.v +++ b/extraction/extraction.v @@ -105,6 +105,8 @@ Extract Constant Compopts.generate_float_constants => "fun _ -> !Clflags.option_ffloatconstprop >= 2". Extract Constant Compopts.optim_tailcalls => "fun _ -> !Clflags.option_ftailcalls". +Extract Constant Compopts.optim_duplicate => + "fun _ -> (if !Clflags.option_fduplicate = -1 then false else true)". Extract Constant Compopts.optim_constprop => "fun _ -> !Clflags.option_fconstprop". Extract Constant Compopts.optim_CSE => -- cgit From c34e25a208e092aff0b7dfa931b199df0ce3bc52 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 1 Apr 2020 11:53:55 +0200 Subject: Fixing packedstruct issue --- test/regression/Makefile | 6 +++--- test/regression/packedstruct1.c | 24 ++++++++++++------------ 2 files changed, 15 insertions(+), 15 deletions(-) diff --git a/test/regression/Makefile b/test/regression/Makefile index ad3ffd99..97c25f6c 100644 --- a/test/regression/Makefile +++ b/test/regression/Makefile @@ -26,10 +26,10 @@ TESTS_COMP?=attribs1 bitfields1 bitfields2 bitfields3 bitfields4 \ varargs1 varargs2 varargs3 sections alias aligned\ packedstruct1 packedstruct2 -# FIXME ifeq ($(ARCH),mppa_k1c) +ifeq ($(ARCH),mppa_k1c) TESTS_COMP:=$(filter-out packedstruct1,$(TESTS_COMP)) TESTS_COMP:=$(filter-out packedstruct2,$(TESTS_COMP)) -# endif +endif # Can run, both in compiled mode and in interpreter mode, # but produce processor-dependent results, so no reference output in Results @@ -37,7 +37,7 @@ TESTS_COMP?=attribs1 bitfields1 bitfields2 bitfields3 bitfields4 \ TESTS_DIFF=NaNs # FIXME ifeq ($(ARCH),mppa_k1c) TESTS_DIFF:=$(filter-out NaNs,$(TESTS_DIFF)) -# endif endif +# endif # Other tests: should compile to .s without errors (but expect warnings) diff --git a/test/regression/packedstruct1.c b/test/regression/packedstruct1.c index ac68c698..b805c92a 100644 --- a/test/regression/packedstruct1.c +++ b/test/regression/packedstruct1.c @@ -23,9 +23,9 @@ void test1(void) struct s1 s1; printf("sizeof(struct s1) = %d\n", szof(s1)); printf("precomputed sizeof(struct s1) = %d\n", bszof(s1)); - printf("offsetOf(x) = %d, offsetOf(y) = %d, offsetOf(z) = %d\n", + printf("offsetof(x) = %d, offsetof(y) = %d, offsetof(z) = %d\n", offsetOf(s1,x), offsetOf(s1,y), offsetOf(s1,z)); - printf("precomputed offsetOf(x) = %d, offsetOf(y) = %d, offsetOf(z) = %d\n", + printf("precomputed offsetof(x) = %d, offsetof(y) = %d, offsetof(z) = %d\n", boffsetof(s1,x), boffsetof(s1,y), boffsetof(s1,z)); s1.x = 123; s1.y = -456; s1.z = 3.14159; printf("s1 = {x = %d, y = %d, z = %.5f}\n\n", s1.x, s1.y, s1.z); @@ -44,9 +44,9 @@ void test2(void) printf("sizeof(struct s2) = %d\n", szof(s2)); printf("precomputed sizeof(struct s2) = %d\n", bszof(s2)); printf("&s2 mod 16 = %d\n", ((int) &s2) & 0xF); - printf("offsetOf(x) = %d, offsetOf(y) = %d, offsetOf(z) = %d\n", + printf("offsetof(x) = %d, offsetof(y) = %d, offsetof(z) = %d\n", offsetOf(s2,x), offsetOf(s2,y), offsetOf(s2,z)); - printf("precomputed offsetOf(x) = %d, offsetOf(y) = %d, offsetOf(z) = %d\n", + printf("precomputed offsetof(x) = %d, offsetof(y) = %d, offsetof(z) = %d\n", boffsetof(s2,x), boffsetof(s2,y), boffsetof(s2,z)); s2.x = 12345; s2.y = -456; s2.z = 3.14159; printf("s2 = {x = %d, y = %d, z = %.5f}\n\n", s2.x, s2.y, s2.z); @@ -73,8 +73,8 @@ void test3(void) printf("sizeof(struct s3) = %d\n", szof(s3)); printf("precomputed sizeof(struct s3) = %d\n", bszof(s3)); - printf("offsetOf(s) = %d\n", offsetOf(s3,s)); - printf("precomputed offsetOf(s) = %d\n", boffsetof(s3,s)); + printf("offsetof(s) = %d\n", offsetOf(s3,s)); + printf("precomputed offsetof(s) = %d\n", boffsetof(s3,s)); s3.x = 123; s3.y = 45678; s3.z = 0x80000001U; @@ -103,9 +103,9 @@ void test4(void) printf("sizeof(struct s4) = %d\n", szof(s4)); printf("precomputed sizeof(struct s4) = %d\n", bszof(s4)); - printf("offsetOf(x) = %d, offsetOf(y) = %d, offsetOf(z) = %d\n", + printf("offsetof(x) = %d, offsetof(y) = %d, offsetof(z) = %d\n", offsetOf(s4,x), offsetOf(s4,y), offsetOf(s4,z)); - printf("precomputed offsetOf(x) = %d, offsetOf(y) = %d, offsetOf(z) = %d\n", + printf("precomputed offsetof(x) = %d, offsetof(y) = %d, offsetof(z) = %d\n", boffsetof(s4,x), boffsetof(s4,y), boffsetof(s4,z)); s4.x = 123; s4.y = -456; s4.z = 3.14159; printf("s4 = {x = %d, y = %d, z = %.5f}\n\n", s4.x, s4.y, s4.z); @@ -121,9 +121,9 @@ void test5(void) printf("sizeof(struct s5) = %d\n", szof(s5)); printf("precomputed sizeof(struct s5) = %d\n", bszof(s5)); - printf("offsetOf(x) = %d, offsetOf(y) = %d, offsetOf(z) = %d\n", + printf("offsetof(x) = %d, offsetof(y) = %d, offsetof(z) = %d\n", offsetOf(s5,x), offsetOf(s5,y), offsetOf(s5,z)); - printf("precomputed offsetOf(x) = %d, offsetOf(y) = %d, offsetOf(z) = %d\n", + printf("precomputed offsetof(x) = %d, offsetof(y) = %d, offsetof(z) = %d\n", boffsetof(s5,x), boffsetof(s5,y), boffsetof(s5,z)); s5.x = 123; s5.y = -456; s5.z = 3.14159; printf("s5 = {x = %d, y = %d, z = %.5f}\n\n", s5.x, s5.y, s5.z); @@ -139,9 +139,9 @@ void test6(void) printf("sizeof(struct s6) = %d\n", szof(s6)); printf("precomputed sizeof(struct s6) = %d\n", bszof(s6)); - printf("offsetOf(x) = %d, offsetOf(y) = %d, offsetOf(z) = %d\n", + printf("offsetof(x) = %d, offsetof(y) = %d, offsetof(z) = %d\n", offsetOf(s6,x), offsetOf(s6,y), offsetOf(s6,z)); - printf("precomputed offsetOf(x) = %d, offsetOf(y) = %d, offsetOf(z) = %d\n", + printf("precomputed offsetof(x) = %d, offsetof(y) = %d, offsetof(z) = %d\n", boffsetof(s6,x), boffsetof(s6,y), boffsetof(s6,z)); s62.x = 123; s62.y = -456; s62.z = 3.14159; printf("s62 = {x = %d, y = %d, z = %.5f}\n\n", s62.x, s62.y, s62.z); -- cgit From 6e7c693e6cfe683b7a44c4f2a3420678fcdcc36f Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 1 Apr 2020 14:24:52 +0200 Subject: Stopping traces at join points --- backend/Linearizeaux.ml | 27 +++++++++++++++++++++++++-- 1 file changed, 25 insertions(+), 2 deletions(-) diff --git a/backend/Linearizeaux.ml b/backend/Linearizeaux.ml index 605a5db5..bfa056ca 100644 --- a/backend/Linearizeaux.ml +++ b/backend/Linearizeaux.ml @@ -150,8 +150,27 @@ let print_plist l = Printf.printf "]" end +(* adapted from the above join_points function, but with PTree *) +let get_join_points code entry = + let reached = ref (PTree.map (fun n i -> false) code) in + let reached_twice = ref (PTree.map (fun n i -> false) code) in + let rec traverse pc = + if get_some @@ PTree.get pc !reached then begin + if not (get_some @@ PTree.get pc !reached_twice) then + reached_twice := PTree.set pc true !reached_twice + end else begin + reached := PTree.set pc true !reached; + traverse_succs (successors_block @@ get_some @@ PTree.get pc code) + end + and traverse_succs = function + | [] -> () + | [pc] -> traverse pc + | pc :: l -> traverse pc; traverse_succs l + in traverse entry; !reached_twice + let forward_sequences code entry = let visited = ref (PTree.map (fun n i -> false) code) in + let join_points = get_join_points code entry in (* returns the list of traversed nodes, and a list of nodes to start traversing next *) let rec traverse_fallthrough code node = (* Printf.printf "Traversing %d..\n" (P.to_int node); *) @@ -164,10 +183,14 @@ let forward_sequences code entry = | Lop _ | Lload _ | Lgetstack _ | Lsetstack _ | Lstore _ | Lcall _ | Lbuiltin _ -> assert false | Ltailcall _ | Lreturn -> begin (* Printf.printf "STOP tailcall/return\n"; *) ([], []) end - | Lbranch n -> let ln, rem = traverse_fallthrough code n in (ln, rem) + | Lbranch n -> + if get_some @@ PTree.get n join_points then ([], [n]) + else let ln, rem = traverse_fallthrough code n in (ln, rem) | Lcond (_, _, ifso, ifnot, info) -> (match info with | None -> begin (* Printf.printf "STOP Lcond None\n"; *) ([], [ifso; ifnot]) end - | Some false -> let ln, rem = traverse_fallthrough code ifnot in (ln, [ifso] @ rem) + | Some false -> + if get_some @@ PTree.get ifnot join_points then ([], [ifso; ifnot]) + else let ln, rem = traverse_fallthrough code ifnot in (ln, [ifso] @ rem) | Some true -> let errstr = Printf.sprintf ("Inconsistency detected in node %d: ifnot is not the preferred branch") (P.to_int node) in failwith errstr) -- cgit From 1a70cffa6080d0d9f90bfa7541e46737c9588212 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Thu, 2 Apr 2020 16:23:10 +0200 Subject: Fixing loop heuristic --- backend/Duplicateaux.ml | 49 ++++++++++++++++++++++++++++++++++--------------- 1 file changed, 34 insertions(+), 15 deletions(-) diff --git a/backend/Duplicateaux.ml b/backend/Duplicateaux.ml index 28ad4266..1f4a693d 100644 --- a/backend/Duplicateaux.ml +++ b/backend/Duplicateaux.ml @@ -105,6 +105,7 @@ let get_loop_headers code entrypoint = begin match (get_some @@ PTree.get node !visited) with | Visited -> () | Processed -> begin + Printf.printf "Node %d is a loop header\n" (P.to_int node); is_loop_header := PTree.set node true !is_loop_header; visited := PTree.set node Visited !visited end @@ -238,19 +239,36 @@ let get_loop_info is_loop_header bfs_order code = | Icond (_,_,s1,s2,_) -> (explore s1 dest) || (explore s2 dest) | Ijumptable _ | Itailcall _ | Ireturn _ -> false end - in match get_some @@ PTree.get s !loop_info with - | None -> begin - match get_some @@ PTree.get s code with - | Icond (_, _, n1, n2, _) -> - let b1 = explore n1 n in - let b2 = explore n2 n in - if (b1 && b2) then () - else if b1 then loop_info := PTree.set s (Some true) !loop_info - else if b2 then loop_info := PTree.set s (Some false) !loop_info - else () - | _ -> () + in let rec advance_to_cb src = + if (get_some @@ PTree.get src !visited) then None + else begin + visited := PTree.set src true !visited; + match get_some @@ PTree.get src code with + | Inop s | Iop (_, _, _, s) | Iload (_,_,_,_,_,s) | Istore (_,_,_,_,s) | Icall (_,_,_,_,s) + | Ibuiltin (_,_,_,s) -> advance_to_cb s + | Icond _ -> Some src + | Ijumptable _ | Itailcall _ | Ireturn _ -> None end - | Some _ -> () + in begin + Printf.printf "Marking path from %d to %d\n" (P.to_int n) (P.to_int s); + match advance_to_cb s with + | None -> (Printf.printf "Nothing found\n") + | Some s -> ( Printf.printf "Advancing to %d\n" (P.to_int s); + match get_some @@ PTree.get s !loop_info with + | None | Some _ -> begin + match get_some @@ PTree.get s code with + | Icond (_, _, n1, n2, _) -> + let b1 = explore n1 n in + let b2 = explore n2 n in + if (b1 && b2) then (Printf.printf "both true\n") + else if b1 then (Printf.printf "true privileged\n"; loop_info := PTree.set s (Some true) !loop_info) + else if b2 then (Printf.printf "false privileged\n"; loop_info := PTree.set s (Some false) !loop_info) + else (Printf.printf "none true\n") + | _ -> ( Printf.printf "not an icond\n" ) + end + (* | Some _ -> ( Printf.printf "already loop info there\n" ) FIXME - we don't know yet whether a branch to a loop head is a backedge or not *) + ) + end in begin List.iter (fun n -> match get_some @@ PTree.get n code with @@ -527,7 +545,7 @@ let rec change_pointers code n n' = function * n': the integer which should contain the duplicate of n * returns: new code, new ptree *) let duplicate code ptree parent n preds n' = - (* Printf.printf "Duplicating node %d into %d..\n" (P.to_int n) (P.to_int n'); *) + Printf.printf "Duplicating node %d into %d..\n" (P.to_int n) (P.to_int n'); match PTree.get n' code with | Some _ -> failwith "The PTree already has a node n'" | None -> @@ -591,8 +609,9 @@ let superblockify_traces code preds traces = | [] -> (code, ptree, 0) | trace :: traces -> let new_code, new_ptree, nb_duplicated = tail_duplicate code preds ptree trace - in if (nb_duplicated < max_nb_duplicated) then f new_code new_ptree traces - else (Printf.printf "Too many duplicated nodes, aborting tail duplication\n"; (code, ptree, 0)) + in if (nb_duplicated < max_nb_duplicated) + then (Printf.printf "End duplication\n"; f new_code new_ptree traces) + else (Printf.printf "Too many duplicated nodes, aborting tail duplication\n"; (code, ptree, 0)) in let new_code, new_ptree, _ = f code ptree traces in (new_code, new_ptree) -- cgit From c6356cdc5f567a317afcb99cb004354cf7dcce0f Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Fri, 3 Apr 2020 11:11:19 +0200 Subject: Changing best_predecessor_of to not take None predictions --- backend/Duplicateaux.ml | 19 +++++++++++++++---- 1 file changed, 15 insertions(+), 4 deletions(-) diff --git a/backend/Duplicateaux.ml b/backend/Duplicateaux.ml index 1f4a693d..98e2f325 100644 --- a/backend/Duplicateaux.ml +++ b/backend/Duplicateaux.ml @@ -413,11 +413,22 @@ let best_successor_of node code is_visited = | Some n -> if not (ptree_get_some n is_visited) then Some n else None (* FIXME - could be improved by selecting in priority the predicted paths *) -let best_predecessor_of node predecessors order is_visited = +let best_predecessor_of node predecessors code order is_visited = match (PTree.get node predecessors) with | None -> failwith "No predecessor list found" - | Some lp -> try Some (List.find (fun n -> (List.mem n lp) && (not (ptree_get_some n is_visited))) order) - with Not_found -> None + | Some lp -> + try Some (List.find (fun n -> + if (List.mem n lp) && (not (ptree_get_some n is_visited)) then + match ptree_get_some n code with + | Icond (_, _, n1, n2, ob) -> (match ob with + | None -> false + | Some false -> n == n2 + | Some true -> n == n1 + ) + | _ -> true + else false + ) order) + with Not_found -> None let print_trace t = print_intlist t @@ -489,7 +500,7 @@ let select_traces_chang code entrypoint = begin current := seed; quit_loop := false; while not !quit_loop do - let s = best_predecessor_of !current predecessors order !is_visited in + let s = best_predecessor_of !current predecessors code order !is_visited in match s with | None -> quit_loop := true (* if (s==0) exit loop *) | Some pred -> begin -- cgit From 7d60bff91b1ede7475f703f2d9eb926d11345bf9 Mon Sep 17 00:00:00 2001 From: David Monniaux Date: Wed, 8 Apr 2020 08:51:28 +0200 Subject: accept Coq 8.11.1 --- configure | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/configure b/configure index f790281c..cb2f52ba 100755 --- a/configure +++ b/configure @@ -568,7 +568,7 @@ missingtools=false echo "Testing Coq... " | tr -d '\n' coq_ver=$(${COQBIN}coqc -v 2>/dev/null | sed -n -e 's/The Coq Proof Assistant, version \([^ ]*\).*$/\1/p') case "$coq_ver" in - 8.9.0|8.9.1|8.10.0|8.10.1|8.10.2|8.11.0) + 8.9.0|8.9.1|8.10.0|8.10.1|8.10.2|8.11.0|8.11.1) echo "version $coq_ver -- good!";; ?*) echo "version $coq_ver -- UNSUPPORTED" -- cgit From ba6453483f7c742a98cd6fcefe015018df1dfea7 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 8 Apr 2020 11:57:37 +0200 Subject: Duplicate: Common rtl_successors function --- backend/Duplicateaux.ml | 94 +++++++++++++++++-------------------------------- 1 file changed, 33 insertions(+), 61 deletions(-) diff --git a/backend/Duplicateaux.ml b/backend/Duplicateaux.ml index 98e2f325..b137e872 100644 --- a/backend/Duplicateaux.ml +++ b/backend/Duplicateaux.ml @@ -1,3 +1,9 @@ +(* Oracle for Duplicate pass. + * - Add static prediction information to Icond nodes + * - Performs tail duplication on interesting traces to form superblocks + * - (TODO: perform partial loop unrolling inside innermost loops) + *) + open RTL open Maps open Camlcoq @@ -6,6 +12,13 @@ let get_some = function | None -> failwith "Did not get some" | Some thing -> thing +let rtl_successors = function +| Itailcall _ | Ireturn _ -> [] +| Icall(_,_,_,_,n) | Ibuiltin(_,_,_,n) | Inop n | Iop (_,_,_,n) +| Iload (_,_,_,_,_,n) | Istore (_,_,_,_,n) -> [n] +| Icond (_,_,n1,n2,_) -> [n1; n2] +| Ijumptable (_,ln) -> ln + let bfs code entrypoint = begin Printf.printf "bfs\n"; flush stdout; let visited = ref (PTree.map (fun n i -> false) code) @@ -22,13 +35,8 @@ let bfs code entrypoint = begin | None -> failwith "No such node" | Some i -> bfs_list := !node :: !bfs_list; - match i with - | Icall(_, _, _, _, n) -> Queue.add n to_visit - | Ibuiltin(_, _, _, n) -> Queue.add n to_visit - | Ijumptable(_, ln) -> List.iter (fun n -> Queue.add n to_visit) ln - | Itailcall _ | Ireturn _ -> () - | Icond (_, _, n1, n2, _) -> Queue.add n1 to_visit; Queue.add n2 to_visit - | Inop n | Iop (_, _, _, n) | Iload (_, _, _, _, _, n) | Istore (_, _, _, _, n) -> Queue.add n to_visit + let succ = rtl_successors i in + List.iter (fun n -> Queue.add n to_visit) succ end done; List.rev !bfs_list @@ -43,12 +51,7 @@ let get_predecessors_rtl code = begin Printf.printf "get_predecessors_rtl\n"; flush stdout; let preds = ref (PTree.map (fun n i -> []) code) in let process_inst (node, i) = - let succ = match i with - | Inop n | Iop (_,_,_,n) | Iload (_, _,_,_,_,n) | Istore (_,_,_,_,n) - | Icall (_,_,_,_,n) | Ibuiltin (_, _, _, n) -> [n] - | Icond (_,_,n1,n2,_) -> [n1;n2] - | Ijumptable (_,ln) -> ln - | Itailcall _ | Ireturn _ -> [] + let succ = rtl_successors i in List.iter (fun s -> let previous_preds = ptree_get_some s !preds in if optbool @@ List.find_opt (fun e -> e == node) previous_preds then () @@ -113,13 +116,7 @@ let get_loop_headers code entrypoint = begin visited := PTree.set node Processed !visited; match PTree.get node code with | None -> failwith "No such node" - | Some i -> let next_visits = (match i with - | Icall (_, _, _, _, n) | Ibuiltin (_, _, _, n) | Inop n | Iop (_, _, _, n) - | Iload (_, _, _, _, _, n) | Istore (_, _, _, _, n) -> [n] - | Icond (_, _, n1, n2, _) -> [n1; n2] - | Itailcall _ | Ireturn _ -> [] - | Ijumptable (_, ln) -> ln - ) in dfs_visit code next_visits; + | Some i -> let next_visits = rtl_successors i in dfs_visit code next_visits; visited := PTree.set node Visited !visited; dfs_visit code ln end @@ -143,16 +140,13 @@ let ptree_printbool pt = * the given predicate *) let rec look_ahead code node is_loop_header predicate = if (predicate node) then true - else match (get_some @@ PTree.get node code) with - | Ireturn _ | Itailcall _ | Icond _ | Ijumptable _ -> false - | Inop n | Iop (_, _, _, n) | Iload (_, _, _, _, _, n) - | Istore (_, _, _, _, n) | Icall (_, _, _, _, n) - | Ibuiltin (_, _, _, n) -> - if (predicate n) then true - else ( - if (get_some @@ PTree.get n is_loop_header) then false - else look_ahead code n is_loop_header predicate - ) + else match (rtl_successors @@ get_some @@ PTree.get node code) with + | [n] -> if (predicate n) then true + else ( + if (get_some @@ PTree.get n is_loop_header) then false + else look_ahead code n is_loop_header predicate + ) + | _ -> false let do_call_heuristic code cond ifso ifnot is_loop_header = begin @@ -233,11 +227,11 @@ let get_loop_info is_loop_header bfs_order code = else if src == dest then true else begin visited := PTree.set src true !visited; - match get_some @@ PTree.get src code with - | Inop s | Iop (_, _, _, s) | Iload (_,_,_,_,_,s) | Istore (_,_,_,_,s) | Icall (_,_,_,_,s) - | Ibuiltin (_,_,_,s) -> explore s dest - | Icond (_,_,s1,s2,_) -> (explore s1 dest) || (explore s2 dest) - | Ijumptable _ | Itailcall _ | Ireturn _ -> false + match rtl_successors @@ get_some @@ PTree.get src code with + | [] -> false + | [s] -> explore s dest + | [s1; s2] -> (explore s1 dest) || (explore s2 dest) + | _ -> false end in let rec advance_to_cb src = if (get_some @@ PTree.get src !visited) then None @@ -275,14 +269,14 @@ let get_loop_info is_loop_header bfs_order code = | Inop s | Iop (_,_,_,s) | Iload (_,_,_,_,_,s) | Istore (_,_,_,_,s) | Icall (_,_,_,_,s) | Ibuiltin (_, _, _, s) -> if get_some @@ PTree.get s is_loop_header then mark_path s n - | Icond _ -> () (* loop backedges are never Icond in CompCert *) + | Icond _ -> () (* loop backedges are never Icond in CompCert RTL.3 *) | Ijumptable _ -> () | Itailcall _ | Ireturn _ -> () ) bfs_order; !loop_info end - (* Remark - compared to the original paper, we don't use the store heuristic *) +(* Remark - compared to the original paper, we don't use the store heuristic *) let get_directions code entrypoint = begin Printf.printf "get_directions\n"; flush stdout; let bfs_order = bfs code entrypoint in @@ -373,24 +367,6 @@ let dfs code entrypoint = begin in dfs_list code [entrypoint] end -(* -let get_predecessors_ttl code = - let preds = ref (PTree.map (fun n i -> []) code) in - let process_inst (node, ti) = match ti with - | Tleaf _ -> () - | Tnext (_, i) -> let succ = match i with - | Inop n | Iop (_,_,_,n) | Iload (_, _,_,_,_,n) | Istore (_,_,_,_,n) - | Icall (_,_,_,_,n) | Ibuiltin (_, _, _, n) -> [n] - | Icond (_,_,n1,n2,_) -> [n1;n2] - | Ijumptable (_,ln) -> ln - | _ -> [] - in List.iter (fun s -> preds := PTree.set s (node::(get_some @@ PTree.get s !preds)) !preds) succ - in begin - List.iter process_inst (PTree.elements code); - !preds - end -*) - let rec select_unvisited_node is_visited = function | [] -> failwith "Empty list" | n :: ln -> if not (ptree_get_some n is_visited) then n else select_unvisited_node is_visited ln @@ -400,12 +376,8 @@ let best_successor_of node code is_visited = | None -> failwith "No such node in the code" | Some i -> let next_node = match i with - | Inop n -> Some n - | Iop (_, _, _, n) -> Some n - | Iload (_, _, _, _, _, n) -> Some n - | Istore (_, _, _, _, n) -> Some n - | Icall (_, _, _, _, n) -> Some n - | Ibuiltin (_, _, _, n) -> Some n + | Inop n | Iop (_,_,_,n) | Iload (_,_,_,_,_,n) | Istore(_,_,_,_,n) + | Icall (_,_,_,_,n) | Ibuiltin (_,_,_,n) -> Some n | Icond (_, _, n1, n2, ob) -> (match ob with None -> None | Some false -> Some n2 | Some true -> Some n1) | _ -> None in match next_node with -- cgit From e326ed9f28a2ed6869f0cb356ef9a8e189cb0a47 Mon Sep 17 00:00:00 2001 From: Cyril SIX Date: Wed, 8 Apr 2020 14:53:50 +0200 Subject: Some cleaning on Linearize and Duplicate --- backend/Duplicateaux.ml | 124 +++++++++++++++++++++++++++--------------------- backend/Linearizeaux.ml | 89 ++++++++++++++++++---------------- 2 files changed, 118 insertions(+), 95 deletions(-) diff --git a/backend/Duplicateaux.ml b/backend/Duplicateaux.ml index b137e872..89f187da 100644 --- a/backend/Duplicateaux.ml +++ b/backend/Duplicateaux.ml @@ -8,6 +8,12 @@ open RTL open Maps open Camlcoq +let debug_flag = ref false + +let debug fmt = + if !debug_flag then Printf.eprintf fmt + else Printf.ifprintf stderr fmt + let get_some = function | None -> failwith "Did not get some" | Some thing -> thing @@ -20,7 +26,7 @@ let rtl_successors = function | Ijumptable (_,ln) -> ln let bfs code entrypoint = begin - Printf.printf "bfs\n"; flush stdout; + debug "bfs\n"; let visited = ref (PTree.map (fun n i -> false) code) and bfs_list = ref [] and to_visit = Queue.create () @@ -48,7 +54,7 @@ let optbool o = match o with Some _ -> true | None -> false let ptree_get_some n ptree = get_some @@ PTree.get n ptree let get_predecessors_rtl code = begin - Printf.printf "get_predecessors_rtl\n"; flush stdout; + debug "get_predecessors_rtl\n"; let preds = ref (PTree.map (fun n i -> []) code) in let process_inst (node, i) = let succ = rtl_successors i @@ -74,19 +80,23 @@ let print_intlist l = | [] -> () | n::ln -> (Printf.printf "%d " (P.to_int n); f ln) in begin - Printf.printf "["; - f l; - Printf.printf "]" + if !debug_flag then begin + Printf.printf "["; + f l; + Printf.printf "]" + end end let print_intset s = let seq = PSet.to_seq s in begin - Printf.printf "{"; - Seq.iter (fun n -> - Printf.printf "%d " (P.to_int n) - ) seq; - Printf.printf "}" + if !debug_flag then begin + Printf.printf "{"; + Seq.iter (fun n -> + Printf.printf "%d " (P.to_int n) + ) seq; + Printf.printf "}" + end end type vstate = Unvisited | Processed | Visited @@ -99,7 +109,7 @@ type vstate = Unvisited | Processed | Visited * If we come accross an edge to a Processed node, it's a loop! *) let get_loop_headers code entrypoint = begin - Printf.printf "get_loop_headers\n"; flush stdout; + debug "get_loop_headers\n"; let visited = ref (PTree.map (fun n i -> Unvisited) code) and is_loop_header = ref (PTree.map (fun n i -> false) code) in let rec dfs_visit code = function @@ -108,7 +118,7 @@ let get_loop_headers code entrypoint = begin match (get_some @@ PTree.get node !visited) with | Visited -> () | Processed -> begin - Printf.printf "Node %d is a loop header\n" (P.to_int node); + debug "Node %d is a loop header\n" (P.to_int node); is_loop_header := PTree.set node true !is_loop_header; visited := PTree.set node Visited !visited end @@ -129,11 +139,13 @@ end let ptree_printbool pt = let elements = PTree.elements pt in begin - Printf.printf "["; - List.iter (fun (n, b) -> - if b then Printf.printf "%d, " (P.to_int n) else () - ) elements; - Printf.printf "]" + if !debug_flag then begin + Printf.printf "["; + List.iter (fun (n, b) -> + if b then Printf.printf "%d, " (P.to_int n) else () + ) elements; + Printf.printf "]" + end end (* Looks ahead (until a branch) to see if a node further down verifies @@ -150,7 +162,7 @@ let rec look_ahead code node is_loop_header predicate = let do_call_heuristic code cond ifso ifnot is_loop_header = begin - Printf.printf "\tCall heuristic..\n"; + debug "\tCall heuristic..\n"; let predicate n = (function | Icall _ -> true | _ -> false) @@ get_some @@ PTree.get n code @@ -164,13 +176,13 @@ let do_call_heuristic code cond ifso ifnot is_loop_header = let do_opcode_heuristic code cond ifso ifnot is_loop_header = begin - Printf.printf "\tOpcode heuristic..\n"; + debug "\tOpcode heuristic..\n"; DuplicateOpcodeHeuristic.opcode_heuristic code cond ifso ifnot is_loop_header end let do_return_heuristic code cond ifso ifnot is_loop_header = begin - Printf.printf "\tReturn heuristic..\n"; + debug "\tReturn heuristic..\n"; let predicate n = (function | Ireturn _ -> true | _ -> false) @@ get_some @@ PTree.get n code @@ -184,7 +196,7 @@ let do_return_heuristic code cond ifso ifnot is_loop_header = let do_store_heuristic code cond ifso ifnot is_loop_header = begin - Printf.printf "\tStore heuristic..\n"; + debug "\tStore heuristic..\n"; let predicate n = (function | Istore _ -> true | _ -> false) @@ get_some @@ PTree.get n code @@ -198,7 +210,7 @@ let do_store_heuristic code cond ifso ifnot is_loop_header = let do_loop_heuristic code cond ifso ifnot is_loop_header = begin - Printf.printf "\tLoop heuristic..\n"; + debug "\tLoop heuristic..\n"; let predicate n = get_some @@ PTree.get n is_loop_header in let ifso_loop = look_ahead code ifso is_loop_header predicate in let ifnot_loop = look_ahead code ifnot is_loop_header predicate in @@ -210,7 +222,7 @@ let do_loop_heuristic code cond ifso ifnot is_loop_header = let do_loop2_heuristic loop_info n code cond ifso ifnot is_loop_header = begin - Printf.printf "\tLoop2 heuristic..\n"; + debug "\tLoop2 heuristic..\n"; match get_some @@ PTree.get n loop_info with | None -> None | Some b -> Some b @@ -244,23 +256,23 @@ let get_loop_info is_loop_header bfs_order code = | Ijumptable _ | Itailcall _ | Ireturn _ -> None end in begin - Printf.printf "Marking path from %d to %d\n" (P.to_int n) (P.to_int s); + debug "Marking path from %d to %d\n" (P.to_int n) (P.to_int s); match advance_to_cb s with - | None -> (Printf.printf "Nothing found\n") - | Some s -> ( Printf.printf "Advancing to %d\n" (P.to_int s); + | None -> (debug "Nothing found\n") + | Some s -> ( debug "Advancing to %d\n" (P.to_int s); match get_some @@ PTree.get s !loop_info with | None | Some _ -> begin match get_some @@ PTree.get s code with | Icond (_, _, n1, n2, _) -> let b1 = explore n1 n in let b2 = explore n2 n in - if (b1 && b2) then (Printf.printf "both true\n") - else if b1 then (Printf.printf "true privileged\n"; loop_info := PTree.set s (Some true) !loop_info) - else if b2 then (Printf.printf "false privileged\n"; loop_info := PTree.set s (Some false) !loop_info) - else (Printf.printf "none true\n") - | _ -> ( Printf.printf "not an icond\n" ) + if (b1 && b2) then (debug "both true\n") + else if b1 then (debug "true privileged\n"; loop_info := PTree.set s (Some true) !loop_info) + else if b2 then (debug "false privileged\n"; loop_info := PTree.set s (Some false) !loop_info) + else (debug "none true\n") + | _ -> ( debug "not an icond\n" ) end - (* | Some _ -> ( Printf.printf "already loop info there\n" ) FIXME - we don't know yet whether a branch to a loop head is a backedge or not *) + (* | Some _ -> ( debug "already loop info there\n" ) FIXME - we don't know yet whether a branch to a loop head is a backedge or not *) ) end in begin @@ -278,34 +290,34 @@ let get_loop_info is_loop_header bfs_order code = (* Remark - compared to the original paper, we don't use the store heuristic *) let get_directions code entrypoint = begin - Printf.printf "get_directions\n"; flush stdout; + debug "get_directions\n"; let bfs_order = bfs code entrypoint in let is_loop_header = get_loop_headers code entrypoint in let loop_info = get_loop_info is_loop_header bfs_order code in let directions = ref (PTree.map (fun n i -> None) code) in (* None <=> no predicted direction *) begin (* ptree_printbool is_loop_header; *) - (* Printf.printf "\n"; *) + (* debug "\n"; *) List.iter (fun n -> match (get_some @@ PTree.get n code) with | Icond (cond, lr, ifso, ifnot, _) -> - (* Printf.printf "Analyzing %d.." (P.to_int n); *) + (* debug "Analyzing %d.." (P.to_int n); *) let heuristics = [ do_opcode_heuristic; do_return_heuristic; do_loop2_heuristic loop_info n; do_loop_heuristic; do_call_heuristic; (* do_store_heuristic *) ] in let preferred = ref None in begin - Printf.printf "Deciding condition for RTL node %d\n" (P.to_int n); + debug "Deciding condition for RTL node %d\n" (P.to_int n); List.iter (fun do_heur -> match !preferred with | None -> preferred := do_heur code cond ifso ifnot is_loop_header | Some _ -> () ) heuristics; directions := PTree.set n !preferred !directions; - (match !preferred with | Some false -> Printf.printf "\tFALLTHROUGH\n" - | Some true -> Printf.printf "\tBRANCH\n" - | None -> Printf.printf "\tUNSURE\n"); - Printf.printf "---------------------------------------\n" + (match !preferred with | Some false -> debug "\tFALLTHROUGH\n" + | Some true -> debug "\tBRANCH\n" + | None -> debug "\tUNSURE\n"); + debug "---------------------------------------\n" end | _ -> () ) bfs_order; @@ -325,12 +337,12 @@ let rec update_direction_rec directions = function (* Uses branch prediction to write prediction annotations in Icond *) let update_directions code entrypoint = begin - Printf.printf "Update_directions\n"; flush stdout; + debug "Update_directions\n"; let directions = get_directions code entrypoint in begin - (* Printf.printf "Ifso directions: "; + (* debug "Ifso directions: "; ptree_printbool directions; - Printf.printf "\n"; *) + debug "\n"; *) update_direction_rec directions (PTree.elements code) end end @@ -345,7 +357,7 @@ let exists_false boolmap = exists_false_rec (PTree.elements boolmap) (* DFS using prediction info to guide the exploration *) let dfs code entrypoint = begin - Printf.printf "dfs\n"; flush stdout; + debug "dfs\n"; let visited = ref (PTree.map (fun n i -> false) code) in let rec dfs_list code = function | [] -> [] @@ -409,9 +421,11 @@ let print_traces traces = | [] -> () | t::lt -> Printf.printf "\n\t"; print_trace t; Printf.printf ",\n"; f lt in begin - Printf.printf "Traces: {"; - f traces; - Printf.printf "}\n"; + if !debug_flag then begin + Printf.printf "Traces: {"; + f traces; + Printf.printf "}\n"; + end end (* Dumb (but linear) trace selection *) @@ -447,12 +461,12 @@ let select_traces_linear code entrypoint = (* Algorithm mostly inspired from Chang and Hwu 1988 * "Trace Selection for Compiling Large C Application Programs to Microcode" *) let select_traces_chang code entrypoint = begin - Printf.printf "select_traces\n"; flush stdout; + debug "select_traces\n"; let order = dfs code entrypoint in let predecessors = get_predecessors_rtl code in let traces = ref [] in let is_visited = ref (PTree.map (fun n i -> false) code) in begin (* mark all nodes visited *) - Printf.printf "Length: %d\n" (List.length order); flush stdout; + debug "Length: %d\n" (List.length order); while exists_false !is_visited do (* while (there are unvisited nodes) *) let seed = select_unvisited_node !is_visited order in let trace = ref [seed] in @@ -485,8 +499,8 @@ let select_traces_chang code entrypoint = begin end end done; - (* Printf.printf "DFS: \t"; print_intlist order; Printf.printf "\n"; *) - Printf.printf "Traces: "; print_traces !traces; + (* debug "DFS: \t"; print_intlist order; debug "\n"; *) + debug "Traces: "; print_traces !traces; !traces end end @@ -528,7 +542,7 @@ let rec change_pointers code n n' = function * n': the integer which should contain the duplicate of n * returns: new code, new ptree *) let duplicate code ptree parent n preds n' = - Printf.printf "Duplicating node %d into %d..\n" (P.to_int n) (P.to_int n'); + debug "Duplicating node %d into %d..\n" (P.to_int n) (P.to_int n'); match PTree.get n' code with | Some _ -> failwith "The PTree already has a node n'" | None -> @@ -593,8 +607,8 @@ let superblockify_traces code preds traces = | trace :: traces -> let new_code, new_ptree, nb_duplicated = tail_duplicate code preds ptree trace in if (nb_duplicated < max_nb_duplicated) - then (Printf.printf "End duplication\n"; f new_code new_ptree traces) - else (Printf.printf "Too many duplicated nodes, aborting tail duplication\n"; (code, ptree, 0)) + then (debug "End duplication\n"; f new_code new_ptree traces) + else (debug "Too many duplicated nodes, aborting tail duplication\n"; (code, ptree, 0)) in let new_code, new_ptree, _ = f code ptree traces in (new_code, new_ptree) @@ -604,7 +618,7 @@ let rec invert_iconds_trace code = function let code' = match ptree_get_some n code with | Icond (c, lr, ifso, ifnot, info) -> (match info with | Some true -> begin - (* Printf.printf "Reversing ifso/ifnot for node %d\n" (P.to_int n); *) + (* debug "Reversing ifso/ifnot for node %d\n" (P.to_int n); *) PTree.set n (Icond (Op.negate_condition c, lr, ifnot, ifso, Some false)) code end | _ -> code) diff --git a/backend/Linearizeaux.ml b/backend/Linearizeaux.ml index bfa056ca..1381877b 100644 --- a/backend/Linearizeaux.ml +++ b/backend/Linearizeaux.ml @@ -13,6 +13,12 @@ open LTL open Maps +let debug_flag = ref false + +let debug fmt = + if !debug_flag then Printf.eprintf fmt + else Printf.ifprintf stderr fmt + (* Trivial enumeration, in decreasing order of PC *) (*** @@ -115,18 +121,11 @@ let enumerate_aux_flat f reach = flatten_blocks (basic_blocks f (join_points f)) (** - * Enumeration based on traces as identified by Duplicate.v - * - * The Duplicate phase heuristically identifies the most frequented paths. Each - * Icond is modified so that the preferred condition is a fallthrough (ifnot) - * rather than a branch (ifso). + * Alternate enumeration based on traces as identified by Duplicate.v * - * The enumeration below takes advantage of this - preferring to layout nodes - * following the fallthroughs of the Lcond branches. - * - * It is slightly adapted from the work of Petris and Hansen 90 on intraprocedural - * code positioning - only we do it on a broader grain, since we don't have the exact - * frequencies (we only know which branch is the preferred one) + * This is a slight alteration to the above heuristic, ensuring that any + * superblock will be contiguous in memory, while still following the original + * heuristic *) let get_some = function @@ -145,9 +144,11 @@ let print_plist l = | [] -> () | n :: l -> Printf.printf "%d, " (P.to_int n); f l in begin - Printf.printf "["; - f l; - Printf.printf "]" + if !debug_flag then begin + Printf.printf "["; + f l; + Printf.printf "]" + end end (* adapted from the above join_points function, but with PTree *) @@ -173,7 +174,7 @@ let forward_sequences code entry = let join_points = get_join_points code entry in (* returns the list of traversed nodes, and a list of nodes to start traversing next *) let rec traverse_fallthrough code node = - (* Printf.printf "Traversing %d..\n" (P.to_int node); *) + (* debug "Traversing %d..\n" (P.to_int node); *) if not (get_some @@ PTree.get node !visited) then begin visited := PTree.set node true !visited; match PTree.get node code with @@ -182,19 +183,19 @@ let forward_sequences code entry = let ln, rem = match (last_element bb) with | Lop _ | Lload _ | Lgetstack _ | Lsetstack _ | Lstore _ | Lcall _ | Lbuiltin _ -> assert false - | Ltailcall _ | Lreturn -> begin (* Printf.printf "STOP tailcall/return\n"; *) ([], []) end + | Ltailcall _ | Lreturn -> begin (* debug "STOP tailcall/return\n"; *) ([], []) end | Lbranch n -> if get_some @@ PTree.get n join_points then ([], [n]) else let ln, rem = traverse_fallthrough code n in (ln, rem) | Lcond (_, _, ifso, ifnot, info) -> (match info with - | None -> begin (* Printf.printf "STOP Lcond None\n"; *) ([], [ifso; ifnot]) end + | None -> begin (* debug "STOP Lcond None\n"; *) ([], [ifso; ifnot]) end | Some false -> if get_some @@ PTree.get ifnot join_points then ([], [ifso; ifnot]) else let ln, rem = traverse_fallthrough code ifnot in (ln, [ifso] @ rem) | Some true -> let errstr = Printf.sprintf ("Inconsistency detected in node %d: ifnot is not the preferred branch") (P.to_int node) in failwith errstr) - | Ljumptable(_, ln) -> begin (* Printf.printf "STOP Ljumptable\n"; *) ([], ln) end + | Ljumptable(_, ln) -> begin (* debug "STOP Ljumptable\n"; *) ([], ln) end in ([node] @ ln, rem) end else ([], []) @@ -355,15 +356,19 @@ end module ISet = Set.Make(Int) let print_iset s = begin - Printf.printf "{"; - ISet.iter (fun e -> Printf.printf "%d, " e) s; - Printf.printf "}" + if !debug_flag then begin + Printf.printf "{"; + ISet.iter (fun e -> Printf.printf "%d, " e) s; + Printf.printf "}" + end end let print_depmap dm = begin - Printf.printf "[|"; - Array.iter (fun s -> print_iset s; Printf.printf ", ") dm; - Printf.printf "|]\n" + if !debug_flag then begin + Printf.printf "[|"; + Array.iter (fun s -> print_iset s; Printf.printf ", ") dm; + Printf.printf "|]\n" + end end let construct_depmap code entry fs = @@ -381,7 +386,7 @@ let construct_depmap code entry fs = !index end in let check_and_update_depmap from target = - (* Printf.printf "From %d to %d\n" (P.to_int from) (P.to_int target); *) + (* debug "From %d to %d\n" (P.to_int from) (P.to_int target); *) if not (ppmap_is_true (from, target) is_loop_edge) then let in_index_fs = find_index_of_node from in let out_index_fs = find_index_of_node target in @@ -423,14 +428,18 @@ let construct_depmap code entry fs = end let print_sequence s = - Printf.printf "["; - List.iter (fun n -> Printf.printf "%d, " (P.to_int n)) s; - Printf.printf "]\n" + if !debug_flag then begin + Printf.printf "["; + List.iter (fun n -> Printf.printf "%d, " (P.to_int n)) s; + Printf.printf "]\n" + end let print_ssequence ofs = - Printf.printf "["; - List.iter (fun s -> print_sequence s) ofs; - Printf.printf "]\n" + if !debug_flag then begin + Printf.printf "["; + List.iter (fun s -> print_sequence s) ofs; + Printf.printf "]\n" + end let order_sequences code entry fs = let fs_a = Array.of_list fs in @@ -442,13 +451,13 @@ let order_sequences code entry fs = assert (not fs_evaluated.(s_id)); ordered_fs := fs_a.(s_id) :: !ordered_fs; fs_evaluated.(s_id) <- true; - (* Printf.printf "++++++\n"; - Printf.printf "Scheduling %d\n" s_id; - Printf.printf "Initial depmap: "; print_depmap depmap; *) + (* debug "++++++\n"; + debug "Scheduling %d\n" s_id; + debug "Initial depmap: "; print_depmap depmap; *) Array.iteri (fun i deps -> depmap.(i) <- ISet.remove s_id deps ) depmap; - (* Printf.printf "Final depmap: "; print_depmap depmap; *) + (* debug "Final depmap: "; print_depmap depmap; *) end in let choose_best_of candidates = let current_best_id = ref None in @@ -478,7 +487,7 @@ let order_sequences code entry fs = begin Array.iteri (fun i deps -> begin - (* Printf.printf "Deps of %d: " i; print_iset deps; Printf.printf "\n"; *) + (* debug "Deps of %d: " i; print_iset deps; debug "\n"; *) (* FIXME - if we keep it that way (no dependency check), remove all the unneeded stuff *) if ((* deps == ISet.empty && *) not fs_evaluated.(i)) then candidates := i :: !candidates @@ -492,14 +501,14 @@ let order_sequences code entry fs = get_some (choose_best_of !candidates) end in begin - Printf.printf "-------------------------------\n"; - Printf.printf "depmap: "; print_depmap depmap; - Printf.printf "forward sequences identified: "; print_ssequence fs; + debug "-------------------------------\n"; + debug "depmap: "; print_depmap depmap; + debug "forward sequences identified: "; print_ssequence fs; while List.length !ordered_fs != List.length fs do let next_id = select_next () in evaluate next_id done; - Printf.printf "forward sequences ordered: "; print_ssequence (List.rev (!ordered_fs)); + debug "forward sequences ordered: "; print_ssequence (List.rev (!ordered_fs)); List.rev (!ordered_fs) end -- cgit